diff --git a/src/Init/Meta.lean b/src/Init/Meta.lean index b8fb5f2801aa..3c2854684bce 100644 --- a/src/Init/Meta.lean +++ b/src/Init/Meta.lean @@ -271,10 +271,11 @@ abbrev CharLit := TSyntax charLitKind abbrev NameLit := TSyntax nameLitKind abbrev ScientificLit := TSyntax scientificLitKind abbrev NumLit := TSyntax numLitKind +abbrev HygieneInfo := TSyntax hygieneInfoKind end Syntax -export Syntax (Term Command Prec Prio Ident StrLit CharLit NameLit ScientificLit NumLit) +export Syntax (Term Command Prec Prio Ident StrLit CharLit NameLit ScientificLit NumLit HygieneInfo) namespace TSyntax @@ -921,6 +922,9 @@ def getChar (s : CharLit) : Char := def getName (s : NameLit) : Name := s.raw.isNameLit?.getD .anonymous +def getHygieneInfo (s : HygieneInfo) : Name := + s.raw[0].getId + namespace Compat scoped instance : CoeTail (Array Syntax) (Syntax.TSepArray k sep) where @@ -930,6 +934,11 @@ end Compat end TSyntax +def HygieneInfo.mkIdent (s : HygieneInfo) (val : Name) (canonical := false) : Ident := + let src := s.raw[0] + let id := { extractMacroScopes src.getId with name := val.eraseMacroScopes }.review + ⟨Syntax.ident (SourceInfo.fromRef src canonical) (toString val).toSubstring id []⟩ + /-- Reflect a runtime datum back to surface syntax (best-effort). -/ class Quote (α : Type) (k : SyntaxNodeKind := `term) where quote : α → TSyntax k diff --git a/src/Init/Notation.lean b/src/Init/Notation.lean index 51b49db630ef..fa83914edb05 100644 --- a/src/Init/Notation.lean +++ b/src/Init/Notation.lean @@ -496,12 +496,6 @@ macro_rules else `(%[ $elems,* | List.nil ]) --- Declare `this` as a keyword that unhygienically binds to a scope-less `this` assumption (or other binding). --- The keyword prevents declaring a `this` binding except through metaprogramming, as is done by `have`/`show`. -/-- Special identifier introduced by "anonymous" `have : ...`, `suffices p ...` etc. -/ -macro tk:"this" : term => - return (⟨(Syntax.ident tk.getHeadInfo "this".toSubstring `this [])⟩ : TSyntax `term) - /-- 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. diff --git a/src/Init/Prelude.lean b/src/Init/Prelude.lean index be52954b44af..c3bb42d9bb9d 100644 --- a/src/Init/Prelude.lean +++ b/src/Init/Prelude.lean @@ -3721,6 +3721,16 @@ abbrev nameLitKind : SyntaxNodeKind := `name /-- `fieldIdx` is the node kind of projection indices like the `2` in `x.2`. -/ abbrev fieldIdxKind : SyntaxNodeKind := `fieldIdx +/-- +`hygieneInfo` is the node kind of the `hygieneInfo` parser, which is an +"invisible token" which captures the hygiene information at the current point +without parsing anything. + +They can be used to generate identifiers (with `Lean.HygieneInfo.mkIdent`) +as if they were introduced by the calling context, not the called macro. +-/ +abbrev hygieneInfoKind : SyntaxNodeKind := `hygieneInfo + /-- `interpolatedStrLitKind` is the node kind of interpolated string literal fragments like `"value = {` and `}"` in `s!"value = {x}"`. diff --git a/src/Init/Tactics.lean b/src/Init/Tactics.lean index 394d0e489425..845a82477345 100644 --- a/src/Init/Tactics.lean +++ b/src/Init/Tactics.lean @@ -511,7 +511,7 @@ macro "refine_lift' " e:term : tactic => `(tactic| focus (refine' no_implicit_la /-- Similar to `have`, but using `refine'` -/ macro "have' " d:haveDecl : tactic => `(tactic| refine_lift' have $d:haveDecl; ?_) /-- Similar to `have`, but using `refine'` -/ -macro (priority := high) "have'" x:ident " := " p:term : tactic => `(tactic| have' $x : _ := $p) +macro (priority := high) "have'" x:ident " := " p:term : tactic => `(tactic| have' $x:ident : _ := $p) /-- Similar to `let`, but using `refine'` -/ macro "let' " d:letDecl : tactic => `(tactic| refine_lift' let $d:letDecl; ?_) diff --git a/src/Lean/Elab/Binders.lean b/src/Lean/Elab/Binders.lean index 7db9a88b6d45..6ba8fbdb84f8 100644 --- a/src/Lean/Elab/Binders.lean +++ b/src/Lean/Elab/Binders.lean @@ -689,7 +689,7 @@ structure LetIdDeclView where value : Syntax def mkLetIdDeclView (letIdDecl : Syntax) : LetIdDeclView := - -- `letIdDecl` is of the form `ident >> many bracketedBinder >> optType >> " := " >> termParser + -- `letIdDecl` is of the form `binderIdent >> many bracketedBinder >> optType >> " := " >> termParser let id := letIdDecl[0] let binders := letIdDecl[1].getArgs let optType := letIdDecl[2] @@ -708,6 +708,7 @@ def elabLetDeclCore (stx : Syntax) (expectedType? : Option Expr) (useLetExpr : B let body := stx[3] if letDecl.getKind == ``Lean.Parser.Term.letIdDecl then let { id, binders, type, value } := mkLetIdDeclView letDecl + let id ← if id.isIdent then pure id else mkFreshIdent id (canonical := true) elabLetDeclAux id binders type value body expectedType? useLetExpr elabBodyFirst usedLetOnly else if letDecl.getKind == ``Lean.Parser.Term.letPatDecl then -- node `Lean.Parser.Term.letPatDecl $ try (termParser >> pushNone >> optType >> " := ") >> termParser @@ -717,7 +718,7 @@ def elabLetDeclCore (stx : Syntax) (expectedType? : Option Expr) (useLetExpr : B let optType := letDecl[2] let val := letDecl[4] if pat.getKind == ``Parser.Term.hole then - -- `let _ := ...` should not be treated at a `letIdDecl` + -- `let _ := ...` should not be treated as a `letIdDecl` let id ← mkFreshIdent pat (canonical := true) let type := expandOptType id optType elabLetDeclAux id #[] type val body expectedType? useLetExpr elabBodyFirst usedLetOnly diff --git a/src/Lean/Elab/BuiltinNotation.lean b/src/Lean/Elab/BuiltinNotation.lean index cfe145dc19fc..e2ff3e907da6 100644 --- a/src/Lean/Elab/BuiltinNotation.lean +++ b/src/Lean/Elab/BuiltinNotation.lean @@ -104,17 +104,25 @@ open Meta @[builtin_macro Lean.Parser.Term.have] def expandHave : Macro := fun stx => match stx with - | `(have $x $bs* $[: $type]? := $val; $body) => `(let_fun $x $bs* $[: $type]? := $val; $body) - | `(have%$tk $[: $type]? := $val; $body) => `(have $(mkIdentFrom tk `this (canonical := true)) $[: $type]? := $val; $body) - | `(have $x $bs* $[: $type]? $alts; $body) => `(let_fun $x $bs* $[: $type]? $alts; $body) - | `(have%$tk $[: $type]? $alts:matchAlts; $body) => `(have $(mkIdentFrom tk `this (canonical := true)) $[: $type]? $alts:matchAlts; $body) - | `(have $pattern:term $[: $type]? := $val:term; $body) => `(let_fun $pattern:term $[: $type]? := $val:term ; $body) - | _ => Macro.throwUnsupported + | `(have $hy:hygieneInfo $bs* $[: $type]? := $val; $body) => + `(have $(HygieneInfo.mkIdent hy `this (canonical := true)) $bs* $[: $type]? := $val; $body) + | `(have $hy:hygieneInfo $bs* $[: $type]? $alts; $body) => + `(have $(HygieneInfo.mkIdent hy `this (canonical := true)) $bs* $[: $type]? $alts; $body) + | `(have $x:ident $bs* $[: $type]? := $val; $body) => `(let_fun $x $bs* $[: $type]? := $val; $body) + | `(have $x:ident $bs* $[: $type]? $alts; $body) => `(let_fun $x $bs* $[: $type]? $alts; $body) + | `(have _%$x $bs* $[: $type]? := $val; $body) => `(let_fun _%$x $bs* $[: $type]? := $val; $body) + | `(have _%$x $bs* $[: $type]? $alts; $body) => `(let_fun _%$x $bs* $[: $type]? $alts; $body) + | `(have $pattern:term $[: $type]? := $val; $body) => `(let_fun $pattern:term $[: $type]? := $val; $body) + | _ => Macro.throwUnsupported @[builtin_macro Lean.Parser.Term.suffices] def expandSuffices : Macro - | `(suffices%$tk $[$x :]? $type from $val; $body) => `(have%$tk $[$x]? : $type := $body; $val) - | `(suffices%$tk $[$x :]? $type by%$b $tac:tacticSeq; $body) => `(have%$tk $[$x]? : $type := $body; by%$b $tac) - | _ => Macro.throwUnsupported + | `(suffices%$tk $x:ident : $type from $val; $body) => `(have%$tk $x : $type := $body; $val) + | `(suffices%$tk _%$x : $type from $val; $body) => `(have%$tk _%$x : $type := $body; $val) + | `(suffices%$tk $hy:hygieneInfo $type from $val; $body) => `(have%$tk $hy:hygieneInfo : $type := $body; $val) + | `(suffices%$tk $x:ident : $type by%$b $tac:tacticSeq; $body) => `(have%$tk $x : $type := $body; by%$b $tac) + | `(suffices%$tk _%$x : $type by%$b $tac:tacticSeq; $body) => `(have%$tk _%$x : $type := $body; by%$b $tac) + | `(suffices%$tk $hy:hygieneInfo $type by%$b $tac:tacticSeq; $body) => `(have%$tk $hy:hygieneInfo : $type := $body; by%$b $tac) + | _ => Macro.throwUnsupported open Lean.Parser in private def elabParserMacroAux (prec e : Term) (withAnonymousAntiquot : Bool) : TermElabM Syntax := do diff --git a/src/Lean/Elab/Do.lean b/src/Lean/Elab/Do.lean index 9160cd89a912..847791bc12f5 100644 --- a/src/Lean/Elab/Do.lean +++ b/src/Lean/Elab/Do.lean @@ -49,7 +49,7 @@ private def letDeclArgHasBinders (letDeclArg : Syntax) : Bool := else if k == ``Parser.Term.letEqnsDecl then true else if k == ``Parser.Term.letIdDecl then - -- letIdLhs := ident >> checkWsBefore "expected space before binders" >> many (ppSpace >> letIdBinder)) >> optType + -- letIdLhs := binderIdent >> checkWsBefore "expected space before binders" >> many (ppSpace >> letIdBinder)) >> optType let binders := letDeclArg[1] binders.getNumArgs > 0 else @@ -584,8 +584,11 @@ def concat (terminal : CodeBlock) (kRef : Syntax) (y? : Option Var) (k : CodeBlo let terminal ← liftMacroM <| convertTerminalActionIntoJmp terminal.code jp xs return { code := attachJP jpDecl terminal, uvars := k.uvars } -def getLetIdDeclVar (letIdDecl : Syntax) : Var := - letIdDecl[0] +def getLetIdDeclVars (letIdDecl : Syntax) : Array Var := + if letIdDecl[0].isIdent then + #[letIdDecl[0]] + else + #[] -- support both regular and syntax match def getPatternVarsEx (pattern : Syntax) : TermElabM (Array Var) := @@ -600,17 +603,20 @@ def getLetPatDeclVars (letPatDecl : Syntax) : TermElabM (Array Var) := do let pattern := letPatDecl[0] getPatternVarsEx pattern -def getLetEqnsDeclVar (letEqnsDecl : Syntax) : Var := - letEqnsDecl[0] +def getLetEqnsDeclVars (letEqnsDecl : Syntax) : Array Var := + if letEqnsDecl[0].isIdent then + #[letEqnsDecl[0]] + else + #[] def getLetDeclVars (letDecl : Syntax) : TermElabM (Array Var) := do let arg := letDecl[0] if arg.getKind == ``Parser.Term.letIdDecl then - return #[getLetIdDeclVar arg] + return getLetIdDeclVars arg else if arg.getKind == ``Parser.Term.letPatDecl then getLetPatDeclVars arg else if arg.getKind == ``Parser.Term.letEqnsDecl then - return #[getLetEqnsDeclVar arg] + return getLetEqnsDeclVars arg else throwError "unexpected kind of let declaration" @@ -618,11 +624,11 @@ def getDoLetVars (doLet : Syntax) : TermElabM (Array Var) := -- leading_parser "let " >> optional "mut " >> letDecl getLetDeclVars doLet[2] -def getHaveIdLhsVar (optIdent : Syntax) : TermElabM Var := - if optIdent.isNone then - `(this) +def getHaveIdLhsVar (optIdent : Syntax) : Var := + if optIdent.getKind == hygieneInfoKind then + HygieneInfo.mkIdent optIdent[0] `this else - pure optIdent[0] + optIdent def getDoHaveVars (doHave : Syntax) : TermElabM (Array Var) := do -- doHave := leading_parser "have " >> Term.haveDecl @@ -630,13 +636,13 @@ def getDoHaveVars (doHave : Syntax) : TermElabM (Array Var) := do let arg := doHave[1][0] if arg.getKind == ``Parser.Term.haveIdDecl then -- haveIdDecl := leading_parser atomic (haveIdLhs >> " := ") >> termParser - -- haveIdLhs := optional (ident >> many (ppSpace >> letIdBinder)) >> optType - return #[← getHaveIdLhsVar arg[0]] + -- haveIdLhs := (binderIdent <|> hygieneInfo) >> many letIdBinder >> optType + return #[getHaveIdLhsVar arg[0]] else if arg.getKind == ``Parser.Term.letPatDecl then getLetPatDeclVars arg else if arg.getKind == ``Parser.Term.haveEqnsDecl then -- haveEqnsDecl := leading_parser haveIdLhs >> matchAlts - return #[← getHaveIdLhsVar arg[0]] + return #[getHaveIdLhsVar arg[0]] else throwError "unexpected kind of have declaration" @@ -672,7 +678,7 @@ def getDoLetArrowVars (doLetArrow : Syntax) : TermElabM (Array Var) := do def getDoReassignVars (doReassign : Syntax) : TermElabM (Array Var) := do let arg := doReassign[0] if arg.getKind == ``Parser.Term.letIdDecl then - return #[getLetIdDeclVar arg] + return getLetIdDeclVars arg else if arg.getKind == ``Parser.Term.letPatDecl then getLetPatDeclVars arg else diff --git a/src/Lean/Elab/LetRec.lean b/src/Lean/Elab/LetRec.lean index 41a5153137f1..73e49fe7df8d 100644 --- a/src/Lean/Elab/LetRec.lean +++ b/src/Lean/Elab/LetRec.lean @@ -37,6 +37,8 @@ private def mkLetRecDeclView (letRec : Syntax) : TermElabM LetRecView := do throwErrorAt decl "patterns are not allowed in 'let rec' expressions" else if decl.isOfKind `Lean.Parser.Term.letIdDecl || decl.isOfKind `Lean.Parser.Term.letEqnsDecl then let declId := decl[0] + unless declId.isIdent do + throwErrorAt declId "'let rec' expressions must be named" let shortDeclName := declId.getId let currDeclName? ← getDeclName? let declName := currDeclName?.getD Name.anonymous ++ shortDeclName diff --git a/src/Lean/Elab/MutualDef.lean b/src/Lean/Elab/MutualDef.lean index 522c15719f7a..e37eb7d206ef 100644 --- a/src/Lean/Elab/MutualDef.lean +++ b/src/Lean/Elab/MutualDef.lean @@ -210,6 +210,7 @@ private def expandWhereStructInst : Macro have : Coe (TSyntax ``letIdBinder) (TSyntax ``funBinder) := ⟨(⟨·⟩)⟩ val ← if binders.size > 0 then `(fun $binders* => $val) else pure val `(structInstField|$id:ident := $val) + | stx@`(letIdDecl|_ $_* $[: $_]? := $_) => Macro.throwErrorAt stx "'_' is not allowed here" | _ => Macro.throwUnsupported let body ← `({ $structInstFields,* }) match whereDecls? with diff --git a/src/Lean/Elab/Syntax.lean b/src/Lean/Elab/Syntax.lean index 165c8e46631d..0e231ddfb87f 100644 --- a/src/Lean/Elab/Syntax.lean +++ b/src/Lean/Elab/Syntax.lean @@ -260,7 +260,7 @@ open Lean.Parser.Command private def declareSyntaxCatQuotParser (catName : Name) : CommandElabM Unit := do if let .str _ suffix := catName then - let quotSymbol := "`(" ++ suffix ++ "|" + let quotSymbol := "`(" ++ suffix ++ "| " let name := catName ++ `quot let cmd ← `( @[term_parser] def $(mkIdent name) : Lean.ParserDescr := diff --git a/src/Lean/Parser.lean b/src/Lean/Parser.lean index 8b3683a99885..7745db80ed20 100644 --- a/src/Lean/Parser.lean +++ b/src/Lean/Parser.lean @@ -27,7 +27,8 @@ builtin_initialize 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 (kind := identKind) ident + register_parser_alias (kind := hygieneInfoKind) hygieneInfo register_parser_alias "colGt" checkColGt { stackSz? := some 0 } register_parser_alias "colGe" checkColGe { stackSz? := some 0 } register_parser_alias "colEq" checkColEq { stackSz? := some 0 } diff --git a/src/Lean/Parser/Basic.lean b/src/Lean/Parser/Basic.lean index 0eb426b3ad19..e7eef8967dd4 100644 --- a/src/Lean/Parser/Basic.lean +++ b/src/Lean/Parser/Basic.lean @@ -1072,6 +1072,20 @@ def identEq (id : Name) : Parser := { info := mkAtomicInfo "ident" } +def hygieneInfoFn : ParserFn := fun c s => + let input := c.input + let pos := s.pos + let str := mkEmptySubstringAt input pos + let info := SourceInfo.original str pos str pos + let ident := mkIdent info str .anonymous + let stx := mkNode hygieneInfoKind #[ident] + s.pushSyntax stx + +def hygieneInfoNoAntiquot : Parser := { + fn := hygieneInfoFn + info := nodeInfo hygieneInfoKind epsilonInfo +} + namespace ParserState def keepTop (s : SyntaxStack) (startStackSize : Nat) : SyntaxStack := diff --git a/src/Lean/Parser/Extra.lean b/src/Lean/Parser/Extra.lean index 10817d1e3239..a00944f6f648 100644 --- a/src/Lean/Parser/Extra.lean +++ b/src/Lean/Parser/Extra.lean @@ -37,6 +37,9 @@ attribute [run_builtin_parser_attribute_hooks] @[run_builtin_parser_attribute_hooks] def rawIdent : Parser := withAntiquot (mkAntiquot "ident" identKind) rawIdentNoAntiquot +@[run_builtin_parser_attribute_hooks] def hygieneInfo : Parser := + withAntiquot (mkAntiquot "hygieneInfo" hygieneInfoKind (anonymous := false)) hygieneInfoNoAntiquot + @[run_builtin_parser_attribute_hooks] def numLit : Parser := withAntiquot (mkAntiquot "num" numLitKind) numLitNoAntiquot diff --git a/src/Lean/Parser/Term.lean b/src/Lean/Parser/Term.lean index 5654ed66ca94..47d0853c0ead 100644 --- a/src/Lean/Parser/Term.lean +++ b/src/Lean/Parser/Term.lean @@ -119,6 +119,7 @@ def optSemicolon (p : Parser) : Parser := "_" @[builtin_term_parser] def syntheticHole := leading_parser "?" >> (ident <|> hole) +def binderIdent : Parser := ident <|> hole /-- A temporary placeholder for a missing proof or value. -/ @[builtin_term_parser] def «sorry» := leading_parser "sorry" @@ -165,7 +166,7 @@ def fromTerm := leading_parser "from " >> termParser def showRhs := fromTerm <|> byTactic' def sufficesDecl := leading_parser - optIdent >> termParser >> ppSpace >> showRhs + (atomic (group (binderIdent >> " : ")) <|> hygieneInfo) >> termParser >> ppSpace >> showRhs @[builtin_term_parser] def «suffices» := leading_parser:leadPrec withPosition ("suffices " >> sufficesDecl) >> optSemicolon termParser @[builtin_term_parser] def «show» := leading_parser:leadPrec "show " >> termParser >> ppSpace >> showRhs @@ -209,7 +210,6 @@ In contrast to regular patterns, `e` may be an arbitrary term of the appropriate -/ @[builtin_term_parser] def inaccessible := leading_parser ".(" >> withoutPosition termParser >> ")" -def binderIdent : Parser := ident <|> hole def binderType (requireType := false) : Parser := if requireType then node nullKind (" : " >> termParser) else optional (" : " >> termParser) def binderTactic := leading_parser @@ -394,7 +394,7 @@ def letIdBinder := binderIdent <|> 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 "" >> "[") + binderIdent >> 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) @@ -467,7 +467,7 @@ It is often used when building macros. withPosition ("let_tmp " >> letDecl) >> optSemicolon termParser /- like `let_fun` but with optional name -/ -def haveIdLhs := optional (ppSpace >> ident >> many (ppSpace >> letIdBinder)) >> optType +def haveIdLhs := ((ppSpace >> binderIdent) <|> hygieneInfo) >> many (ppSpace >> letIdBinder) >> optType def haveIdDecl := leading_parser (withAnonymousAntiquot := false) atomic (haveIdLhs >> " := ") >> termParser def haveEqnsDecl := leading_parser (withAnonymousAntiquot := false) diff --git a/src/Lean/PrettyPrinter/Formatter.lean b/src/Lean/PrettyPrinter/Formatter.lean index 2f9beb8f60dd..da8effa217f0 100644 --- a/src/Lean/PrettyPrinter/Formatter.lean +++ b/src/Lean/PrettyPrinter/Formatter.lean @@ -489,6 +489,7 @@ def sepByNoAntiquot.formatter (p pSep : Formatter) : Formatter := do @[combinator_formatter skip] def skip.formatter : Formatter := pure () @[combinator_formatter pushNone] def pushNone.formatter : Formatter := goLeft +@[combinator_formatter hygieneInfoNoAntiquot] def hygieneInfoNoAntiquot.formatter : Formatter := goLeft @[combinator_formatter interpolatedStr] def interpolatedStr.formatter (p : Formatter) : Formatter := do diff --git a/src/Lean/PrettyPrinter/Parenthesizer.lean b/src/Lean/PrettyPrinter/Parenthesizer.lean index 988905f76529..b007273b652b 100644 --- a/src/Lean/PrettyPrinter/Parenthesizer.lean +++ b/src/Lean/PrettyPrinter/Parenthesizer.lean @@ -492,6 +492,7 @@ def sepByNoAntiquot.parenthesizer (p pSep : Parenthesizer) : Parenthesizer := do @[combinator_parenthesizer skip] def skip.parenthesizer : Parenthesizer := pure () @[combinator_parenthesizer pushNone] def pushNone.parenthesizer : Parenthesizer := goLeft +@[combinator_parenthesizer hygieneInfoNoAntiquot] def hygieneInfoNoAntiquot.parenthesizer : Parenthesizer := goLeft @[combinator_parenthesizer interpolatedStr] def interpolatedStr.parenthesizer (p : Parenthesizer) : Parenthesizer := do diff --git a/stage0/src/library/profiling.cpp b/stage0/src/library/profiling.cpp index 46dcb55af075..d48bbceeff61 100644 --- a/stage0/src/library/profiling.cpp +++ b/stage0/src/library/profiling.cpp @@ -7,40 +7,23 @@ Author: Gabriel Ebner #include "library/profiling.h" #include "util/option_declarations.h" -#ifndef LEAN_DEFAULT_PROFILER -#define LEAN_DEFAULT_PROFILER false -#endif - -#ifndef LEAN_DEFAULT_PROFILER_THRESHOLD -#define LEAN_DEFAULT_PROFILER_THRESHOLD 100 -#endif - namespace lean { -static name * g_profiler = nullptr; -static name * g_profiler_threshold = nullptr; - +extern "C" uint8_t lean_get_profiler(obj_arg opts); bool get_profiler(options const & opts) { - return opts.get_bool(*g_profiler, LEAN_DEFAULT_PROFILER); + return lean_get_profiler(opts.to_obj_arg()); } +extern "C" double lean_get_profiler_threshold(obj_arg opts); second_duration get_profiling_threshold(options const & opts) { - return second_duration(static_cast(opts.get_unsigned(*g_profiler_threshold, LEAN_DEFAULT_PROFILER_THRESHOLD))/1000.0); + double ms = lean_get_profiler_threshold(opts.to_obj_arg()); + return second_duration(ms); } void initialize_profiling() { - g_profiler = new name{"profiler"}; - mark_persistent(g_profiler->raw()); - g_profiler_threshold = new name{"profiler", "threshold"}; - mark_persistent(g_profiler_threshold->raw()); - register_bool_option(*g_profiler, LEAN_DEFAULT_PROFILER, "(profiler) profile tactics and vm_eval command"); - register_unsigned_option(*g_profiler_threshold, LEAN_DEFAULT_PROFILER_THRESHOLD, - "(profiler) threshold in milliseconds, profiling times under threshold will not be reported"); } void finalize_profiling() { - delete g_profiler; - delete g_profiler_threshold; } } diff --git a/stage0/src/runtime/io.cpp b/stage0/src/runtime/io.cpp index 8d378f02c36d..4cc452c70bee 100644 --- a/stage0/src/runtime/io.cpp +++ b/stage0/src/runtime/io.cpp @@ -410,6 +410,7 @@ extern "C" LEAN_EXPORT obj_res lean_io_get_random_bytes (size_t nbytes, obj_arg obj_res res = lean_alloc_sarray(1, 0, nbytes); size_t remain = nbytes; + uint8_t *dst = lean_sarray_cptr(res); while (remain > 0) { #if defined(LEAN_WINDOWS) @@ -417,7 +418,7 @@ extern "C" LEAN_EXPORT obj_res lean_io_get_random_bytes (size_t nbytes, obj_arg size_t read_sz = std::min(remain, static_cast(std::numeric_limits::max())); NTSTATUS status = BCryptGenRandom( NULL, - lean_sarray_cptr(res), + dst, static_cast(read_sz), BCRYPT_USE_SYSTEM_PREFERRED_RNG ); @@ -426,6 +427,7 @@ extern "C" LEAN_EXPORT obj_res lean_io_get_random_bytes (size_t nbytes, obj_arg return io_result_mk_error("BCryptGenRandom failed"); } remain -= read_sz; + dst += read_sz; #else #if defined(LEAN_EMSCRIPTEN) // `Crypto.getRandomValues` documents `dest` should be at most 65536 bytes. @@ -433,7 +435,7 @@ extern "C" LEAN_EXPORT obj_res lean_io_get_random_bytes (size_t nbytes, obj_arg #else size_t read_sz = remain; #endif - ssize_t nread = read(fd_urandom, lean_sarray_cptr(res), read_sz); + ssize_t nread = read(fd_urandom, dst, read_sz); if (nread < 0) { if (errno != EINTR) { close(fd_urandom); @@ -442,6 +444,7 @@ extern "C" LEAN_EXPORT obj_res lean_io_get_random_bytes (size_t nbytes, obj_arg } } else { remain -= nread; + dst += nread; } #endif } diff --git a/stage0/stdlib/Init/Conv.c b/stage0/stdlib/Init/Conv.c index d50043e91d07..fb1e8e702cae 100644 --- a/stage0/stdlib/Init/Conv.c +++ b/stage0/stdlib/Init/Conv.c @@ -13,6 +13,7 @@ #ifdef __cplusplus extern "C" { #endif +static lean_object* l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__1; 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_simp___closed__5; static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convApply____1___closed__1; @@ -22,8 +23,8 @@ static lean_object* l_Lean_Parser_Tactic_Conv_first___closed__9; static lean_object* l_Lean_Parser_Tactic_Conv_arg___closed__8; static lean_object* l_Lean_Parser_Tactic_Conv_convRepeat_____closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_occsIndexed___closed__5; +static lean_object* l_Lean_Parser_Tactic_Conv_first___closed__25; static lean_object* l_Lean_Parser_Tactic_Conv_convSeqBracketed___closed__11; -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e__; static lean_object* l_Lean_Parser_Tactic_Conv_nestedTactic___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_dsimp; static lean_object* l_Lean_Parser_Tactic_Conv_simpMatch___closed__2; @@ -32,20 +33,20 @@ lean_object* lean_mk_empty_array_with_capacity(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_rewrite; static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__anyGoals__1___closed__1; static lean_object* l_Lean_Parser_Tactic_Conv_change___closed__2; +static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext_____x3d_x3e____1___closed__1; +static lean_object* l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__4; static lean_object* l_Lean_Parser_Tactic_Conv_convArgs___closed__3; -static lean_object* l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__10; static lean_object* l_Lean_Parser_Tactic_Conv_arg___closed__4; static lean_object* l_Lean_Parser_Tactic_Conv_case_x27___closed__1; -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_convIntro____; -static lean_object* l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__4; +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convEnter_x5b___x5d__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_allGoals; static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__case__1___closed__2; static lean_object* l_Lean_Parser_Tactic_Conv_rewrite___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_convRw____; static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convRfl__1___closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_change; +static lean_object* l_Lean_Parser_Tactic_Conv_first___closed__20; 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*); -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convEnter_x5b_____x5d__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv_case___closed__2; 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_case_x27___closed__3; @@ -53,7 +54,6 @@ static lean_object* l_Lean_Parser_Tactic_Conv_conv_xb7_x2e_____closed__9; static lean_object* l_Lean_Parser_Tactic_Conv_convSeq1Indented___closed__4; static lean_object* l_Lean_Parser_Tactic_Conv_convApply_____closed__1; static lean_object* l_Lean_Parser_Tactic_Conv_paren___closed__5; -static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext___x3d_x3e____1___closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_first___closed__5; static lean_object* l_Lean_Parser_Tactic_Conv_whnf___closed__4; static lean_object* l_Lean_Parser_Tactic_Conv_dsimp___closed__10; @@ -63,9 +63,9 @@ static lean_object* l_Lean_Parser_Tactic_Conv_case___closed__8; static lean_object* l_Lean_Parser_Tactic_Conv_convSeqBracketed___closed__2; static lean_object* l_Lean_Parser_Tactic_Conv_first___closed__8; static lean_object* l_Lean_Parser_Tactic_Conv_skip___closed__2; -static lean_object* l_Lean_Parser_Tactic_Conv_dsimp___closed__17; static lean_object* l_Lean_Parser_Tactic_Conv_convSeq___closed__5; static lean_object* l_Lean_Parser_Tactic_Conv_simpMatch___closed__4; +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d; static lean_object* l_Lean_Parser_Tactic_Conv_arg___closed__11; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_zeta; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_convSeq1Indented; @@ -80,16 +80,14 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_lhs; static lean_object* l_Lean_Parser_Tactic_Conv_conv___closed__13; static lean_object* l_Lean_Parser_Tactic_Conv_anyGoals___closed__6; static lean_object* l_Lean_Parser_Tactic_Conv_reduce___closed__4; -static lean_object* l_Lean_Parser_Tactic_Conv_conv___closed__16; +static lean_object* l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_occs___closed__17; static lean_object* l_Lean_Parser_Tactic_Conv_conv_xb7_x2e_____closed__14; static lean_object* l_Lean_Parser_Tactic_Conv_whnf___closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_conv___closed__15; static lean_object* l_Lean_Parser_Tactic_Conv_withAnnotateState___closed__8; -static lean_object* l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convRw______1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv_congr___closed__2; -static lean_object* l_Lean_Parser_Tactic_Conv_dsimp___closed__15; static lean_object* l_Lean_Parser_Tactic_Conv_congr___closed__4; extern lean_object* l_Lean_Parser_Tactic_simpStar; static lean_object* l_Lean_Parser_Tactic_Conv_delta___closed__1; @@ -107,9 +105,11 @@ static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules_ static lean_object* l_Lean_Parser_Tactic_Conv_nestedTacticCore___closed__11; static lean_object* l_Lean_Parser_Tactic_Conv_convConvSeq___closed__4; static lean_object* l_Lean_Parser_Tactic_Conv_nestedTactic___closed__1; +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext_____x3d_x3e____1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv_simp___closed__16; static lean_object* l_Lean_Parser_Tactic_Conv_convSeq1Indented___closed__8; static lean_object* l_Lean_Parser_Tactic_Conv_conv_quot___closed__16; +static lean_object* l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__6; static lean_object* l_Lean_Parser_Tactic_Conv_occs___closed__9; static lean_object* l_Lean_Parser_Tactic_Conv_dsimp___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_occs; @@ -121,12 +121,14 @@ static lean_object* l_Lean_Parser_Tactic_Conv_convSeqBracketed___closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_whnf___closed__1; lean_object* l_Lean_Syntax_TSepArray_getElems___rarg(lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv_ext___closed__3; +static lean_object* l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__anyGoals__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_convTactic; static lean_object* l_Lean_Parser_Tactic_Conv_conv___x3c_x3b_x3e_____closed__4; static lean_object* l_Lean_Parser_Tactic_Conv_occs___closed__6; static lean_object* l_Lean_Parser_Tactic_Conv_case___closed__1; static lean_object* l_Lean_Parser_Tactic_Conv_conv_xb7_x2e_____closed__6; +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e__; static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv___x3c_x3b_x3e____1___closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_delta___closed__4; static lean_object* l_Lean_Parser_Tactic_Conv_failIfSuccess___closed__1; @@ -144,7 +146,6 @@ lean_object* l_Lean_Name_mkStr5(lean_object*, lean_object*, lean_object*, lean_o static lean_object* l_Lean_Parser_Tactic_Conv_conv___x3c_x3b_x3e_____closed__2; static lean_object* l_Lean_Parser_Tactic_Conv_conv_xb7_x2e_____closed__12; static lean_object* l_Lean_Parser_Tactic_Conv_conv_quot___closed__18; -static lean_object* l_Lean_Parser_Tactic_Conv_convIntro_______closed__6; static lean_object* l_Lean_Parser_Tactic_Conv_pattern___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_paren; static lean_object* l_Lean_Parser_Tactic_Conv_conv___closed__6; @@ -155,35 +156,34 @@ static lean_object* l_Lean_Parser_Tactic_Conv_convRfl___closed__2; static lean_object* l_Lean_Parser_Tactic_Conv_convErw_____closed__1; 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_occsIndexed___closed__6; -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convEnter_x5b_____x5d__1___lambda__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv_convRfl___closed__5; static lean_object* l_Lean_Parser_Tactic_Conv_conv_quot___closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_occs___closed__16; static lean_object* l_Lean_Parser_Tactic_Conv_unfold___closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_conv_quot___closed__14; -static lean_object* l_Lean_Parser_Tactic_Conv_dsimp___closed__19; +static lean_object* l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_convSeq; static lean_object* l_Lean_Parser_Tactic_Conv_convLeft___closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_convTactic___closed__5; static lean_object* l_Lean_Parser_Tactic_Conv_rhs___closed__4; static lean_object* l_Lean_Parser_Tactic_Conv_conv_xb7_x2e_____closed__4; -static lean_object* l_Lean_Parser_Tactic_Conv_convIntro_______closed__2; static lean_object* l_Lean_Parser_Tactic_Conv_first___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_convSeqBracketed; static lean_object* l_Lean_Parser_Tactic_Conv_enterArg___closed__1; -static lean_object* l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__1; static lean_object* l_Lean_Parser_Tactic_Conv_rewrite___closed__7; static lean_object* l_Lean_Parser_Tactic_Conv_dsimp___closed__4; static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convErw____1___closed__25; static lean_object* l_Lean_Parser_Tactic_Conv_conv_quot___closed__6; static lean_object* l_Lean_Parser_Tactic_Conv_simp___closed__25; static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__case_x27__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_convIntro______; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__case_x27__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv_nestedTactic___closed__4; static lean_object* l_Lean_Parser_Tactic_Conv_anyGoals___closed__4; lean_object* l_Lean_Syntax_getNumArgs(lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv_unfold___closed__5; static lean_object* l_Lean_Parser_Tactic_Conv_zeta___closed__2; +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_first___closed__6; extern lean_object* l_Lean_Parser_Tactic_discharger; static lean_object* l_Lean_Parser_Tactic_Conv_simp___closed__26; @@ -191,8 +191,8 @@ static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules_ LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_whnf; static lean_object* l_Lean_Parser_Tactic_Conv_occs___closed__18; static lean_object* l_Lean_Parser_Tactic_Conv_withAnnotateState___closed__12; -static lean_object* l_Lean_Parser_Tactic_Conv_conv___closed__17; static lean_object* l_Lean_Parser_Tactic_Conv_conv_xb7_x2e_____closed__16; +static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext_____x3d_x3e____1___closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_withAnnotateState___closed__7; static lean_object* l_Lean_Parser_Tactic_Conv_convRfl___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_convTrace__state; @@ -213,6 +213,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_convLeft; static lean_object* l_Lean_Parser_Tactic_Conv_withAnnotateState___closed__13; static lean_object* l_Lean_Parser_Tactic_Conv_focus___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_ext; +static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext_____x3d_x3e____1___closed__2; static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__case__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__case__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv_convConvSeq___closed__2; @@ -225,29 +226,28 @@ static lean_object* l_Lean_Parser_Tactic_Conv_convRw_______closed__5; static lean_object* l_Lean_Parser_Tactic_Conv_convRepeat_____closed__6; static lean_object* l_Lean_Parser_Tactic_Conv_change___closed__1; static lean_object* l_Lean_Parser_Tactic_Conv_simpMatch___closed__3; +static lean_object* l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__7; 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_nestedTactic___closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_conv___x3c_x3b_x3e_____closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_conv_quot; -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext___x3d_x3e____1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv_occsIndexed___closed__9; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_convApply__; -static lean_object* l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__5; static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convErw____1___closed__12; lean_object* l_Lean_SourceInfo_fromRef(lean_object*, uint8_t); -static lean_object* l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__6; static lean_object* l_Lean_Parser_Tactic_Conv_occsIndexed___closed__4; static lean_object* l_Lean_Parser_Tactic_Conv_convTactic___closed__1; static lean_object* l_Lean_Parser_Tactic_Conv_arg___closed__3; lean_object* l_Lean_Syntax_node6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv_anyGoals___closed__1; +static lean_object* l_Lean_Parser_Tactic_Conv_first___closed__23; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_failIfSuccess; static lean_object* l_Lean_Parser_Tactic_Conv_case_x27___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__allGoals__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv_case_x27___closed__4; static lean_object* l_Lean_Parser_Tactic_Conv_congr___closed__3; -static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext___x3d_x3e____1___closed__4; static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convRfl__1___closed__5; +static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext_____x3d_x3e____1___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_arg; static lean_object* l_Lean_Parser_Tactic_Conv_convApply_____closed__2; static lean_object* l_Lean_Parser_Tactic_Conv_case_x27___closed__7; @@ -257,21 +257,22 @@ static lean_object* l_Lean_Parser_Tactic_Conv_first___closed__14; static lean_object* l_Lean_Parser_Tactic_Conv_change___closed__9; static lean_object* l_Lean_Parser_Tactic_Conv_convConvSeq___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convErw____1(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__7; -static lean_object* l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__8; static lean_object* l_Lean_Parser_Tactic_Conv_simp___closed__2; extern lean_object* l_Lean_Parser_Tactic_rwRuleSeq; static lean_object* l_Lean_Parser_Tactic_Conv_case___closed__4; static lean_object* l_Lean_Parser_Tactic_Conv_lhs___closed__1; +static lean_object* l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__9; static lean_object* l_Lean_Parser_Tactic_Conv_convSeq___closed__4; static lean_object* l_Lean_Parser_Tactic_Conv_allGoals___closed__1; extern lean_object* l_Lean_Parser_Tactic_simpLemma; static lean_object* l_Lean_Parser_Tactic_Conv_failIfSuccess___closed__5; static lean_object* l_Lean_Parser_Tactic_Conv_convTactic___closed__2; +static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext_____x3d_x3e____1___closed__5; static lean_object* l_Lean_Parser_Tactic_Conv_rewrite___closed__6; static lean_object* l_Lean_Parser_Tactic_Conv_convLeft___closed__5; static lean_object* l_Lean_Parser_Tactic_Conv_rewrite___closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_convSeqBracketed___closed__7; +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convEnter_x5b___x5d__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv_conv___x3c_x3b_x3e_____closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_simp___closed__17; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_reduce; @@ -287,6 +288,7 @@ static lean_object* l_Lean_Parser_Tactic_Conv_focus___closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_conv_xb7_x2e_____closed__15; static lean_object* l_Lean_Parser_Tactic_Conv_failIfSuccess___closed__4; static lean_object* l_Lean_Parser_Tactic_Conv_convSeqBracketed___closed__8; +static lean_object* l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__2; static lean_object* l_Lean_Parser_Tactic_Conv_convRepeat_____closed__4; lean_object* l_Lean_Syntax_getOptional_x3f(lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv_arg___closed__9; @@ -300,7 +302,6 @@ static lean_object* l_Lean_Parser_Tactic_Conv_arg___closed__12; static lean_object* l_Lean_Parser_Tactic_Conv_conv_xb7_x2e_____closed__8; static lean_object* l_Lean_Parser_Tactic_Conv_convTactic___closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_convTrace__state___closed__4; -static lean_object* l_Lean_Parser_Tactic_Conv_convIntro_______closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_convSeqBracketed___closed__12; static lean_object* l_Lean_Parser_Tactic_Conv_convRepeat_____closed__1; static lean_object* l_Lean_Parser_Tactic_Conv_first___closed__16; @@ -308,6 +309,7 @@ static lean_object* l_Lean_Parser_Tactic_Conv_convSeq___closed__2; static lean_object* l_Lean_Parser_Tactic_Conv_dsimp___closed__2; static lean_object* l_Lean_Parser_Tactic_Conv_delta___closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_conv___x3c_x3b_x3e_____closed__1; +static lean_object* l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__5; lean_object* l_Lean_Syntax_node3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv_conv_xb7_x2e_____closed__2; static lean_object* l_Lean_Parser_Tactic_Conv_convSeq1Indented___closed__1; @@ -317,12 +319,10 @@ static lean_object* l_Lean_Parser_Tactic_Conv_convRepeat_____closed__2; static lean_object* l_Lean_Parser_Tactic_Conv_simpMatch___closed__1; static lean_object* l_Lean_Parser_Tactic_Conv_convSeqBracketed___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_delta; -static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convEnter_x5b_____x5d__1___closed__1; static lean_object* l_Lean_Parser_Tactic_Conv_convArgs___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_convArgs; static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__focus__1___closed__1; static lean_object* l_Lean_Parser_Tactic_Conv_arg___closed__5; -static lean_object* l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__9; static lean_object* l_Lean_Parser_Tactic_Conv_convTry_____closed__1; static lean_object* l_Lean_Parser_Tactic_Conv_nestedTacticCore___closed__6; static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convErw____1___closed__10; @@ -331,24 +331,29 @@ static lean_object* l_Lean_Parser_Tactic_Conv_lhs___closed__2; static lean_object* l_Lean_Parser_Tactic_Conv_allGoals___closed__5; static lean_object* l_Lean_Parser_Tactic_Conv_convRw_______closed__1; static lean_object* l_Lean_Parser_Tactic_Conv_lhs___closed__3; -static lean_object* l_Lean_Parser_Tactic_Conv_convIntro_______closed__4; static lean_object* l_Lean_Parser_Tactic_Conv_withAnnotateState___closed__4; static lean_object* l_Lean_Parser_Tactic_Conv_arg___closed__2; static lean_object* l_Lean_Parser_Tactic_Conv_simp___closed__6; static lean_object* l_Lean_Parser_Tactic_Conv_occs___closed__2; static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convTrace__state__1___closed__2; +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convEnter_x5b___x5d__1___lambda__1(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv_change___closed__5; static lean_object* l_Lean_Parser_Tactic_Conv_withAnnotateState___closed__3; +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_convIntro_________closed__5; static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convErw____1___closed__13; static lean_object* l_Lean_Parser_Tactic_Conv_paren___closed__3; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv_occs___closed__4; lean_object* l_Lean_Syntax_node2(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic_Conv_convIntro_________closed__2; 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_nestedTacticCore___closed__7; +static lean_object* l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__7; lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv_ext___closed__11; +static lean_object* l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__8; static lean_object* l_Lean_Parser_Tactic_Conv_reduce___closed__1; static lean_object* l_Lean_Parser_Tactic_Conv_occsIndexed___closed__1; static lean_object* l_Lean_Parser_Tactic_Conv_rhs___closed__2; @@ -356,6 +361,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_nestedTactic; static lean_object* l_Lean_Parser_Tactic_Conv_occsIndexed___closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_conv_quot___closed__7; uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__10; static lean_object* l_Lean_Parser_Tactic_Conv_convRw_______closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_paren___closed__4; static lean_object* l_Lean_Parser_Tactic_Conv_conv___x3c_x3b_x3e_____closed__5; @@ -363,17 +369,17 @@ static lean_object* l_Lean_Parser_Tactic_Conv_first___closed__7; static lean_object* l_Lean_Parser_Tactic_Conv_convRight___closed__1; static lean_object* l_Lean_Parser_Tactic_Conv_allGoals___closed__4; static lean_object* l_Lean_Parser_Tactic_Conv_convApply_____closed__5; -static lean_object* l_Lean_Parser_Tactic_Conv_convIntro_______closed__1; +static lean_object* l_Lean_Parser_Tactic_Conv_first___closed__22; static lean_object* l_Lean_Parser_Tactic_Conv_ext___closed__14; static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__failIfSuccess__1___closed__2; static lean_object* l_Lean_Parser_Tactic_Conv_delta___closed__6; static lean_object* l_Lean_Parser_Tactic_Conv_convSeq1Indented___closed__7; -static lean_object* l_Lean_Parser_Tactic_Conv_unfold___closed__6; static lean_object* l_Lean_Parser_Tactic_Conv_nestedTactic___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_convRight; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_occsIndexed; static lean_object* l_Lean_Parser_Tactic_Conv_lhs___closed__4; static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convApply____1___closed__2; +static lean_object* l_Lean_Parser_Tactic_Conv_convIntro_________closed__1; static lean_object* l_Lean_Parser_Tactic_Conv_simp___closed__23; static lean_object* l_Lean_Parser_Tactic_Conv_conv_quot___closed__4; static lean_object* l_Lean_Parser_Tactic_Conv_rhs___closed__3; @@ -387,12 +393,13 @@ static lean_object* l_Lean_Parser_Tactic_Conv_ext___closed__15; static lean_object* l_Lean_Parser_Tactic_Conv_delta___closed__5; static lean_object* l_Lean_Parser_Tactic_Conv_conv_xb7_x2e_____closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_conv___closed__8; -static lean_object* l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__8; +static lean_object* l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_case_x27___closed__5; lean_object* l_Array_append___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convRw______1___closed__1; static lean_object* l_Lean_Parser_Tactic_Conv_convRight___closed__4; static lean_object* l_Lean_Parser_Tactic_Conv_convRw_______closed__7; +static lean_object* l_Lean_Parser_Tactic_Conv_ext___closed__16; static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convErw____1___closed__16; static lean_object* l_Lean_Parser_Tactic_Conv_convArgs___closed__1; static lean_object* l_Lean_Parser_Tactic_Conv_focus___closed__6; @@ -404,13 +411,11 @@ extern lean_object* l_Lean_Parser_Tactic_simpErase; static lean_object* l_Lean_Parser_Tactic_Conv_case_x27___closed__6; static lean_object* l_Lean_Parser_Tactic_Conv_first___closed__11; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_occsWildcard; -static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext___x3d_x3e____1___closed__1; static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__anyGoals__1___closed__2; lean_object* l_Array_extract___rarg(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_node4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv_failIfSuccess___closed__6; static lean_object* l_Lean_Parser_Tactic_Conv_occs___closed__14; -static lean_object* l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__1; static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convErw____1___closed__22; static lean_object* l_Lean_Parser_Tactic_Conv_convDone___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_case; @@ -433,7 +438,6 @@ static lean_object* l_Lean_Parser_Tactic_Conv_first___closed__10; static lean_object* l_Lean_Parser_Tactic_Conv_enterArg___closed__6; static lean_object* l_Lean_Parser_Tactic_Conv_pattern___closed__8; 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__convEnter_x5b_____x5d__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv_conv_quot___closed__11; 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__convRfl__1___closed__2; @@ -446,6 +450,7 @@ static lean_object* l_Lean_Parser_Tactic_Conv_nestedTacticCore___closed__4; static lean_object* l_Lean_Parser_Tactic_Conv_ext___closed__2; static lean_object* l_Lean_Parser_Tactic_Conv_simp___closed__10; static lean_object* l_Lean_Parser_Tactic_Conv_simp___closed__19; +static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext_____x3d_x3e____1___closed__7; static lean_object* l_Lean_Parser_Tactic_Conv_ext___closed__6; static lean_object* l_Lean_Parser_Tactic_Conv_zeta___closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_anyGoals___closed__5; @@ -474,9 +479,8 @@ static lean_object* l_Lean_Parser_Tactic_Conv_convConvSeq___closed__5; static lean_object* l_Lean_Parser_Tactic_Conv_enterArg___closed__4; static lean_object* l_Lean_Parser_Tactic_Conv_change___closed__8; static lean_object* l_Lean_Parser_Tactic_Conv_conv_quot___closed__5; -static lean_object* l_Lean_Parser_Tactic_Conv_dsimp___closed__14; static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convRfl__1___closed__3; -static lean_object* l_Lean_Parser_Tactic_Conv_delta___closed__7; +static lean_object* l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__1; 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_occsIndexed___closed__2; static lean_object* l_Lean_Parser_Tactic_Conv_convSeqBracketed___closed__9; @@ -489,10 +493,10 @@ static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules_ lean_object* l_Lean_Syntax_node1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv_conv___closed__3; lean_object* l_Lean_mkSepArray(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__6; +static lean_object* l_Lean_Parser_Tactic_Conv_convIntro_________closed__4; static lean_object* l_Lean_Parser_Tactic_Conv_conv_xb7_x2e_____closed__11; static lean_object* l_Lean_Parser_Tactic_Conv_dsimp___closed__7; -static lean_object* l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__3; +static lean_object* l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__9; extern lean_object* l_Lean_binderIdent; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_nestedConv; static lean_object* l_Lean_Parser_Tactic_Conv_simp___closed__18; @@ -506,6 +510,7 @@ static lean_object* l_Lean_Parser_Tactic_Conv_rewrite___closed__5; static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__allGoals__1___closed__2; static lean_object* l_Lean_Parser_Tactic_Conv_conv_xb7_x2e_____closed__17; static lean_object* l_Lean_Parser_Tactic_Conv_first___closed__12; +static lean_object* l_Lean_Parser_Tactic_Conv_first___closed__24; static lean_object* l_Lean_Parser_Tactic_Conv_case___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_convRepeat__; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_convErw__; @@ -519,14 +524,12 @@ static lean_object* l_Lean_Parser_Tactic_Conv_change___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_congr; static lean_object* l_Lean_Parser_Tactic_Conv_allGoals___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_rhs; -static lean_object* l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__4; -static lean_object* l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__5; static lean_object* l_Lean_Parser_Tactic_Conv_allGoals___closed__6; +static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext_____x3d_x3e____1___closed__6; extern lean_object* l_Lean_Parser_Tactic_caseArg; static lean_object* l_Lean_Parser_Tactic_Conv_case___closed__11; static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__conv___x3c_x3b_x3e____1___closed__1; static lean_object* l_Lean_Parser_Tactic_Conv_skip___closed__4; -static lean_object* l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_simp; static lean_object* l_Lean_Parser_Tactic_Conv_paren___closed__6; static lean_object* l_Lean_Parser_Tactic_Conv_simp___closed__9; @@ -538,12 +541,12 @@ static lean_object* l_Lean_Parser_Tactic_Conv_paren___closed__2; static lean_object* l_Lean_Parser_Tactic_Conv_nestedTacticCore___closed__8; static lean_object* l_Lean_Parser_Tactic_Conv_convSeq___closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_conv_xb7_x2e_____closed__10; +static lean_object* l_Lean_Parser_Tactic_Conv_convIntro_________closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_convLeft___closed__2; static lean_object* l_Lean_Parser_Tactic_Conv_convApply_____closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_conv___closed__9; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_focus; static lean_object* l_Lean_Parser_Tactic_Conv_conv_xb7_x2e_____closed__5; -static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext___x3d_x3e____1___closed__6; static lean_object* l_Lean_Parser_Tactic_Conv_ext___closed__10; LEAN_EXPORT lean_object* l_Lean_Parser_Category_conv; static lean_object* l_Lean_Parser_Tactic_Conv_dsimp___closed__1; @@ -558,23 +561,20 @@ static lean_object* l_Lean_Parser_Tactic_Conv_zeta___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_convTry__; static lean_object* l_Lean_Parser_Tactic_Conv_ext___closed__4; static lean_object* l_Lean_Parser_Tactic_Conv_convErw_____closed__2; -static lean_object* l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__9; static lean_object* l_Lean_Parser_Tactic_Conv_convArgs___closed__4; static lean_object* l_Lean_Parser_Tactic_Conv_skip___closed__1; static lean_object* l_Lean_Parser_Tactic_Conv_failIfSuccess___closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_convErw_____closed__4; -static lean_object* l_Lean_Parser_Tactic_Conv_dsimp___closed__16; static lean_object* l_Lean_Parser_Tactic_Conv_change___closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_first___closed__13; static lean_object* l_Lean_Parser_Tactic_Conv_simp___closed__4; static lean_object* l_Lean_Parser_Tactic_Conv_conv___closed__14; -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convEnter_x5b_____x5d__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Parser_Tactic_config; static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convDone__1___closed__1; +static lean_object* l_Lean_Parser_Tactic_Conv_first___closed__19; static lean_object* l_Lean_Parser_Tactic_Conv_anyGoals___closed__2; 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_convRfl___closed__1; -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d; lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv_pattern___closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_nestedTacticCore___closed__3; @@ -585,7 +585,6 @@ static lean_object* l_Lean_Parser_Tactic_Conv_focus___closed__4; static lean_object* l_Lean_Parser_Tactic_Conv_paren___closed__1; static lean_object* l_Lean_Parser_Tactic_Conv_conv_quot___closed__19; static lean_object* l_Lean_Parser_Tactic_Conv_convRw_______closed__2; -static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext___x3d_x3e____1___closed__5; static lean_object* l_Lean_Parser_Tactic_Conv_simp___closed__24; static lean_object* l_Lean_Parser_Tactic_Conv_occsWildcard___closed__4; static lean_object* l_Lean_Parser_Tactic_Conv_case_x27___closed__8; @@ -593,23 +592,22 @@ static lean_object* l_Lean_Parser_Tactic_Conv_simp___closed__22; static lean_object* l_Lean_Parser_Tactic_Conv_simp___closed__12; static lean_object* l_Lean_Parser_Tactic_Conv_occsWildcard___closed__2; static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convErw____1___closed__1; -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_convEnter_x5b_____x5d___closed__2; static lean_object* l_Lean_Parser_Tactic_Conv_first___closed__15; -static lean_object* l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__7; -static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext___x3d_x3e____1___closed__7; static lean_object* l_Lean_Parser_Tactic_Conv_convRw_______closed__4; static lean_object* l_Lean_Parser_Tactic_Conv_convTry_____closed__3; +static lean_object* l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__10; static lean_object* l_Lean_Parser_Tactic_Conv_convDone___closed__1; -static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext___x3d_x3e____1___closed__2; static lean_object* l_Lean_Parser_Tactic_Conv_withAnnotateState___closed__5; static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convRfl__1___closed__8; static lean_object* l_Lean_Parser_Tactic_Conv_convTrace__state___closed__5; static lean_object* l_Lean_Parser_Tactic_Conv_conv___closed__5; +static lean_object* l_Lean_Parser_Tactic_Conv_first___closed__18; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convLeft__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv_occs___closed__8; uint8_t lean_nat_dec_le(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic_Conv_convIntro_________closed__6; static lean_object* l_Lean_Parser_Tactic_Conv_pattern___closed__5; +static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convEnter_x5b___x5d__1___closed__1; static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convErw____1___closed__4; static lean_object* l_Lean_Parser_Tactic_Conv_conv___closed__10; static lean_object* l_Lean_Parser_Tactic_Conv_case___closed__3; @@ -621,24 +619,28 @@ static lean_object* l_Lean_Parser_Tactic_Conv_nestedConv___closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_conv_quot___closed__9; static lean_object* l_Lean_Parser_Tactic_Conv_arg___closed__7; 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_convNext_____x3d_x3e_____closed__8; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convArgs__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__21; static lean_object* l_Lean_Parser_Tactic_Conv_conv___closed__4; +static lean_object* l_Lean_Parser_Tactic_Conv_first___closed__21; static lean_object* l_Lean_Parser_Tactic_Conv_simp___closed__14; -static lean_object* l_Lean_Parser_Tactic_Conv_dsimp___closed__18; +static lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convEnter_x5b___x5d__1___lambda__2___closed__1; static lean_object* l_Lean_Parser_Tactic_Conv_ext___closed__5; static lean_object* l_Lean_Parser_Tactic_Conv_convSeqBracketed___closed__5; static lean_object* l_Lean_Parser_Tactic_Conv_unfold___closed__2; static lean_object* l_Lean_Parser_Tactic_Conv_occs___closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_rhs___closed__1; +static lean_object* l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__4; lean_object* l_String_toSubstring_x27(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_first; static lean_object* l_Lean_Parser_Tactic_Conv_convConvSeq___closed__3; static lean_object* l_Lean_Parser_Tactic_Conv_convRw_______closed__6; static lean_object* l_Lean_Parser_Tactic_Conv_occsIndexed___closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv_anyGoals; +static lean_object* l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__6; +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convEnter_x5b___x5d__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_Conv_convLeft___closed__1; -static lean_object* l_Lean_Parser_Tactic_Conv_convIntro_______closed__5; static lean_object* l_Lean_Parser_Tactic_Conv_case___closed__7; 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*); LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convTry____1(lean_object*, lean_object*, lean_object*); @@ -2347,38 +2349,52 @@ return x_2; static lean_object* _init_l_Lean_Parser_Tactic_Conv_ext___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__9; +x_2 = l_Lean_Parser_Tactic_Conv_withAnnotateState___closed__11; +x_3 = l_Lean_Parser_Tactic_Conv_ext___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); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_Conv_ext___closed__10() { +_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_Conv_ext___closed__10() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_ext___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_Conv_ext___closed__9; +x_2 = l_Lean_Parser_Tactic_Conv_ext___closed__10; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_ext___closed__11() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_ext___closed__12() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic_Conv_ext___closed__10; +x_1 = l_Lean_Parser_Tactic_Conv_ext___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_Conv_ext___closed__12() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_ext___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_Conv_conv_quot___closed__9; -x_2 = l_Lean_Parser_Tactic_Conv_ext___closed__8; -x_3 = l_Lean_Parser_Tactic_Conv_ext___closed__11; +x_2 = l_Lean_Parser_Tactic_Conv_ext___closed__9; +x_3 = l_Lean_Parser_Tactic_Conv_ext___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); @@ -2386,25 +2402,25 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_ext___closed__13() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_ext___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Tactic_Conv_ext___closed__5; -x_2 = l_Lean_Parser_Tactic_Conv_ext___closed__12; +x_2 = l_Lean_Parser_Tactic_Conv_ext___closed__13; 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_Conv_ext___closed__14() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_ext___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_Conv_conv_quot___closed__9; x_2 = l_Lean_Parser_Tactic_Conv_ext___closed__3; -x_3 = l_Lean_Parser_Tactic_Conv_ext___closed__13; +x_3 = l_Lean_Parser_Tactic_Conv_ext___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); @@ -2412,13 +2428,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_ext___closed__15() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_ext___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_Conv_ext___closed__2; x_2 = lean_unsigned_to_nat(1022u); -x_3 = l_Lean_Parser_Tactic_Conv_ext___closed__14; +x_3 = l_Lean_Parser_Tactic_Conv_ext___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); @@ -2430,7 +2446,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_ext() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Tactic_Conv_ext___closed__15; +x_1 = l_Lean_Parser_Tactic_Conv_ext___closed__16; return x_1; } } @@ -2565,16 +2581,8 @@ return x_6; static lean_object* _init_l_Lean_Parser_Tactic_Conv_delta___closed__3() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("delta ", 6); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Tactic_Conv_delta___closed__4() { -_start: -{ lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_Conv_delta___closed__3; +x_1 = l_Lean_Parser_Tactic_Conv_delta___closed__1; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -2582,25 +2590,25 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_delta___closed__5() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_delta___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Tactic_Conv_occsIndexed___closed__4; -x_2 = l_Lean_Parser_Tactic_Conv_ext___closed__12; +x_2 = l_Lean_Parser_Tactic_Conv_ext___closed__13; 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_Conv_delta___closed__6() { +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__9; -x_2 = l_Lean_Parser_Tactic_Conv_delta___closed__4; -x_3 = l_Lean_Parser_Tactic_Conv_delta___closed__5; +x_2 = l_Lean_Parser_Tactic_Conv_delta___closed__3; +x_3 = l_Lean_Parser_Tactic_Conv_delta___closed__4; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -2608,13 +2616,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_delta___closed__7() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_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_Conv_delta___closed__2; x_2 = lean_unsigned_to_nat(1022u); -x_3 = l_Lean_Parser_Tactic_Conv_delta___closed__6; +x_3 = l_Lean_Parser_Tactic_Conv_delta___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); @@ -2626,7 +2634,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_delta() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Tactic_Conv_delta___closed__7; +x_1 = l_Lean_Parser_Tactic_Conv_delta___closed__6; return x_1; } } @@ -2654,16 +2662,8 @@ return x_6; static lean_object* _init_l_Lean_Parser_Tactic_Conv_unfold___closed__3() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("unfold ", 7); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Tactic_Conv_unfold___closed__4() { -_start: -{ lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_Conv_unfold___closed__3; +x_1 = l_Lean_Parser_Tactic_Conv_unfold___closed__1; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -2671,13 +2671,13 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_unfold___closed__5() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_unfold___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__9; -x_2 = l_Lean_Parser_Tactic_Conv_unfold___closed__4; -x_3 = l_Lean_Parser_Tactic_Conv_delta___closed__5; +x_2 = l_Lean_Parser_Tactic_Conv_unfold___closed__3; +x_3 = l_Lean_Parser_Tactic_Conv_delta___closed__4; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -2685,13 +2685,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_unfold___closed__6() { +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_unfold___closed__2; x_2 = lean_unsigned_to_nat(1022u); -x_3 = l_Lean_Parser_Tactic_Conv_unfold___closed__5; +x_3 = l_Lean_Parser_Tactic_Conv_unfold___closed__4; x_4 = lean_alloc_ctor(3, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -2703,7 +2703,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_unfold() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Tactic_Conv_unfold___closed__6; +x_1 = l_Lean_Parser_Tactic_Conv_unfold___closed__5; return x_1; } } @@ -3242,16 +3242,8 @@ return x_6; static lean_object* _init_l_Lean_Parser_Tactic_Conv_dsimp___closed__3() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("dsimp ", 6); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Tactic_Conv_dsimp___closed__4() { -_start: -{ lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_Conv_dsimp___closed__3; +x_1 = l_Lean_Parser_Tactic_Conv_dsimp___closed__1; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3259,12 +3251,12 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_dsimp___closed__5() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_dsimp___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__9; -x_2 = l_Lean_Parser_Tactic_Conv_dsimp___closed__4; +x_2 = l_Lean_Parser_Tactic_Conv_dsimp___closed__3; x_3 = l_Lean_Parser_Tactic_Conv_rewrite___closed__4; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -3273,12 +3265,12 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_dsimp___closed__6() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_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_Conv_conv_quot___closed__9; -x_2 = l_Lean_Parser_Tactic_Conv_dsimp___closed__5; +x_2 = l_Lean_Parser_Tactic_Conv_dsimp___closed__4; x_3 = l_Lean_Parser_Tactic_Conv_simp___closed__5; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -3287,45 +3279,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_dsimp___closed__7() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("only ", 5); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Tactic_Conv_dsimp___closed__8() { -_start: -{ -lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_Conv_dsimp___closed__7; -x_2 = 0; -x_3 = lean_alloc_ctor(6, 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_Conv_dsimp___closed__9() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_Conv_arg___closed__6; -x_2 = l_Lean_Parser_Tactic_Conv_dsimp___closed__8; -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_Conv_dsimp___closed__10() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_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_Conv_conv_quot___closed__9; -x_2 = l_Lean_Parser_Tactic_Conv_dsimp___closed__6; -x_3 = l_Lean_Parser_Tactic_Conv_dsimp___closed__9; +x_2 = l_Lean_Parser_Tactic_Conv_dsimp___closed__5; +x_3 = l_Lean_Parser_Tactic_Conv_simp___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); @@ -3333,25 +3293,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_dsimp___closed__11() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("[", 1); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Tactic_Conv_dsimp___closed__12() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic_Conv_dsimp___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_dsimp___closed__13() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_dsimp___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; uint8_t x_4; lean_object* x_5; @@ -3367,25 +3309,25 @@ lean_ctor_set_uint8(x_5, sizeof(void*)*3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_dsimp___closed__14() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_dsimp___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Tactic_Conv_convSeqBracketed___closed__6; -x_2 = l_Lean_Parser_Tactic_Conv_dsimp___closed__13; +x_2 = l_Lean_Parser_Tactic_Conv_dsimp___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_Conv_dsimp___closed__15() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_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_Conv_conv_quot___closed__9; -x_2 = l_Lean_Parser_Tactic_Conv_dsimp___closed__12; -x_3 = l_Lean_Parser_Tactic_Conv_dsimp___closed__14; +x_2 = l_Lean_Parser_Tactic_Conv_simp___closed__12; +x_3 = l_Lean_Parser_Tactic_Conv_dsimp___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); @@ -3393,12 +3335,12 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_dsimp___closed__16() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_dsimp___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__9; -x_2 = l_Lean_Parser_Tactic_Conv_dsimp___closed__15; +x_2 = l_Lean_Parser_Tactic_Conv_dsimp___closed__9; x_3 = l_Lean_Parser_Tactic_Conv_simp___closed__22; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -3407,25 +3349,25 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_dsimp___closed__17() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_dsimp___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Tactic_Conv_arg___closed__6; -x_2 = l_Lean_Parser_Tactic_Conv_dsimp___closed__16; +x_2 = l_Lean_Parser_Tactic_Conv_dsimp___closed__10; 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_Conv_dsimp___closed__18() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_dsimp___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__9; -x_2 = l_Lean_Parser_Tactic_Conv_dsimp___closed__10; -x_3 = l_Lean_Parser_Tactic_Conv_dsimp___closed__17; +x_2 = l_Lean_Parser_Tactic_Conv_dsimp___closed__6; +x_3 = l_Lean_Parser_Tactic_Conv_dsimp___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); @@ -3433,13 +3375,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_dsimp___closed__19() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_dsimp___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_Conv_dsimp___closed__2; x_2 = lean_unsigned_to_nat(1022u); -x_3 = l_Lean_Parser_Tactic_Conv_dsimp___closed__18; +x_3 = l_Lean_Parser_Tactic_Conv_dsimp___closed__12; x_4 = lean_alloc_ctor(3, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -3451,7 +3393,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_dsimp() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Tactic_Conv_dsimp___closed__19; +x_1 = l_Lean_Parser_Tactic_Conv_dsimp___closed__13; return x_1; } } @@ -5377,15 +5319,15 @@ return x_53; } } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__1() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("convNext_=>_", 12); +x_1 = lean_mk_string_from_bytes("convNext__=>_", 13); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__2() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__2() { _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; @@ -5393,24 +5335,24 @@ x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__1; x_2 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__2; x_3 = l_Lean_Parser_Tactic_Conv_convSeq1Indented___closed__1; x_4 = l_Lean_Parser_Tactic_Conv_convSeq1Indented___closed__2; -x_5 = l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__1; +x_5 = l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__1; x_6 = l_Lean_Name_mkStr5(x_1, x_2, x_3, x_4, x_5); return x_6; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__3() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__3() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("next ", 5); +x_1 = lean_mk_string_from_bytes("next", 4); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__4() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__4() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__3; +x_1 = l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__3; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -5418,25 +5360,39 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__5() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_convNext_____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_Conv_conv_quot___closed__9; +x_2 = l_Lean_Parser_Tactic_Conv_withAnnotateState___closed__11; +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); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Tactic_Conv_ext___closed__5; -x_2 = l_Lean_binderIdent; +x_2 = l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____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_Conv_convNext___x3d_x3e_____closed__6() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_convNext_____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_Conv_conv_quot___closed__9; -x_2 = l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__4; -x_3 = l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__5; +x_2 = l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__4; +x_3 = l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____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); @@ -5444,12 +5400,12 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__7() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_convNext_____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_Conv_conv_quot___closed__9; -x_2 = l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__6; +x_2 = l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__7; x_3 = l_Lean_Parser_Tactic_Conv_nestedTacticCore___closed__6; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -5458,12 +5414,12 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__8() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____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__9; -x_2 = l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__7; +x_2 = l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__8; x_3 = l_Lean_Parser_Tactic_Conv_convSeq; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -5472,13 +5428,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__9() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____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_convNext___x3d_x3e_____closed__2; +x_1 = l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__2; x_2 = lean_unsigned_to_nat(1022u); -x_3 = l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__8; +x_3 = l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____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); @@ -5486,15 +5442,15 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e__() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e__() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__9; +x_1 = l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__10; return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext___x3d_x3e____1___closed__1() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext_____x3d_x3e____1___closed__1() { _start: { lean_object* x_1; @@ -5502,19 +5458,19 @@ x_1 = lean_mk_string_from_bytes("caseArg", 7); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext___x3d_x3e____1___closed__2() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext_____x3d_x3e____1___closed__2() { _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_Tactic_Conv_conv_quot___closed__1; x_2 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__2; x_3 = l_Lean_Parser_Tactic_Conv_convSeq1Indented___closed__1; -x_4 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext___x3d_x3e____1___closed__1; +x_4 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext_____x3d_x3e____1___closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext___x3d_x3e____1___closed__3() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext_____x3d_x3e____1___closed__3() { _start: { lean_object* x_1; @@ -5522,17 +5478,17 @@ x_1 = lean_mk_string_from_bytes("binderIdent", 11); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext___x3d_x3e____1___closed__4() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext_____x3d_x3e____1___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__1; -x_2 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext___x3d_x3e____1___closed__3; +x_2 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext_____x3d_x3e____1___closed__3; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext___x3d_x3e____1___closed__5() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext_____x3d_x3e____1___closed__5() { _start: { lean_object* x_1; @@ -5540,19 +5496,19 @@ x_1 = lean_mk_string_from_bytes("hole", 4); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext___x3d_x3e____1___closed__6() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext_____x3d_x3e____1___closed__6() { _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_Tactic_Conv_conv_quot___closed__1; x_2 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__2; x_3 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__3; -x_4 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext___x3d_x3e____1___closed__5; +x_4 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext_____x3d_x3e____1___closed__5; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext___x3d_x3e____1___closed__7() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext_____x3d_x3e____1___closed__7() { _start: { lean_object* x_1; @@ -5560,11 +5516,11 @@ x_1 = lean_mk_string_from_bytes("_", 1); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext___x3d_x3e____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__convNext_____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_Conv_convNext___x3d_x3e_____closed__2; +x_4 = l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__2; lean_inc(x_1); x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); if (x_5 == 0) @@ -5598,15 +5554,15 @@ 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_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext___x3d_x3e____1___closed__7; +x_18 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext_____x3d_x3e____1___closed__7; 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_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext___x3d_x3e____1___closed__6; +x_20 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext_____x3d_x3e____1___closed__6; lean_inc(x_15); x_21 = l_Lean_Syntax_node1(x_15, x_20, x_19); -x_22 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext___x3d_x3e____1___closed__4; +x_22 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext_____x3d_x3e____1___closed__4; lean_inc(x_15); x_23 = l_Lean_Syntax_node1(x_15, x_22, x_21); x_24 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__case__1___closed__2; @@ -5617,7 +5573,7 @@ x_27 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_27, 0, x_15); lean_ctor_set(x_27, 1, x_26); lean_ctor_set(x_27, 2, x_25); -x_28 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext___x3d_x3e____1___closed__2; +x_28 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext_____x3d_x3e____1___closed__2; lean_inc(x_15); x_29 = l_Lean_Syntax_node2(x_15, x_28, x_23, x_27); lean_inc(x_15); @@ -7287,15 +7243,15 @@ return x_15; } } } -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__", 11); +x_1 = lean_mk_string_from_bytes("convIntro___", 12); 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; lean_object* x_4; lean_object* x_5; lean_object* x_6; @@ -7303,12 +7259,12 @@ x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__1; x_2 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__2; x_3 = l_Lean_Parser_Tactic_Conv_convSeq1Indented___closed__1; x_4 = l_Lean_Parser_Tactic_Conv_convSeq1Indented___closed__2; -x_5 = l_Lean_Parser_Tactic_Conv_convIntro_______closed__1; +x_5 = l_Lean_Parser_Tactic_Conv_convIntro_________closed__1; x_6 = l_Lean_Name_mkStr5(x_1, x_2, x_3, x_4, x_5); return x_6; } } -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; @@ -7316,11 +7272,11 @@ x_1 = lean_mk_string_from_bytes("intro", 5); 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); @@ -7328,13 +7284,13 @@ 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__9; -x_2 = l_Lean_Parser_Tactic_Conv_convIntro_______closed__4; -x_3 = l_Lean_Parser_Tactic_Conv_ext___closed__13; +x_2 = l_Lean_Parser_Tactic_Conv_convIntro_________closed__4; +x_3 = l_Lean_Parser_Tactic_Conv_ext___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); @@ -7342,13 +7298,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); @@ -7356,19 +7312,19 @@ 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; +x_1 = l_Lean_Parser_Tactic_Conv_convIntro_________closed__6; 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) @@ -7487,7 +7443,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_enterArg___closed__7() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Tactic_Conv_convSeq___closed__4; -x_2 = l_Lean_Parser_Tactic_Conv_ext___closed__11; +x_2 = l_Lean_Parser_Tactic_Conv_ext___closed__12; x_3 = l_Lean_Parser_Tactic_Conv_enterArg___closed__6; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -7518,15 +7474,15 @@ x_1 = l_Lean_Parser_Tactic_Conv_enterArg___closed__8; return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__1() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("convEnter[__]", 13); +x_1 = lean_mk_string_from_bytes("convEnter[_]", 12); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__2() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__2() { _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; @@ -7534,12 +7490,12 @@ x_1 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__1; x_2 = l_Lean_Parser_Tactic_Conv_conv_quot___closed__2; x_3 = l_Lean_Parser_Tactic_Conv_convSeq1Indented___closed__1; x_4 = l_Lean_Parser_Tactic_Conv_convSeq1Indented___closed__2; -x_5 = l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__1; +x_5 = l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__1; x_6 = l_Lean_Name_mkStr5(x_1, x_2, x_3, x_4, x_5); return x_6; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__3() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__3() { _start: { lean_object* x_1; @@ -7547,11 +7503,11 @@ x_1 = lean_mk_string_from_bytes("enter", 5); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__4() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__4() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__3; +x_1 = l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__3; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -7559,12 +7515,12 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__5() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___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__9; -x_2 = l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__4; +x_2 = l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__4; x_3 = l_Lean_Parser_Tactic_Conv_simp___closed__12; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -7573,25 +7529,11 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___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__9; -x_2 = l_Lean_Parser_Tactic_Conv_ext___closed__8; -x_3 = l_Lean_Parser_Tactic_Conv_enterArg; -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_Conv_convEnter_x5b_____x5d___closed__7() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__6() { _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_Conv_convEnter_x5b_____x5d___closed__6; +x_1 = l_Lean_Parser_Tactic_Conv_enterArg; x_2 = l_Lean_Parser_Tactic_Conv_simp___closed__17; x_3 = l_Lean_Parser_Tactic_Conv_simp___closed__16; x_4 = 0; @@ -7603,13 +7545,25 @@ lean_ctor_set_uint8(x_5, sizeof(void*)*3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__8() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Tactic_Conv_convSeqBracketed___closed__6; +x_2 = l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___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_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___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__9; -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_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); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -7617,12 +7571,12 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__9() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___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__9; -x_2 = l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__8; +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); lean_ctor_set(x_4, 0, x_1); @@ -7631,13 +7585,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__10() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___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_convEnter_x5b_____x5d___closed__2; +x_1 = l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__9; +x_3 = l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___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); @@ -7645,15 +7599,15 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__10; +x_1 = l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__10; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convEnter_x5b_____x5d__1___lambda__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__convEnter_x5b___x5d__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; @@ -7664,7 +7618,15 @@ lean_ctor_set(x_5, 1, x_3); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convEnter_x5b_____x5d__1___lambda__2(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_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convEnter_x5b___x5d__1___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("[", 1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convEnter_x5b___x5d__1___lambda__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; 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; @@ -7678,12 +7640,12 @@ lean_inc(x_8); x_10 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_10, 0, x_8); lean_ctor_set(x_10, 1, x_9); -x_11 = l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__3; +x_11 = l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__3; lean_inc(x_8); x_12 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_12, 0, x_8); lean_ctor_set(x_12, 1, x_11); -x_13 = l_Lean_Parser_Tactic_Conv_dsimp___closed__11; +x_13 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convEnter_x5b___x5d__1___lambda__2___closed__1; lean_inc(x_8); x_14 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_14, 0, x_8); @@ -7737,19 +7699,19 @@ lean_ctor_set(x_35, 1, x_5); return x_35; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convEnter_x5b_____x5d__1___closed__1() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convEnter_x5b___x5d__1___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convEnter_x5b_____x5d__1___lambda__1___boxed), 3, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convEnter_x5b___x5d__1___lambda__1___boxed), 3, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convEnter_x5b_____x5d__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__convEnter_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_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__2; +x_4 = l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__2; lean_inc(x_1); x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); if (x_5 == 0) @@ -7769,7 +7731,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_ 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_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convEnter_x5b_____x5d__1___closed__1; +x_10 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convEnter_x5b___x5d__1___closed__1; x_11 = lean_unsigned_to_nat(1u); lean_inc(x_9); x_12 = l_Lean_Syntax_matchesNull(x_9, x_11); @@ -7807,7 +7769,7 @@ lean_ctor_set(x_24, 1, x_23); lean_ctor_set(x_24, 2, x_21); x_25 = l_Lean_Syntax_getArgs(x_24); lean_dec(x_24); -x_26 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convEnter_x5b_____x5d__1___lambda__2(x_4, x_18, x_25, x_2, x_3); +x_26 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convEnter_x5b___x5d__1___lambda__2(x_4, x_18, x_25, x_2, x_3); return x_26; } } @@ -7852,7 +7814,7 @@ lean_ctor_set(x_40, 1, x_39); lean_ctor_set(x_40, 2, x_37); x_41 = l_Lean_Syntax_getArgs(x_40); lean_dec(x_40); -x_42 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convEnter_x5b_____x5d__1___lambda__2(x_4, x_28, x_41, x_2, x_3); +x_42 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convEnter_x5b___x5d__1___lambda__2(x_4, x_28, x_41, x_2, x_3); return x_42; } } @@ -7866,7 +7828,7 @@ x_45 = l_Lean_Syntax_isOfKind(x_43, x_44); if (x_45 == 0) { lean_object* x_46; uint8_t x_47; -x_46 = l_Lean_Parser_Tactic_Conv_ext___closed__10; +x_46 = l_Lean_Parser_Tactic_Conv_ext___closed__11; lean_inc(x_43); x_47 = l_Lean_Syntax_isOfKind(x_43, x_46); if (x_47 == 0) @@ -7903,7 +7865,7 @@ lean_ctor_set(x_57, 1, x_56); lean_ctor_set(x_57, 2, x_54); x_58 = l_Lean_Syntax_getArgs(x_57); lean_dec(x_57); -x_59 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convEnter_x5b_____x5d__1___lambda__2(x_4, x_28, x_58, x_2, x_3); +x_59 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convEnter_x5b___x5d__1___lambda__2(x_4, x_28, x_58, x_2, x_3); return x_59; } } @@ -7977,7 +7939,7 @@ lean_ctor_set(x_82, 1, x_81); lean_ctor_set(x_82, 2, x_79); x_83 = l_Lean_Syntax_getArgs(x_82); lean_dec(x_82); -x_84 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convEnter_x5b_____x5d__1___lambda__2(x_4, x_28, x_83, x_2, x_3); +x_84 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convEnter_x5b___x5d__1___lambda__2(x_4, x_28, x_83, x_2, x_3); return x_84; } } @@ -8057,7 +8019,7 @@ lean_ctor_set(x_110, 1, x_109); lean_ctor_set(x_110, 2, x_107); x_111 = l_Lean_Syntax_getArgs(x_110); lean_dec(x_110); -x_112 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convEnter_x5b_____x5d__1___lambda__2(x_4, x_28, x_111, x_2, x_3); +x_112 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convEnter_x5b___x5d__1___lambda__2(x_4, x_28, x_111, x_2, x_3); return x_112; } } @@ -8097,11 +8059,11 @@ return x_123; } } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convEnter_x5b_____x5d__1___lambda__1___boxed(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__convEnter_x5b___x5d__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convEnter_x5b_____x5d__1___lambda__1(x_1, x_2, x_3); +x_4 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convEnter_x5b___x5d__1___lambda__1(x_1, x_2, x_3); lean_dec(x_2); lean_dec(x_1); return x_4; @@ -8332,7 +8294,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_first___closed__7() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("colGe", 5); +x_1 = lean_mk_string_from_bytes("ppDedent", 8); return x_1; } } @@ -8349,30 +8311,110 @@ return x_3; static lean_object* _init_l_Lean_Parser_Tactic_Conv_first___closed__9() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("ppLine", 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_Conv_first___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_Conv_first___closed__9; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_Conv_first___closed__11() { +_start: +{ lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Tactic_Conv_first___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_Parser_Tactic_Conv_first___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Tactic_Conv_first___closed__8; +x_2 = l_Lean_Parser_Tactic_Conv_first___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_Parser_Tactic_Conv_first___closed__13() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("colGe", 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_Conv_first___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_Conv_first___closed__13; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_Conv_first___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Tactic_Conv_first___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_Conv_first___closed__10() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_first___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_Conv_conv_quot___closed__9; +x_2 = l_Lean_Parser_Tactic_Conv_first___closed__12; +x_3 = l_Lean_Parser_Tactic_Conv_first___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); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_Conv_first___closed__17() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("| ", 2); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_Conv_first___closed__18() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__case__1___closed__3; +x_1 = l_Lean_Parser_Tactic_Conv_first___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_first___closed__11() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_first___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__9; -x_2 = l_Lean_Parser_Tactic_Conv_first___closed__9; -x_3 = l_Lean_Parser_Tactic_Conv_first___closed__10; +x_2 = l_Lean_Parser_Tactic_Conv_first___closed__16; +x_3 = l_Lean_Parser_Tactic_Conv_first___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); @@ -8380,12 +8422,12 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_first___closed__12() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_first___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__9; -x_2 = l_Lean_Parser_Tactic_Conv_first___closed__11; +x_2 = l_Lean_Parser_Tactic_Conv_first___closed__19; x_3 = l_Lean_Parser_Tactic_Conv_convSeq; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -8394,49 +8436,49 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_first___closed__13() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_first___closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Tactic_Conv_enterArg___closed__4; -x_2 = l_Lean_Parser_Tactic_Conv_first___closed__12; +x_2 = l_Lean_Parser_Tactic_Conv_first___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_Parser_Tactic_Conv_first___closed__14() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_first___closed__22() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Tactic_Conv_occsIndexed___closed__4; -x_2 = l_Lean_Parser_Tactic_Conv_first___closed__13; +x_2 = l_Lean_Parser_Tactic_Conv_first___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); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_first___closed__15() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_first___closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Tactic_Conv_first___closed__6; -x_2 = l_Lean_Parser_Tactic_Conv_first___closed__14; +x_2 = l_Lean_Parser_Tactic_Conv_first___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_Parser_Tactic_Conv_first___closed__16() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_first___closed__24() { _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__9; x_2 = l_Lean_Parser_Tactic_Conv_first___closed__4; -x_3 = l_Lean_Parser_Tactic_Conv_first___closed__15; +x_3 = l_Lean_Parser_Tactic_Conv_first___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); @@ -8444,13 +8486,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_first___closed__17() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_first___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_first___closed__2; x_2 = lean_unsigned_to_nat(1022u); -x_3 = l_Lean_Parser_Tactic_Conv_first___closed__16; +x_3 = l_Lean_Parser_Tactic_Conv_first___closed__24; x_4 = lean_alloc_ctor(3, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -8462,7 +8504,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_first() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Tactic_Conv_first___closed__17; +x_1 = l_Lean_Parser_Tactic_Conv_first___closed__25; return x_1; } } @@ -8884,7 +8926,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_convRepeat_____closed__3() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("repeat", 6); +x_1 = lean_mk_string_from_bytes("repeat ", 7); return x_1; } } @@ -8936,6 +8978,14 @@ x_1 = l_Lean_Parser_Tactic_Conv_convRepeat_____closed__6; return x_1; } } +static lean_object* _init_l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convRepeat____1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("repeat", 6); +return x_1; +} +} LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convRepeat____1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { @@ -8994,7 +9044,7 @@ 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_Lean_Parser_Tactic_Conv_convRepeat_____closed__3; +x_25 = l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convRepeat____1___closed__1; lean_inc(x_12); x_26 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_26, 0, x_12); @@ -9058,47 +9108,27 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv___closed__2() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("conv ", 5); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv___closed__3() { -_start: -{ -lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_Conv_conv___closed__2; -x_2 = 0; -x_3 = lean_alloc_ctor(6, 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_Conv_conv___closed__4() { -_start: -{ -lean_object* x_1; x_1 = lean_mk_string_from_bytes(" at ", 4); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv___closed__5() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv___closed__3() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic_Conv_conv___closed__4; +x_1 = l_Lean_Parser_Tactic_Conv_conv___closed__2; 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___closed__6() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv___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__9; -x_2 = l_Lean_Parser_Tactic_Conv_conv___closed__5; -x_3 = l_Lean_Parser_Tactic_Conv_ext___closed__11; +x_2 = l_Lean_Parser_Tactic_Conv_conv___closed__3; +x_3 = l_Lean_Parser_Tactic_Conv_ext___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); @@ -9106,25 +9136,25 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv___closed__7() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Tactic_Conv_arg___closed__6; -x_2 = l_Lean_Parser_Tactic_Conv_conv___closed__6; +x_2 = l_Lean_Parser_Tactic_Conv_conv___closed__4; 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_Conv_conv___closed__8() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv___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__9; -x_2 = l_Lean_Parser_Tactic_Conv_conv___closed__3; -x_3 = l_Lean_Parser_Tactic_Conv_conv___closed__7; +x_2 = l_Lean_Parser_Tactic_Conv_convConvSeq___closed__3; +x_3 = l_Lean_Parser_Tactic_Conv_conv___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); @@ -9132,7 +9162,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv___closed__9() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv___closed__7() { _start: { lean_object* x_1; @@ -9140,22 +9170,22 @@ x_1 = lean_mk_string_from_bytes(" in ", 4); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv___closed__10() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv___closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic_Conv_conv___closed__9; +x_1 = l_Lean_Parser_Tactic_Conv_conv___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_Parser_Tactic_Conv_conv___closed__11() { +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__9; -x_2 = l_Lean_Parser_Tactic_Conv_conv___closed__10; +x_2 = l_Lean_Parser_Tactic_Conv_conv___closed__8; x_3 = l_Lean_Parser_Tactic_Conv_pattern___closed__5; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -9164,12 +9194,12 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv___closed__12() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv___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__9; -x_2 = l_Lean_Parser_Tactic_Conv_conv___closed__11; +x_2 = l_Lean_Parser_Tactic_Conv_conv___closed__9; x_3 = l_Lean_Parser_Tactic_Conv_change___closed__7; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -9178,25 +9208,25 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv___closed__13() { +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; x_1 = l_Lean_Parser_Tactic_Conv_arg___closed__6; -x_2 = l_Lean_Parser_Tactic_Conv_conv___closed__12; +x_2 = l_Lean_Parser_Tactic_Conv_conv___closed__10; 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_Conv_conv___closed__14() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv___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__9; -x_2 = l_Lean_Parser_Tactic_Conv_conv___closed__8; -x_3 = l_Lean_Parser_Tactic_Conv_conv___closed__13; +x_2 = l_Lean_Parser_Tactic_Conv_conv___closed__6; +x_3 = l_Lean_Parser_Tactic_Conv_conv___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); @@ -9204,12 +9234,12 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv___closed__15() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv___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_Conv_conv_quot___closed__9; -x_2 = l_Lean_Parser_Tactic_Conv_conv___closed__14; +x_2 = l_Lean_Parser_Tactic_Conv_conv___closed__12; x_3 = l_Lean_Parser_Tactic_Conv_nestedTacticCore___closed__6; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -9218,12 +9248,12 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv___closed__16() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv___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_Conv_conv_quot___closed__9; -x_2 = l_Lean_Parser_Tactic_Conv_conv___closed__15; +x_2 = l_Lean_Parser_Tactic_Conv_conv___closed__13; x_3 = l_Lean_Parser_Tactic_Conv_convSeq; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -9232,13 +9262,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv___closed__17() { +static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv___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_Conv_conv___closed__1; x_2 = lean_unsigned_to_nat(1022u); -x_3 = l_Lean_Parser_Tactic_Conv_conv___closed__16; +x_3 = l_Lean_Parser_Tactic_Conv_conv___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); @@ -9250,7 +9280,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_Conv_conv() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Tactic_Conv_conv___closed__17; +x_1 = l_Lean_Parser_Tactic_Conv_conv___closed__15; return x_1; } } @@ -9595,6 +9625,8 @@ l_Lean_Parser_Tactic_Conv_ext___closed__14 = _init_l_Lean_Parser_Tactic_Conv_ext lean_mark_persistent(l_Lean_Parser_Tactic_Conv_ext___closed__14); l_Lean_Parser_Tactic_Conv_ext___closed__15 = _init_l_Lean_Parser_Tactic_Conv_ext___closed__15(); lean_mark_persistent(l_Lean_Parser_Tactic_Conv_ext___closed__15); +l_Lean_Parser_Tactic_Conv_ext___closed__16 = _init_l_Lean_Parser_Tactic_Conv_ext___closed__16(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_ext___closed__16); l_Lean_Parser_Tactic_Conv_ext = _init_l_Lean_Parser_Tactic_Conv_ext(); lean_mark_persistent(l_Lean_Parser_Tactic_Conv_ext); l_Lean_Parser_Tactic_Conv_change___closed__1 = _init_l_Lean_Parser_Tactic_Conv_change___closed__1(); @@ -9629,8 +9661,6 @@ l_Lean_Parser_Tactic_Conv_delta___closed__5 = _init_l_Lean_Parser_Tactic_Conv_de lean_mark_persistent(l_Lean_Parser_Tactic_Conv_delta___closed__5); l_Lean_Parser_Tactic_Conv_delta___closed__6 = _init_l_Lean_Parser_Tactic_Conv_delta___closed__6(); lean_mark_persistent(l_Lean_Parser_Tactic_Conv_delta___closed__6); -l_Lean_Parser_Tactic_Conv_delta___closed__7 = _init_l_Lean_Parser_Tactic_Conv_delta___closed__7(); -lean_mark_persistent(l_Lean_Parser_Tactic_Conv_delta___closed__7); l_Lean_Parser_Tactic_Conv_delta = _init_l_Lean_Parser_Tactic_Conv_delta(); lean_mark_persistent(l_Lean_Parser_Tactic_Conv_delta); l_Lean_Parser_Tactic_Conv_unfold___closed__1 = _init_l_Lean_Parser_Tactic_Conv_unfold___closed__1(); @@ -9643,8 +9673,6 @@ l_Lean_Parser_Tactic_Conv_unfold___closed__4 = _init_l_Lean_Parser_Tactic_Conv_u lean_mark_persistent(l_Lean_Parser_Tactic_Conv_unfold___closed__4); l_Lean_Parser_Tactic_Conv_unfold___closed__5 = _init_l_Lean_Parser_Tactic_Conv_unfold___closed__5(); lean_mark_persistent(l_Lean_Parser_Tactic_Conv_unfold___closed__5); -l_Lean_Parser_Tactic_Conv_unfold___closed__6 = _init_l_Lean_Parser_Tactic_Conv_unfold___closed__6(); -lean_mark_persistent(l_Lean_Parser_Tactic_Conv_unfold___closed__6); l_Lean_Parser_Tactic_Conv_unfold = _init_l_Lean_Parser_Tactic_Conv_unfold(); lean_mark_persistent(l_Lean_Parser_Tactic_Conv_unfold); l_Lean_Parser_Tactic_Conv_pattern___closed__1 = _init_l_Lean_Parser_Tactic_Conv_pattern___closed__1(); @@ -9761,18 +9789,6 @@ l_Lean_Parser_Tactic_Conv_dsimp___closed__12 = _init_l_Lean_Parser_Tactic_Conv_d lean_mark_persistent(l_Lean_Parser_Tactic_Conv_dsimp___closed__12); l_Lean_Parser_Tactic_Conv_dsimp___closed__13 = _init_l_Lean_Parser_Tactic_Conv_dsimp___closed__13(); lean_mark_persistent(l_Lean_Parser_Tactic_Conv_dsimp___closed__13); -l_Lean_Parser_Tactic_Conv_dsimp___closed__14 = _init_l_Lean_Parser_Tactic_Conv_dsimp___closed__14(); -lean_mark_persistent(l_Lean_Parser_Tactic_Conv_dsimp___closed__14); -l_Lean_Parser_Tactic_Conv_dsimp___closed__15 = _init_l_Lean_Parser_Tactic_Conv_dsimp___closed__15(); -lean_mark_persistent(l_Lean_Parser_Tactic_Conv_dsimp___closed__15); -l_Lean_Parser_Tactic_Conv_dsimp___closed__16 = _init_l_Lean_Parser_Tactic_Conv_dsimp___closed__16(); -lean_mark_persistent(l_Lean_Parser_Tactic_Conv_dsimp___closed__16); -l_Lean_Parser_Tactic_Conv_dsimp___closed__17 = _init_l_Lean_Parser_Tactic_Conv_dsimp___closed__17(); -lean_mark_persistent(l_Lean_Parser_Tactic_Conv_dsimp___closed__17); -l_Lean_Parser_Tactic_Conv_dsimp___closed__18 = _init_l_Lean_Parser_Tactic_Conv_dsimp___closed__18(); -lean_mark_persistent(l_Lean_Parser_Tactic_Conv_dsimp___closed__18); -l_Lean_Parser_Tactic_Conv_dsimp___closed__19 = _init_l_Lean_Parser_Tactic_Conv_dsimp___closed__19(); -lean_mark_persistent(l_Lean_Parser_Tactic_Conv_dsimp___closed__19); l_Lean_Parser_Tactic_Conv_dsimp = _init_l_Lean_Parser_Tactic_Conv_dsimp(); lean_mark_persistent(l_Lean_Parser_Tactic_Conv_dsimp); l_Lean_Parser_Tactic_Conv_simpMatch___closed__1 = _init_l_Lean_Parser_Tactic_Conv_simpMatch___closed__1(); @@ -10017,40 +10033,42 @@ l_Lean_Parser_Tactic_Conv_case_x27 = _init_l_Lean_Parser_Tactic_Conv_case_x27(); lean_mark_persistent(l_Lean_Parser_Tactic_Conv_case_x27); l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__case_x27__1___closed__1 = _init_l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__case_x27__1___closed__1(); lean_mark_persistent(l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__case_x27__1___closed__1); -l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__1 = _init_l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__1(); -lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__1); -l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__2 = _init_l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__2(); -lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__2); -l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__3 = _init_l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__3(); -lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__3); -l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__4 = _init_l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__4(); -lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__4); -l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__5 = _init_l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__5(); -lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__5); -l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__6 = _init_l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__6(); -lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__6); -l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__7 = _init_l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__7(); -lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__7); -l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__8 = _init_l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__8(); -lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__8); -l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__9 = _init_l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__9(); -lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e_____closed__9); -l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e__ = _init_l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e__(); -lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convNext___x3d_x3e__); -l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext___x3d_x3e____1___closed__1 = _init_l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext___x3d_x3e____1___closed__1(); -lean_mark_persistent(l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext___x3d_x3e____1___closed__1); -l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext___x3d_x3e____1___closed__2 = _init_l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext___x3d_x3e____1___closed__2(); -lean_mark_persistent(l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext___x3d_x3e____1___closed__2); -l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext___x3d_x3e____1___closed__3 = _init_l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext___x3d_x3e____1___closed__3(); -lean_mark_persistent(l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext___x3d_x3e____1___closed__3); -l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext___x3d_x3e____1___closed__4 = _init_l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext___x3d_x3e____1___closed__4(); -lean_mark_persistent(l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext___x3d_x3e____1___closed__4); -l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext___x3d_x3e____1___closed__5 = _init_l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext___x3d_x3e____1___closed__5(); -lean_mark_persistent(l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext___x3d_x3e____1___closed__5); -l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext___x3d_x3e____1___closed__6 = _init_l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext___x3d_x3e____1___closed__6(); -lean_mark_persistent(l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext___x3d_x3e____1___closed__6); -l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext___x3d_x3e____1___closed__7 = _init_l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext___x3d_x3e____1___closed__7(); -lean_mark_persistent(l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext___x3d_x3e____1___closed__7); +l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__1 = _init_l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__1(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__1); +l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__2 = _init_l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__2(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__2); +l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__3 = _init_l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__3(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__3); +l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__4 = _init_l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__4(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__4); +l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__5 = _init_l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__5(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__5); +l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__6 = _init_l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__6(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__6); +l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__7 = _init_l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__7(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__7); +l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__8 = _init_l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__8(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__8); +l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__9 = _init_l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__9(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__9); +l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__10 = _init_l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__10(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e_____closed__10); +l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e__ = _init_l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e__(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convNext_____x3d_x3e__); +l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext_____x3d_x3e____1___closed__1 = _init_l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext_____x3d_x3e____1___closed__1(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext_____x3d_x3e____1___closed__1); +l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext_____x3d_x3e____1___closed__2 = _init_l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext_____x3d_x3e____1___closed__2(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext_____x3d_x3e____1___closed__2); +l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext_____x3d_x3e____1___closed__3 = _init_l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext_____x3d_x3e____1___closed__3(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext_____x3d_x3e____1___closed__3); +l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext_____x3d_x3e____1___closed__4 = _init_l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext_____x3d_x3e____1___closed__4(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext_____x3d_x3e____1___closed__4); +l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext_____x3d_x3e____1___closed__5 = _init_l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext_____x3d_x3e____1___closed__5(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext_____x3d_x3e____1___closed__5); +l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext_____x3d_x3e____1___closed__6 = _init_l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext_____x3d_x3e____1___closed__6(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext_____x3d_x3e____1___closed__6); +l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext_____x3d_x3e____1___closed__7 = _init_l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext_____x3d_x3e____1___closed__7(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convNext_____x3d_x3e____1___closed__7); l_Lean_Parser_Tactic_Conv_focus___closed__1 = _init_l_Lean_Parser_Tactic_Conv_focus___closed__1(); lean_mark_persistent(l_Lean_Parser_Tactic_Conv_focus___closed__1); l_Lean_Parser_Tactic_Conv_focus___closed__2 = _init_l_Lean_Parser_Tactic_Conv_focus___closed__2(); @@ -10253,20 +10271,20 @@ 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_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(); @@ -10285,30 +10303,32 @@ l_Lean_Parser_Tactic_Conv_enterArg___closed__8 = _init_l_Lean_Parser_Tactic_Conv lean_mark_persistent(l_Lean_Parser_Tactic_Conv_enterArg___closed__8); l_Lean_Parser_Tactic_Conv_enterArg = _init_l_Lean_Parser_Tactic_Conv_enterArg(); lean_mark_persistent(l_Lean_Parser_Tactic_Conv_enterArg); -l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__1 = _init_l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__1(); -lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__1); -l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__2 = _init_l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__2(); -lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__2); -l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__3 = _init_l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__3(); -lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__3); -l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__4 = _init_l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__4(); -lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__4); -l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__5 = _init_l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__5(); -lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__5); -l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__6 = _init_l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__6(); -lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__6); -l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__7 = _init_l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__7(); -lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__7); -l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__8 = _init_l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__8(); -lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__8); -l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__9 = _init_l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__9(); -lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__9); -l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__10 = _init_l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__10(); -lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d___closed__10); -l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d = _init_l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d(); -lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convEnter_x5b_____x5d); -l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convEnter_x5b_____x5d__1___closed__1 = _init_l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convEnter_x5b_____x5d__1___closed__1(); -lean_mark_persistent(l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convEnter_x5b_____x5d__1___closed__1); +l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__1 = _init_l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__1(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__1); +l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__2 = _init_l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__2(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__2); +l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__3 = _init_l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__3(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__3); +l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__4 = _init_l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__4(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__4); +l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__5 = _init_l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__5(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__5); +l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__6 = _init_l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__6(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__6); +l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__7 = _init_l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__7(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__7); +l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__8 = _init_l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__8(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__8); +l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__9 = _init_l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__9(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__9); +l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__10 = _init_l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__10(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d___closed__10); +l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d = _init_l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convEnter_x5b___x5d); +l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convEnter_x5b___x5d__1___lambda__2___closed__1 = _init_l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convEnter_x5b___x5d__1___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convEnter_x5b___x5d__1___lambda__2___closed__1); +l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convEnter_x5b___x5d__1___closed__1 = _init_l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convEnter_x5b___x5d__1___closed__1(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convEnter_x5b___x5d__1___closed__1); l_Lean_Parser_Tactic_Conv_convApply_____closed__1 = _init_l_Lean_Parser_Tactic_Conv_convApply_____closed__1(); lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convApply_____closed__1); l_Lean_Parser_Tactic_Conv_convApply_____closed__2 = _init_l_Lean_Parser_Tactic_Conv_convApply_____closed__2(); @@ -10361,6 +10381,22 @@ l_Lean_Parser_Tactic_Conv_first___closed__16 = _init_l_Lean_Parser_Tactic_Conv_f lean_mark_persistent(l_Lean_Parser_Tactic_Conv_first___closed__16); l_Lean_Parser_Tactic_Conv_first___closed__17 = _init_l_Lean_Parser_Tactic_Conv_first___closed__17(); lean_mark_persistent(l_Lean_Parser_Tactic_Conv_first___closed__17); +l_Lean_Parser_Tactic_Conv_first___closed__18 = _init_l_Lean_Parser_Tactic_Conv_first___closed__18(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_first___closed__18); +l_Lean_Parser_Tactic_Conv_first___closed__19 = _init_l_Lean_Parser_Tactic_Conv_first___closed__19(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_first___closed__19); +l_Lean_Parser_Tactic_Conv_first___closed__20 = _init_l_Lean_Parser_Tactic_Conv_first___closed__20(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_first___closed__20); +l_Lean_Parser_Tactic_Conv_first___closed__21 = _init_l_Lean_Parser_Tactic_Conv_first___closed__21(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_first___closed__21); +l_Lean_Parser_Tactic_Conv_first___closed__22 = _init_l_Lean_Parser_Tactic_Conv_first___closed__22(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_first___closed__22); +l_Lean_Parser_Tactic_Conv_first___closed__23 = _init_l_Lean_Parser_Tactic_Conv_first___closed__23(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_first___closed__23); +l_Lean_Parser_Tactic_Conv_first___closed__24 = _init_l_Lean_Parser_Tactic_Conv_first___closed__24(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_first___closed__24); +l_Lean_Parser_Tactic_Conv_first___closed__25 = _init_l_Lean_Parser_Tactic_Conv_first___closed__25(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv_first___closed__25); l_Lean_Parser_Tactic_Conv_first = _init_l_Lean_Parser_Tactic_Conv_first(); lean_mark_persistent(l_Lean_Parser_Tactic_Conv_first); l_Lean_Parser_Tactic_Conv_convTry_____closed__1 = _init_l_Lean_Parser_Tactic_Conv_convTry_____closed__1(); @@ -10413,6 +10449,8 @@ l_Lean_Parser_Tactic_Conv_convRepeat_____closed__6 = _init_l_Lean_Parser_Tactic_ lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convRepeat_____closed__6); l_Lean_Parser_Tactic_Conv_convRepeat__ = _init_l_Lean_Parser_Tactic_Conv_convRepeat__(); lean_mark_persistent(l_Lean_Parser_Tactic_Conv_convRepeat__); +l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convRepeat____1___closed__1 = _init_l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convRepeat____1___closed__1(); +lean_mark_persistent(l_Lean_Parser_Tactic_Conv___aux__Init__Conv______macroRules__Lean__Parser__Tactic__Conv__convRepeat____1___closed__1); l_Lean_Parser_Tactic_Conv_conv___closed__1 = _init_l_Lean_Parser_Tactic_Conv_conv___closed__1(); lean_mark_persistent(l_Lean_Parser_Tactic_Conv_conv___closed__1); l_Lean_Parser_Tactic_Conv_conv___closed__2 = _init_l_Lean_Parser_Tactic_Conv_conv___closed__2(); @@ -10443,10 +10481,6 @@ l_Lean_Parser_Tactic_Conv_conv___closed__14 = _init_l_Lean_Parser_Tactic_Conv_co lean_mark_persistent(l_Lean_Parser_Tactic_Conv_conv___closed__14); l_Lean_Parser_Tactic_Conv_conv___closed__15 = _init_l_Lean_Parser_Tactic_Conv_conv___closed__15(); lean_mark_persistent(l_Lean_Parser_Tactic_Conv_conv___closed__15); -l_Lean_Parser_Tactic_Conv_conv___closed__16 = _init_l_Lean_Parser_Tactic_Conv_conv___closed__16(); -lean_mark_persistent(l_Lean_Parser_Tactic_Conv_conv___closed__16); -l_Lean_Parser_Tactic_Conv_conv___closed__17 = _init_l_Lean_Parser_Tactic_Conv_conv___closed__17(); -lean_mark_persistent(l_Lean_Parser_Tactic_Conv_conv___closed__17); l_Lean_Parser_Tactic_Conv_conv = _init_l_Lean_Parser_Tactic_Conv_conv(); lean_mark_persistent(l_Lean_Parser_Tactic_Conv_conv); return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Init/Data/String/Basic.c b/stage0/stdlib/Init/Data/String/Basic.c index c77675b1c976..80b2fe998d4a 100644 --- a/stage0/stdlib/Init/Data/String/Basic.c +++ b/stage0/stdlib/Init/Data/String/Basic.c @@ -58,7 +58,6 @@ LEAN_EXPORT lean_object* l_String_set___boxed(lean_object*, lean_object*, lean_o LEAN_EXPORT lean_object* l_String_mkIterator(lean_object*); LEAN_EXPORT lean_object* l_String_foldrAux___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_String_foldlAux(lean_object*); -LEAN_EXPORT uint32_t l___private_Init_Data_String_Basic_0__String_utf8GetAux(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_String_atEnd___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Substring_take(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_repeatTR_loop___at_String_pushn___spec__1___boxed(lean_object*, lean_object*, lean_object*); @@ -66,19 +65,17 @@ LEAN_EXPORT lean_object* l_Substring_foldr___rarg___boxed(lean_object*, lean_obj LEAN_EXPORT lean_object* l_String_trim(lean_object*); LEAN_EXPORT lean_object* l_String_isNat___lambda__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Substring_trimRight(lean_object*); +LEAN_EXPORT uint32_t l_String_utf8GetAux(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_String_Iterator_next(lean_object*); LEAN_EXPORT lean_object* l_String_isPrefixOf___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Substring_front___boxed(lean_object*); LEAN_EXPORT uint8_t l_String_contains(lean_object*, uint32_t); LEAN_EXPORT lean_object* l_String_length___boxed(lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_String_Basic_0__String_utf8SetAux___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_String_Basic_0__String_utf8SetAux(uint32_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Substring_posOf(lean_object*, uint32_t); LEAN_EXPORT lean_object* l_String_foldl(lean_object*); LEAN_EXPORT lean_object* l_String_Iterator_remainingBytes___boxed(lean_object*); LEAN_EXPORT lean_object* l_String_Iterator_remainingToString(lean_object*); LEAN_EXPORT uint8_t l_Substring_atEnd(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_String_Basic_0__String_utf8PrevAux___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_String_revFind(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Substring_posOf___boxed(lean_object*, lean_object*); static lean_object* l_Substring_extract___closed__1; @@ -96,12 +93,16 @@ LEAN_EXPORT lean_object* l_String_dropRightWhile___boxed(lean_object*, lean_obje lean_object* lean_string_push(lean_object*, uint32_t); LEAN_EXPORT lean_object* l_String_pushn(lean_object*, uint32_t, lean_object*); LEAN_EXPORT lean_object* l_String_substrEq___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_String_utf8GetAux___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Substring_next(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_String_utf8GetAux_x3f___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_String_intercalate_go(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_String_Iterator_extract(lean_object*, lean_object*); LEAN_EXPORT uint32_t l_String_Iterator_curr(lean_object*); LEAN_EXPORT lean_object* l_Substring_all___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_String_foldrAux_loop(lean_object*); +LEAN_EXPORT lean_object* l_String_utf8SetAux___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_String_Basic_0__Substring_takeWhileAux(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_repeatTR_loop___at_String_pushn___spec__1(uint32_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_String_replace_loop___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -120,7 +121,6 @@ uint8_t l_instDecidableNot___rarg(uint8_t); LEAN_EXPORT lean_object* l_String_get_x3f___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_String_toNat_x3f___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Substring_all(lean_object*, lean_object*); -LEAN_EXPORT uint8_t l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(lean_object*, lean_object*); lean_object* lean_string_utf8_next(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_String_Iterator_curr___boxed(lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); @@ -136,6 +136,7 @@ LEAN_EXPORT uint8_t l_String_anyAux(lean_object*, lean_object*, lean_object*, le LEAN_EXPORT lean_object* l_String_instInhabitedString; LEAN_EXPORT lean_object* l_Substring_splitOn_loop(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Substring_takeWhile(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_String_foldrAux_loop___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Substring_contains(lean_object*, uint32_t); LEAN_EXPORT lean_object* l_Substring_get___boxed(lean_object*, lean_object*); @@ -175,7 +176,6 @@ LEAN_EXPORT lean_object* l_String_foldlAux___rarg___boxed(lean_object*, lean_obj LEAN_EXPORT lean_object* l___private_Init_Data_String_Basic_0__Substring_takeWhileAux___at_String_nextUntil___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_String_offsetOfPosAux(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_String_takeRight(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_String_Basic_0__String_utf8GetAux___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_String_extract_go_u2081(lean_object*, lean_object*, lean_object*, lean_object*); uint32_t lean_string_utf8_get_fast(lean_object*, lean_object*); lean_object* lean_string_utf8_prev(lean_object*, lean_object*); @@ -190,6 +190,7 @@ LEAN_EXPORT lean_object* l___private_Init_Data_String_Basic_0__Substring_takeWhi LEAN_EXPORT lean_object* l_String_firstDiffPos_loop(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_String_all___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_String_Iterator_prevn(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_String_utf8SetAux(uint32_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Substring_splitOn_loop___closed__1; LEAN_EXPORT lean_object* l_String_revPosOfAux(lean_object*, uint32_t, lean_object*); LEAN_EXPORT lean_object* l_String_singleton___boxed(lean_object*); @@ -204,7 +205,6 @@ LEAN_EXPORT lean_object* l_String_Iterator_extract___boxed(lean_object*, lean_ob LEAN_EXPORT lean_object* l_String_modify(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_String_revPosOf(lean_object*, uint32_t); LEAN_EXPORT lean_object* l_String_Iterator_atEnd___boxed(lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_String_Basic_0__String_utf8GetAux_x3f___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Substring_splitOn_loop___closed__2; LEAN_EXPORT uint8_t l_String_all___lambda__1(lean_object*, uint32_t); lean_object* l_String_csize(uint32_t); @@ -224,7 +224,6 @@ LEAN_EXPORT lean_object* l_String_findAux(lean_object*, lean_object*, lean_objec LEAN_EXPORT lean_object* l_String_replace(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Substring_beq___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_String_Iterator_remainingBytes(lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_String_Basic_0__String_utf8PrevAux(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_String_takeRightWhile___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_String_get_x21___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_String_nextUntil(lean_object*, lean_object*, lean_object*); @@ -236,12 +235,10 @@ static lean_object* l_String_isNat___closed__1; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_String_Basic_0__Substring_takeRightWhileAux___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_String_foldl___rarg___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_String_Basic_0__String_utf8GetAux_x3f(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_String_prev___boxed(lean_object*, lean_object*); uint8_t lean_uint32_dec_eq(uint32_t, uint32_t); LEAN_EXPORT lean_object* l_Substring_toString___boxed(lean_object*); LEAN_EXPORT lean_object* l_String_extract___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_String_contains___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_String_split___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_String_map(lean_object*, lean_object*); @@ -271,11 +268,13 @@ LEAN_EXPORT lean_object* l_String_revFindAux___boxed(lean_object*, lean_object*, LEAN_EXPORT lean_object* l_Substring_foldl(lean_object*); lean_object* lean_nat_mul(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_String_instDecidableEqIterator(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_String_utf8PrevAux(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint32_t l_Substring_front(lean_object*); LEAN_EXPORT lean_object* l_Substring_extract___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_List_reverse___rarg(lean_object*); lean_object* lean_string_utf8_next_fast(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_String_posOf___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_String_utf8PrevAux___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_String_isEmpty___boxed(lean_object*); static lean_object* l_String_toNat_x3f___closed__1; LEAN_EXPORT lean_object* l_String_intercalate(lean_object*, lean_object*); @@ -292,6 +291,7 @@ LEAN_EXPORT lean_object* l_String_toNat_x3f___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_String_Basic_0__Substring_takeWhileAux___at_Substring_trimLeft___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_String_Pos_min___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_String_posOf(lean_object*, uint32_t); +LEAN_EXPORT lean_object* l_String_utf8GetAux_x3f(lean_object*, lean_object*, lean_object*); static lean_object* l_Substring_hasBeq___closed__1; LEAN_EXPORT uint8_t l_String_Iterator_atEnd(lean_object*); LEAN_EXPORT lean_object* l_String_trimRight(lean_object*); @@ -400,7 +400,7 @@ x_2 = lean_string_data(x_1); return x_2; } } -LEAN_EXPORT uint32_t l___private_Init_Data_String_Basic_0__String_utf8GetAux(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT uint32_t l_String_utf8GetAux(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -444,11 +444,11 @@ return x_12; } } } -LEAN_EXPORT lean_object* l___private_Init_Data_String_Basic_0__String_utf8GetAux___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_String_utf8GetAux___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { uint32_t x_4; lean_object* x_5; -x_4 = l___private_Init_Data_String_Basic_0__String_utf8GetAux(x_1, x_2, x_3); +x_4 = l_String_utf8GetAux(x_1, x_2, x_3); lean_dec(x_3); x_5 = lean_box_uint32(x_4); return x_5; @@ -465,12 +465,13 @@ x_4 = lean_box_uint32(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Init_Data_String_Basic_0__String_utf8GetAux_x3f(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_String_utf8GetAux_x3f(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) { lean_object* x_4; +lean_dec(x_2); x_4 = lean_box(0); return x_4; } @@ -485,36 +486,35 @@ lean_dec(x_1); x_7 = lean_nat_dec_eq(x_2, x_3); if (x_7 == 0) { -uint32_t x_8; lean_object* x_9; lean_object* x_10; uint32_t x_11; lean_object* x_12; lean_object* x_13; +uint32_t x_8; lean_object* x_9; lean_object* x_10; x_8 = lean_unbox_uint32(x_5); lean_dec(x_5); x_9 = l_String_csize(x_8); x_10 = lean_nat_add(x_2, x_9); lean_dec(x_9); -x_11 = l___private_Init_Data_String_Basic_0__String_utf8GetAux(x_6, x_10, x_3); -x_12 = lean_box_uint32(x_11); -x_13 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_13, 0, x_12); -return x_13; +lean_dec(x_2); +x_1 = x_6; +x_2 = x_10; +goto _start; } else { -lean_object* x_14; +lean_object* x_12; lean_dec(x_6); -x_14 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_14, 0, x_5); -return x_14; +lean_dec(x_2); +x_12 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_12, 0, x_5); +return x_12; } } } } -LEAN_EXPORT lean_object* l___private_Init_Data_String_Basic_0__String_utf8GetAux_x3f___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_String_utf8GetAux_x3f___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Init_Data_String_Basic_0__String_utf8GetAux_x3f(x_1, x_2, x_3); +x_4 = l_String_utf8GetAux_x3f(x_1, x_2, x_3); lean_dec(x_3); -lean_dec(x_2); return x_4; } } @@ -539,7 +539,7 @@ x_4 = lean_box_uint32(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Init_Data_String_Basic_0__String_utf8SetAux(uint32_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_String_utf8SetAux(uint32_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { if (lean_obj_tag(x_2) == 0) @@ -565,7 +565,7 @@ x_10 = lean_unbox_uint32(x_7); x_11 = l_String_csize(x_10); x_12 = lean_nat_add(x_3, x_11); lean_dec(x_11); -x_13 = l___private_Init_Data_String_Basic_0__String_utf8SetAux(x_1, x_8, x_12, x_4); +x_13 = l_String_utf8SetAux(x_1, x_8, x_12, x_4); lean_dec(x_12); lean_ctor_set(x_2, 1, x_13); return x_2; @@ -595,7 +595,7 @@ x_18 = lean_unbox_uint32(x_15); x_19 = l_String_csize(x_18); x_20 = lean_nat_add(x_3, x_19); lean_dec(x_19); -x_21 = l___private_Init_Data_String_Basic_0__String_utf8SetAux(x_1, x_16, x_20, x_4); +x_21 = l_String_utf8SetAux(x_1, x_16, x_20, x_4); lean_dec(x_20); x_22 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_22, 0, x_15); @@ -616,13 +616,13 @@ return x_24; } } } -LEAN_EXPORT lean_object* l___private_Init_Data_String_Basic_0__String_utf8SetAux___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_String_utf8SetAux___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { uint32_t x_5; lean_object* x_6; x_5 = lean_unbox_uint32(x_1); lean_dec(x_1); -x_6 = l___private_Init_Data_String_Basic_0__String_utf8SetAux(x_5, x_2, x_3, x_4); +x_6 = l_String_utf8SetAux(x_5, x_2, x_3, x_4); lean_dec(x_4); lean_dec(x_3); return x_6; @@ -671,7 +671,7 @@ lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l___private_Init_Data_String_Basic_0__String_utf8PrevAux(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_String_utf8PrevAux(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -711,11 +711,11 @@ return x_2; } } } -LEAN_EXPORT lean_object* l___private_Init_Data_String_Basic_0__String_utf8PrevAux___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_String_utf8PrevAux___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Init_Data_String_Basic_0__String_utf8PrevAux(x_1, x_2, x_3); +x_4 = l_String_utf8PrevAux(x_1, x_2, x_3); lean_dec(x_3); return x_4; } @@ -1822,7 +1822,7 @@ lean_dec(x_1); return x_3; } } -LEAN_EXPORT uint8_t l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(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; @@ -1845,11 +1845,11 @@ return x_9; } } } -LEAN_EXPORT lean_object* l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_2); +x_3 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -1860,7 +1860,7 @@ LEAN_EXPORT uint8_t l_String_instDecidableEqIterator(lean_object* x_1, lean_obje _start: { uint8_t x_3; -x_3 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_2); +x_3 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_2); return x_3; } } diff --git a/stage0/stdlib/Init/Meta.c b/stage0/stdlib/Init/Meta.c index 9dbef0b148c6..38ab6b07bd17 100644 --- a/stage0/stdlib/Init/Meta.c +++ b/stage0/stdlib/Init/Meta.c @@ -16,15 +16,17 @@ extern "C" { LEAN_EXPORT lean_object* l_Lean_Syntax_decodeQuotedChar___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instQuoteListMkStr1(lean_object*); static lean_object* l_Lean_Name_reprPrec___closed__7; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__9; static lean_object* l_Lean_Parser_Tactic_simpAllAutoUnfold___closed__7; static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__3; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__4; static lean_object* l_Lean_Name_escapePart___closed__2; LEAN_EXPORT lean_object* l_Lean___aux__Init__Meta______macroRules__Lean__termEval__prio____1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__26; LEAN_EXPORT lean_object* l_Lean_evalPrio(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_quoteNameMk___closed__7; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__10; lean_object* lean_string_utf8_extract(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__6; LEAN_EXPORT uint8_t l_Lean_Name_instDecidableEqName(lean_object*, lean_object*); static lean_object* l_Lean_instQuoteSubstringMkStr1___closed__2; static lean_object* l_Lean_mkGroupNode___closed__2; @@ -35,25 +37,26 @@ LEAN_EXPORT lean_object* l_Lean_TSyntax_Compat_instCoeTailArraySyntaxTSepArray__ static lean_object* l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__4; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeNameLit(lean_object*); LEAN_EXPORT lean_object* l_Lean_mkCIdentFromRef___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Init_Meta_0__Lean_Syntax_beqPreresolved____x40_Init_Meta___hyg_3135_(lean_object*, lean_object*); static lean_object* l_Lean_Syntax_mkNameLit___closed__2; LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeNumLitTerm(lean_object*); static lean_object* l_Lean_Meta_Simp_neutralConfig___closed__1; LEAN_EXPORT lean_object* l_Lean_Syntax_mkStrLit___boxed(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__83; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18276____closed__1; LEAN_EXPORT lean_object* l_Lean_TSyntax_Compat_instCoeTailArraySyntaxTSyntaxArray___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_instQuoteProdMkStr1___rarg(lean_object*, lean_object*, lean_object*); lean_object* l___private_Init_Data_Repr_0__reprSourceInfo____x40_Init_Data_Repr___hyg_2311_(lean_object*, lean_object*); static lean_object* l_Lean_origin___closed__1; static lean_object* l_Lean_Parser_Tactic_simpArith___closed__9; static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__24; -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000_(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__26; +static lean_object* l_List_foldr___at_Substring_toName___spec__1___closed__4; LEAN_EXPORT lean_object* l_Lean_Syntax_getHead_x3f___lambda__1(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__65; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__50; LEAN_EXPORT lean_object* l_Lean_Syntax_mkNameLit(lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__14; static lean_object* l_Lean_Parser_Tactic_simpAllAutoUnfold___closed__2; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__63; static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__16; @@ -62,19 +65,22 @@ static lean_object* l_Lean_Parser_Tactic_simpAllArithAutoUnfold___closed__6; 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*); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_splitNameLitAux(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__6; -static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__6; +LEAN_EXPORT lean_object* l_Lean_HygieneInfo_mkIdent___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Macro_throwErrorAt___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___closed__1; uint32_t lean_string_utf8_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_getSepElems___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_decodeScientificLitVal_x3f_decode___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__8; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__9; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_updateFirst___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_beqConfig____x40_Init_Meta___hyg_12968____boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__10; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeStrLit(lean_object*); +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__5; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__1; LEAN_EXPORT lean_object* l_Lean_Syntax_isScientificLit_x3f(lean_object*); 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_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__13; static lean_object* l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__1; lean_object* lean_mk_empty_array_with_capacity(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_TSepArray_push___rarg(lean_object*, lean_object*, lean_object*); @@ -82,18 +88,14 @@ static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean LEAN_EXPORT lean_object* l_Lean_Syntax_setHeadInfo(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_isNameLit_x3f___boxed(lean_object*); static lean_object* l_Lean_instQuoteBoolMkStr1___closed__4; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_dsimpKind; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeScientificLitVal_x3f_decodeExp___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpArith___closed__5; LEAN_EXPORT lean_object* l_Lean_Syntax_instReprTSyntax(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__12; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__3; -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__9; static lean_object* l_Lean_Parser_Tactic_simpAllAutoUnfold___closed__11; -static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__12; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__7; static lean_object* l_Lean_TSyntax_expandInterpolatedStr___closed__5; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18276____closed__2; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__10; LEAN_EXPORT lean_object* l_Lean_TSyntax_getId(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__18; @@ -101,13 +103,13 @@ LEAN_EXPORT lean_object* l_Lean_githash; LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeTSyntaxConsSyntaxNodeKindNil___boxed(lean_object*, lean_object*); static lean_object* l_Lean_versionString___closed__1; LEAN_EXPORT lean_object* l_Lean_TSyntax_getName(lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____boxed(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__84; lean_object* lean_uint32_to_nat(uint32_t); LEAN_EXPORT lean_object* l_Lean_evalPrec(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_getTailInfo___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__40; LEAN_EXPORT lean_object* l_Lean_instInhabitedNameGenerator; -static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__17; 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___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__52; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2377____closed__20; @@ -115,9 +117,12 @@ static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__25; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_version_getMajor___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Name_toStringWithSep_maybeEscape___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_filterSepElemsM___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__22; +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__28; LEAN_EXPORT lean_object* l_List_foldl___at___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2236____spec__4(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__25; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__107; +static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__10; static lean_object* l_Lean_Name_reprPrec___closed__5; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__48; LEAN_EXPORT lean_object* lean_name_append_after(lean_object*, lean_object*); @@ -126,14 +131,16 @@ static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean static lean_object* l_Lean_termEval__prio_____closed__6; LEAN_EXPORT uint8_t l_Lean_Meta_instInhabitedEtaStructMode; static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__15; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__4; static lean_object* l_Lean_termEval__prec_____closed__8; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__39; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__100; extern lean_object* l_Lean_maxRecDepthErrorMessage; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__35; static lean_object* l_Lean_Syntax_instCoeIdentTSyntaxConsSyntaxNodeKindMkStr4Nil___closed__3; +static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__3; LEAN_EXPORT lean_object* l_Lean_Name_toString___boxed(lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__16; +static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__12; lean_object* l_Lean_Syntax_getHeadInfo(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_Parser_Tactic_simpAutoUnfold___closed__2; @@ -148,7 +155,6 @@ static lean_object* l_Lean_Parser_Tactic_simpArith___closed__4; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeScientificLitVal_x3f_decode(lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Parser_Tactic_location; LEAN_EXPORT lean_object* l_Lean_Meta_TransparencyMode_noConfusion___rarg___lambda__1(lean_object*); -LEAN_EXPORT lean_object* l_List_foldr___at_Lean_Syntax_decodeNameLit___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_evalPrec___closed__1; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_updateFirst___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Name_instReprName; @@ -158,7 +164,6 @@ static lean_object* l_Lean_Syntax_mkNumLit___closed__1; 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___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__80; LEAN_EXPORT lean_object* l_Lean_TSyntax_getString(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__3; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__37; LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeNumLitPrec(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__9; @@ -169,19 +174,17 @@ LEAN_EXPORT lean_object* l_Lean_Syntax_structEq___boxed(lean_object*, lean_objec static lean_object* l_Lean_TSyntax_Compat_instCoeTailArraySyntaxTSyntaxArray___closed__1; extern lean_object* l_Lean_reservedMacroScope; static lean_object* l_List_repr_x27___at___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2236____spec__1___closed__5; -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_12336____boxed(lean_object*, lean_object*); 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_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__7; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__13; lean_object* l_String_quote(lean_object*); LEAN_EXPORT lean_object* l_Lean_Name_toString(lean_object*, uint8_t); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__12; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__11; static lean_object* l_Lean_Syntax_decodeNatLitVal_x3f___closed__1; lean_object* l_Lean_Syntax_getId(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__134; static lean_object* l_Lean_Parser_Tactic_simpAllArith___closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_TransparencyMode_toCtorIdx___boxed(lean_object*); -LEAN_EXPORT uint8_t l_List_beq___at___private_Init_Meta_0__Lean_Syntax_beqPreresolved____x40_Init_Meta___hyg_3131____spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__27; 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_origin; lean_object* l_Lean_TSyntaxArray_rawImpl___rarg___boxed(lean_object*); @@ -191,44 +194,47 @@ LEAN_EXPORT lean_object* l_Lean_NameGenerator_curr(lean_object*); LEAN_EXPORT lean_object* l_Lean_monadNameGeneratorLift___rarg___lambda__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_version_major___closed__1; 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*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__18; lean_object* lean_array_push(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeStrLitTerm(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473_(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__11; LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeTSyntaxArrayTSepArray___boxed(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__25; LEAN_EXPORT lean_object* l_Lean_TSyntax_expandInterpolatedStr(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____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_13174____closed__17; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__4; static lean_object* l_Lean_TSyntax_expandInterpolatedStr___closed__2; 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_simpAllKind___closed__5; -static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__22; 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*); LEAN_EXPORT lean_object* l_Lean_mkIdentFrom___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprTSyntax____x40_Init_Meta___hyg_2642____rarg___closed__12; LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeArraySyntaxSepArray(lean_object*); static lean_object* l_Lean_Syntax_isFieldIdx_x3f___closed__1; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__7; static lean_object* l_Lean_TSyntax_expandInterpolatedStr___closed__4; -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19075_(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_mkHole___closed__1; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__46; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__15; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__18; static lean_object* l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__6; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__91; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__6; extern lean_object* l_Lean_Parser_Tactic_simpStar; LEAN_EXPORT lean_object* l_Lean_instQuoteStringStrLitKind(lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_DSimp_Config_zeta___default; uint8_t lean_usize_dec_eq(size_t, size_t); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__32; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__2; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__110; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__73; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__79; LEAN_EXPORT lean_object* l_Lean_mkCIdentFrom___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_dsimpAutoUnfold___closed__2; +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534_(uint8_t, lean_object*); 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_Lean_Name_escapePart___closed__4; static lean_object* l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__10; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__58; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__2; LEAN_EXPORT lean_object* l_Lean_Syntax_isScientificLit_x3f___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__12; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2377____closed__18; @@ -237,6 +243,8 @@ lean_object* l_Lean_Syntax_getArgs(lean_object*); static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2377____closed__11; static lean_object* l_Lean_version_specialDesc___closed__1; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__19; +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__1; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__9; static lean_object* l_Lean_Parser_Tactic_simpAllArithAutoUnfold___closed__9; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__69; static lean_object* l_Lean_Syntax_instReprTSyntax___closed__1; @@ -244,11 +252,13 @@ LEAN_EXPORT lean_object* l_Lean_Syntax_decodeScientificLitVal_x3f_decodeAfterExp static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2377____closed__15; LEAN_EXPORT lean_object* l_Lean_TSyntax_Compat_instCoeTailArraySyntaxTSepArray(lean_object*); LEAN_EXPORT lean_object* l_Lean_TSyntax_Compat_instCoeTailSyntaxTSyntax(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__21; LEAN_EXPORT lean_object* l_Lean_mkIdentFromRef___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_replaceRef(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____closed__7; +LEAN_EXPORT uint8_t l_List_beq___at___private_Init_Meta_0__Lean_Syntax_beqPreresolved____x40_Init_Meta___hyg_3135____spec__1(lean_object*, lean_object*); lean_object* l_String_trim(lean_object*); LEAN_EXPORT lean_object* l_Lean_isLetterLike___boxed(lean_object*); +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__7; LEAN_EXPORT lean_object* l_Lean_NameGenerator_idx___default; static lean_object* l_Lean_Parser_Tactic_simpAllArithAutoUnfold___closed__1; LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeOutTSepArrayTSyntaxArray(lean_object*, lean_object*); @@ -261,12 +271,11 @@ LEAN_EXPORT lean_object* l_Array_filterSepElemsM(lean_object*); lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_reprTSyntax____x40_Init_Meta___hyg_2642_(lean_object*); static lean_object* l_Lean_Parser_Tactic_tacticErw_______closed__5; -static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__21; +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__20; LEAN_EXPORT lean_object* l_Lean_Syntax_instEmptyCollectionSepArray(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__22; LEAN_EXPORT lean_object* l_repr___at___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2236____spec__2___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_Config_maxDischargeDepth___default; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__4; uint8_t l_String_contains(lean_object*, uint32_t); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_isNatLitAux___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_TSepArray_getElems___rarg(lean_object*); @@ -277,6 +286,7 @@ LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_ LEAN_EXPORT uint8_t l_Lean_isIdBeginEscape(uint32_t); LEAN_EXPORT lean_object* l_Lean_Syntax_decodeScientificLitVal_x3f(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_simpAllArith; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__8; LEAN_EXPORT lean_object* l_Lean_Syntax_SepArray_getElems___rarg___boxed(lean_object*); LEAN_EXPORT lean_object* l_Array_filterSepElems___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__87; @@ -286,7 +296,7 @@ LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_reprPreresolved___ static lean_object* l_Lean_Parser_Tactic_simpAllArith___closed__3; LEAN_EXPORT lean_object* l_Lean_Syntax_isFieldIdx_x3f___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__88; -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__21; +static lean_object* l_List_foldr___at_Substring_toName___spec__1___closed__1; LEAN_EXPORT uint8_t l_Lean_Syntax_instBEqTSyntax___rarg(lean_object*, lean_object*); static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2236____closed__1; uint8_t l_Char_isAlpha(uint32_t); @@ -297,11 +307,11 @@ LEAN_EXPORT lean_object* l_Lean_TSyntax_expandInterpolatedStrChunks(lean_object* LEAN_EXPORT lean_object* l_Lean_mkIdentFrom(lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Syntax_decodeScientificLitVal_x3f___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_instQuoteTermMkStr1; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__15; static lean_object* l_Lean_termEval__prec_____closed__1; -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____boxed(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__86; LEAN_EXPORT lean_object* l_Lean_Syntax_isLit_x3f___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_foldr___at_Substring_toName___spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__3; static lean_object* l_Lean_Parser_Tactic_dsimpKind___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__112; @@ -311,21 +321,20 @@ static lean_object* l_Lean_Syntax_unsetTrailing___closed__1; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_updateLast___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_setTailInfoAux(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isSubScriptAlnum___boxed(lean_object*); -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__10; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__17; LEAN_EXPORT lean_object* l_repr___at___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2377____spec__1(lean_object*); static lean_object* l_Lean_Parser_Tactic_simpArith___closed__1; LEAN_EXPORT lean_object* l_Lean_Syntax_TSepArray_getElems(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_beq___at_Lean_Syntax_structEq___spec__2___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__8; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteArray_go(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__25; lean_object* l_Lean_Syntax_node5(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__declareSimpLikeTactic__1___lambda__1___closed__48; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__61; uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__11; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__1; +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_beqPreresolved____x40_Init_Meta___hyg_3135____boxed(lean_object*, lean_object*); static lean_object* l_Lean___aux__Init__Meta______macroRules__Lean__Parser__Syntax__subPrec__1___closed__2; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__15; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2377____closed__12; LEAN_EXPORT uint32_t l_Lean_idBeginEscape; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__25; @@ -333,88 +342,93 @@ LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeTSyntaxConsSyntaxNodeKind___boxed LEAN_EXPORT lean_object* l_Lean_isGreek___boxed(lean_object*); static lean_object* l_List_repr_x27___at___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2236____spec__1___closed__6; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__6; -static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__10; static lean_object* l_Lean_Parser_Tactic_simpAllAutoUnfold___closed__12; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__15; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__114; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2377____closed__8; static lean_object* l_Lean_versionString___closed__6; LEAN_EXPORT uint8_t l_Lean_Name_escapePart___lambda__1(uint32_t); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__19; +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__27; static lean_object* l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_simpArithAutoUnfold; +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201_(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_mkSynthetic(lean_object*); static lean_object* l_Lean_Option_hasQuote___rarg___closed__6; LEAN_EXPORT lean_object* lean_name_append_before(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeTSyntaxConsSyntaxNodeKindNil___rarg___boxed(lean_object*); +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__11; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__21; static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__10; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__14; uint8_t lean_string_dec_eq(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703_(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_dsimpAutoUnfold___closed__1; static lean_object* l_Lean_toolchain___closed__6; LEAN_EXPORT lean_object* l_Lean_TSyntax_Compat_instCoeTailArraySyntaxTSepArray___rarg(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__5; LEAN_EXPORT uint8_t l_Lean_Meta_Simp_Config_autoUnfold___default; static lean_object* l_Lean_termEval__prio_____closed__5; static lean_object* l_Lean_termEval__prec_____closed__9; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__6; LEAN_EXPORT uint8_t l_Lean_Syntax_isAtom(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__8; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2377____closed__3; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__47; -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__11; LEAN_EXPORT lean_object* l_Lean_Meta_instReprEtaStructMode; LEAN_EXPORT uint8_t l_Lean_Meta_Rewrite_Config_offsetCnstrs___default; LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeOutTSyntaxArrayArraySyntax___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__3; LEAN_EXPORT lean_object* l_Lean___aux__Init__Meta______macroRules__Lean__Parser__Syntax__subPrec__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__19; +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__24; 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*); -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_beqEtaStructMode____x40_Init_Meta___hyg_12541____boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____lambda__1___closed__1; LEAN_EXPORT uint8_t l_Lean_version_isRelease; lean_object* lean_string_utf8_byte_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeCharLitTerm___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeInterpStrQuotedChar___boxed__const__1; static lean_object* l_Lean_Parser_Tactic_simpAllArith___closed__1; -static lean_object* l_List_foldr___at_Lean_Syntax_decodeNameLit___spec__1___closed__1; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__3; lean_object* lean_string_push(lean_object*, uint32_t); static lean_object* l_Lean_NameGenerator_namePrefix___default___closed__2; +static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__19; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__59; static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__2; static lean_object* l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__5; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__74; +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18519_(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__10; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__18; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__12; LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeTSyntaxConsSyntaxNodeKindNil(lean_object*, lean_object*); static lean_object* l_Lean_Meta_instBEqTransparencyMode___closed__1; LEAN_EXPORT lean_object* l_Lean_mkFreshId(lean_object*); +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18519____closed__1; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__71; LEAN_EXPORT lean_object* l_Lean_mkCIdentFromRef___rarg(lean_object*, lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Syntax_decodeStrLitAux(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2377____closed__13; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__14; LEAN_EXPORT lean_object* l_List_repr___at___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2377____spec__4(lean_object*, 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_simpAllAutoUnfold___closed__8; LEAN_EXPORT lean_object* l_Lean_Syntax_mkApp(lean_object*, lean_object*); extern lean_object* l_Lean_Parser_Tactic_discharger; static lean_object* l_Lean_Parser_Tactic_simpAllArithAutoUnfold___closed__4; +static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__7; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__14; static lean_object* l_Lean_Parser_Tactic_dsimpKind___closed__4; LEAN_EXPORT lean_object* l_Lean_Syntax_instEmptyCollectionTSepArray(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_decodeScientificLitVal_x3f_decodeAfterDot___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__4; LEAN_EXPORT lean_object* l_Lean_mkFreshId___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_dsimpKind___closed__5; LEAN_EXPORT lean_object* l_Lean_Name_toStringWithSep_maybeEscape(uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeNumLitPrio___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__7; +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__7; LEAN_EXPORT lean_object* l_Lean_Syntax_isInterpolatedStrLit_x3f(lean_object*); static lean_object* l_Lean_versionStringCore___closed__2; LEAN_EXPORT lean_object* l_Lean_Syntax_splitNameLit(lean_object*); static lean_object* l_Lean_TSyntax_expandInterpolatedStr___closed__1; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_updateFirst(lean_object*); -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__1; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__33; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__23; LEAN_EXPORT lean_object* l_Lean_Syntax_isNatLit_x3f___boxed(lean_object*); @@ -423,12 +437,13 @@ LEAN_EXPORT uint8_t l_Lean_Meta_DSimp_Config_iota___default; LEAN_EXPORT lean_object* l_Lean_Syntax_instReprPreresolved; LEAN_EXPORT lean_object* l_Lean_TSyntax_expandInterpolatedStr___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Meta_0__Lean_quoteList___rarg___closed__1; -static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__18; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeHexLitAux___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__71; static lean_object* l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__5; static lean_object* l_Lean_Parser_Tactic_tacticErw_______closed__4; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__12; +LEAN_EXPORT uint8_t l___private_Init_Meta_0__Lean_Meta_DSimp_beqConfig____x40_Init_Meta___hyg_12968_(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19318_(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__3; static lean_object* l_Lean_versionStringCore___closed__3; LEAN_EXPORT lean_object* l_Lean_TSyntax_expandInterpolatedStr___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -452,7 +467,6 @@ static lean_object* l_List_repr_x27___at___private_Init_Meta_0__Lean_Syntax_repr static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__27; static lean_object* l_Lean_Parser_Tactic_dsimpAutoUnfold___closed__5; static lean_object* l_Lean_Parser_Tactic_tacticErw_______closed__11; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__3; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__96; static lean_object* l_Lean_Meta_Simp_instReprConfig___closed__1; LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2236____spec__3(lean_object*, lean_object*); @@ -466,29 +480,29 @@ static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__5; LEAN_EXPORT lean_object* l_Lean_toolchain; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__3; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__6; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__23; +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____boxed(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__30; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____closed__5; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__37; uint8_t l_instDecidableNot___rarg(uint8_t); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__124; LEAN_EXPORT lean_object* l_Lean_Syntax_copyHeadTailInfoFrom(lean_object*, lean_object*); 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_13000____closed__12; +static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__17; static lean_object* l_Array_getSepElems___rarg___closed__1; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__8; LEAN_EXPORT lean_object* l_Lean_Meta_EtaStructMode_noConfusion(lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_mkSepArray___spec__1___closed__1; +LEAN_EXPORT uint8_t l___private_Init_Meta_0__Lean_Meta_Simp_beqConfig____x40_Init_Meta___hyg_13540_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2377____spec__6(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_DSimp_Config_proj___default; static lean_object* l_Lean_instQuoteProdMkStr1___rarg___closed__3; -static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__15; static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__25; static lean_object* l_Lean___aux__Init__Meta______macroRules__Lean__Parser__Syntax__subPrio__1___closed__2; lean_object* lean_string_utf8_next(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__14; +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__89; +LEAN_EXPORT uint8_t l___private_Init_Meta_0__Lean_Meta_beqEtaStructMode____x40_Init_Meta___hyg_12715_(uint8_t, uint8_t); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__105; uint8_t l_Char_isAlphanum(uint32_t); LEAN_EXPORT lean_object* l_Lean_Meta_instReprTransparencyMode; @@ -501,16 +515,17 @@ static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean static lean_object* l_Lean_version_patch___closed__1; LEAN_EXPORT lean_object* l_Lean_mkFreshId___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_instQuoteBoolMkStr1___closed__2; -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21536_(lean_object*, lean_object*, lean_object*); +static lean_object* l_List_foldr___at_Substring_toName___spec__1___closed__2; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__73; -static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__24; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_neutralConfig; static lean_object* l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Syntax_TSepArray_getElems___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Option_hasQuote___rarg___closed__5; +LEAN_EXPORT uint8_t l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_12510_(uint8_t, uint8_t); LEAN_EXPORT lean_object* l_Lean_Syntax_isFieldIdx_x3f(lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllArith___closed__7; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprTSyntax____x40_Init_Meta___hyg_2642____rarg___closed__7; +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_12510____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instQuoteNatNumLitKind(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__90; LEAN_EXPORT lean_object* l_Lean_Syntax_getOptionalIdent_x3f___boxed(lean_object*); @@ -520,7 +535,7 @@ static lean_object* l_List_repr_x27___at___private_Init_Meta_0__Lean_Syntax_repr LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_updateLast(lean_object*); lean_object* l_Lean_SourceInfo_fromRef(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Syntax_getSepArgs___boxed(lean_object*); -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____closed__8; +static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__17; LEAN_EXPORT uint8_t l_List_beq___at_Lean_Syntax_structEq___spec__2(lean_object*, lean_object*); uint8_t lean_uint32_dec_le(uint32_t, uint32_t); static lean_object* l_Lean_Parser_Tactic_dsimpAutoUnfold___closed__3; @@ -528,10 +543,10 @@ LEAN_EXPORT lean_object* l_Lean_Syntax_instEmptyCollectionSepArray___boxed(lean_ static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__81; lean_object* lean_nat_to_int(lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__3; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____lambda__1___closed__1; static lean_object* l_Lean_quoteNameMk___closed__6; 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__declareSimpLikeTactic__1___lambda__1___closed__64; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__24; static lean_object* l_Lean_Syntax_instCoeOutTSepArrayTSyntaxArray___closed__1; lean_object* l_Lean_Syntax_node6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Syntax_isCharLit_x3f___closed__2; @@ -542,21 +557,21 @@ static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__92; static lean_object* l_Lean_TSyntax_expandInterpolatedStrChunks___closed__1; LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeIdentTSyntaxConsSyntaxNodeKindMkStr4Nil(lean_object*); -LEAN_EXPORT uint8_t l___private_Init_Meta_0__Lean_Meta_DSimp_beqConfig____x40_Init_Meta___hyg_12794_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Name_reprPrec___boxed(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_Simp_Config_contextual___default; LEAN_EXPORT lean_object* l_Lean_mkHole(lean_object*, uint8_t); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__102; LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeDepTermMkConsSyntaxNodeKindMkStr1NilIdentIdent(lean_object*, lean_object*, lean_object*, 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___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__18; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeBinLitAux___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__13; -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeHexDigit___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__20; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__95; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__7; static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__11; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__5; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__23; LEAN_EXPORT lean_object* l_Lean_Syntax_mkStrLit(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_instInhabitedConfig___closed__1; @@ -564,32 +579,40 @@ static lean_object* l_Lean_Syntax_instCoeIdentTSyntaxConsSyntaxNodeKindMkStr4Nil LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_mkSepArray___spec__1___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_DSimp_Config_etaStruct___default; LEAN_EXPORT uint32_t l_Lean_TSyntax_getChar(lean_object*); +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__1; LEAN_EXPORT uint8_t l_Lean_Meta_Simp_Config_memoize___default; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__85; static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__7; +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__6; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__130; LEAN_EXPORT lean_object* l_Lean_versionString; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__9; extern lean_object* l_Lean_Parser_Tactic_rwRuleSeq; +static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__31; static lean_object* l_Lean_Parser_Tactic_tacticErw_______closed__7; static lean_object* l_Lean_Syntax_isCharLit_x3f___closed__1; static lean_object* l_Lean_termEval__prio_____closed__9; static lean_object* l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__2; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__13; extern lean_object* l_Lean_Parser_Tactic_simpLemma; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeScientificLitVal_x3f_decodeAfterExp(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); lean_object* l_Lean_Syntax_getKind(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__20; LEAN_EXPORT lean_object* l_Lean_mkCIdentFromRef___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__58; static lean_object* l_Lean_Parser_Tactic_tacticErw_______closed__2; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____closed__6; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____closed__8; LEAN_EXPORT lean_object* l_Lean_mkCIdentFrom(lean_object*, lean_object*, uint8_t); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__132; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__21; LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeOutSepArrayArraySyntax___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__2; +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__15; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__4; +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__12; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__26; 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_22652____closed__4; LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeTSyntaxArray(lean_object*, lean_object*, lean_object*, lean_object*); static uint8_t l_Lean_versionString___closed__2; static lean_object* l_Lean_Name_reprPrec___closed__9; @@ -598,15 +621,12 @@ static lean_object* l_Lean_Syntax_getHead_x3f___closed__3; static lean_object* l_Lean___aux__Init__Meta______macroRules__Lean__Parser__Syntax__addPrec__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeScientificLitVal_x3f_decodeAfterDot(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Name_beq_match__1_splitter___rarg___boxed(lean_object*, 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_13174____closed__16; static lean_object* l_Lean_Name_escapePart___closed__1; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__81; -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__22; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeDecimalLitAux(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_termEval__prec_____closed__6; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__56; -LEAN_EXPORT uint8_t l___private_Init_Meta_0__Lean_Meta_beqEtaStructMode____x40_Init_Meta___hyg_12541_(uint8_t, uint8_t); -LEAN_EXPORT lean_object* l_List_beq___at___private_Init_Meta_0__Lean_Syntax_beqPreresolved____x40_Init_Meta___hyg_3131____spec__1___boxed(lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__3; static lean_object* l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__9; LEAN_EXPORT lean_object* l_Lean_quoteNameMk(lean_object*); LEAN_EXPORT lean_object* l_Lean_mkGroupNode(lean_object*); @@ -614,27 +634,30 @@ LEAN_EXPORT lean_object* l_Array_isEqvAux___at_Lean_Syntax_structEq___spec__1___ static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__88; LEAN_EXPORT lean_object* l_Lean_Name_replacePrefix(lean_object*, lean_object*, lean_object*); 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_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__18; +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__13; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2377____closed__5; 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_NameGenerator_namePrefix___default___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__declareSimpLikeTactic__1___lambda__1___closed__131; -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_beqPreresolved____x40_Init_Meta___hyg_3131____boxed(lean_object*, lean_object*); static lean_object* l_Lean_Syntax_instBEqPreresolved___closed__1; 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_Name_appendIndexAfter___closed__1; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__70; +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__23; static lean_object* l_Lean_Parser_Tactic_dsimpKind___closed__1; -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__11; +static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__15; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__69; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____closed__7; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__4; uint8_t l_List_isEmpty___rarg(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__14; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__1; +static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__32; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic; static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__9; +static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__6; lean_object* l_Lean_Syntax_getOptional_x3f(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__8; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_simpAllArithAutoUnfold; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeInterpStrLit___boxed(lean_object*); uint8_t l_String_anyAux_loop(lean_object*, lean_object*, lean_object*, lean_object*); @@ -651,7 +674,7 @@ static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean lean_object* l_Lean_TSyntaxArray_mkImpl___rarg___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_getHead_x3f___lambda__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_isStrLit_x3f(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__20; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__9; static lean_object* l_Lean_Syntax_mkNumLit___closed__2; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__22; lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); @@ -663,11 +686,13 @@ static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean LEAN_EXPORT lean_object* l_Lean_Syntax_mkSep(lean_object*, lean_object*); static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2377____closed__16; LEAN_EXPORT lean_object* l_List_repr___at___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2377____spec__4___boxed(lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__21; LEAN_EXPORT lean_object* l_Array_filterSepElemsM___at_Array_filterSepElems___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withHeadRefOnly___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_decodeQuotedChar___boxed__const__3; +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_beqEtaStructMode____x40_Init_Meta___hyg_12715____boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____closed__3; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2377____closed__4; +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__22; static lean_object* l_Lean_Option_hasQuote___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_Syntax_getTrailingSize(lean_object*); uint8_t lean_string_utf8_at_end(lean_object*, lean_object*); @@ -678,7 +703,7 @@ static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_ LEAN_EXPORT lean_object* l_Lean_Syntax_SepArray_getElems___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__138; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__41; -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__15; +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Name_getRoot(lean_object*); static lean_object* l___private_Init_Meta_0__Lean_quoteList___rarg___closed__4; LEAN_EXPORT lean_object* l_Lean_Syntax_setTailInfo(lean_object*, lean_object*); @@ -687,18 +712,16 @@ LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeInterpStrLit 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 lean_object* l_Array_mapSepElemsM___at_Array_mapSepElems___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Syntax_instCoeTSyntaxArray___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Substring_toName(lean_object*); LEAN_EXPORT lean_object* l_Lean_withHeadRefOnly___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__127; static lean_object* l_Lean_mkOptionalNode___closed__1; static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__12; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__2; 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_Parser_Tactic_declareSimpLikeTactic___closed__18; static lean_object* l_Lean_Parser_Tactic_simpAllArithAutoUnfold___closed__7; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__5; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__11; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Name_beq_match__1_splitter(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__7; LEAN_EXPORT lean_object* l_Lean_Syntax_mkLit(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_getSepElems___rarg___boxed(lean_object*); LEAN_EXPORT lean_object* l_repr___at___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2377____spec__5(lean_object*); @@ -709,13 +732,11 @@ LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeTSyntaxArrayTSepArray(lean_object* static uint8_t l_Lean_versionString___closed__3; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__90; LEAN_EXPORT lean_object* l_Lean_Syntax_hasArgs___boxed(lean_object*); -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__7; static lean_object* l_Lean_Name_reprPrec___closed__2; -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_versionString___closed__10; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeQuotedChar___boxed__const__5; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__18; LEAN_EXPORT lean_object* l_Lean_mkFreshId___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__14; LEAN_EXPORT lean_object* l_Lean_isIdEndEscape___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__121; LEAN_EXPORT lean_object* l_Lean_Syntax_instBEqPreresolved; @@ -728,6 +749,7 @@ static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean LEAN_EXPORT lean_object* l_Lean_Syntax_decodeQuotedChar___boxed__const__2; LEAN_EXPORT lean_object* l_Lean_Syntax_getSepArgs(lean_object*); lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____boxed(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_Parser_Tactic_dsimpKind___closed__9; lean_object* l_Lean_extractMacroScopes(lean_object*); @@ -737,15 +759,12 @@ LEAN_EXPORT lean_object* l_Lean_Syntax_instBEqTSyntax___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeNumLitPrio(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__57; LEAN_EXPORT lean_object* l_Lean_Syntax_getHead_x3f(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__8; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__20; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__8; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__11; LEAN_EXPORT lean_object* l_Lean_Syntax_TSepArray_push___boxed(lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__5; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__109; -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__5; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__4; LEAN_EXPORT lean_object* l_Lean_Syntax_instBEqTSyntax(lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeIdentLevel(lean_object*); @@ -755,7 +774,6 @@ static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__9; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeQuotedChar___boxed__const__6; static lean_object* l_List_repr_x27___at___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2236____spec__1___closed__2; -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18276_(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Internal_isStage0___boxed(lean_object*); static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2377____closed__21; LEAN_EXPORT uint8_t l_Lean_version_getIsRelease(lean_object*); @@ -781,16 +799,17 @@ LEAN_EXPORT lean_object* l_Lean_Syntax_getTailInfo(lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_Simp_Config_zeta___default; static lean_object* l_Lean_toolchain___closed__8; lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____closed__4; lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_beqConfig____x40_Init_Meta___hyg_12794____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Name_escapePart___lambda__1___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20990_(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__51; -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__17; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__15; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__78; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__5; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__2; +static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__23; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__28; static lean_object* l_Lean_mkSepArray___closed__2; +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652_(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__123; static lean_object* l_List_repr_x27___at___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2236____spec__1___closed__1; static lean_object* l_Lean_Parser_Tactic_simpAllArithAutoUnfold___closed__3; @@ -800,6 +819,7 @@ LEAN_EXPORT lean_object* l_Lean_isIdBeginEscape___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__19; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteArray_go___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_findSomeRevM_x3f_find___at_Lean_Syntax_getTailInfo_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__9; static lean_object* l_Lean_Parser_Tactic_dsimpKind___closed__3; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__61; LEAN_EXPORT lean_object* l_Array_mapSepElemsM(lean_object*); @@ -810,7 +830,6 @@ LEAN_EXPORT lean_object* l___private_Init_Data_String_Basic_0__Substring_takeWhi static lean_object* l_Lean_Parser_Tactic_tacticErw_______closed__9; LEAN_EXPORT lean_object* l_Lean_version_getSpecialDesc(lean_object*); LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2377____spec__2(lean_object*, lean_object*); -static lean_object* l_List_foldr___at_Lean_Syntax_decodeNameLit___spec__1___closed__3; static lean_object* l_Lean_Option_hasQuote___rarg___closed__2; LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeScientificLitTerm(lean_object*); LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeTSyntaxConsSyntaxNodeKind___rarg___boxed(lean_object*); @@ -821,6 +840,7 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Array_getSepElems___spec static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__52; LEAN_EXPORT lean_object* l_Lean_instQuoteSubstringMkStr1(lean_object*); LEAN_EXPORT lean_object* l_Lean_instQuoteNameMkStr1(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__19; LEAN_EXPORT lean_object* lean_mk_syntax_ident(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__41; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__56; @@ -829,9 +849,12 @@ static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__22; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__46; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2377____closed__6; LEAN_EXPORT lean_object* l_Lean_Name_capitalize(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__6; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__14; static lean_object* l_Lean_versionString___closed__5; static lean_object* l_Lean_instQuoteBoolMkStr1___closed__3; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__77; +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__10; static lean_object* l_Lean_Syntax_isInterpolatedStrLit_x3f___closed__1; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__62; static lean_object* l_Lean_toolchain___closed__3; @@ -839,38 +862,37 @@ static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprTSyntax____x40_Init_Meta___hyg_2642____rarg___closed__9; LEAN_EXPORT lean_object* l_Lean_Syntax_isToken___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_filterSepElems(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____closed__5; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_version_getMinor___boxed(lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360_(uint8_t, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__3; +LEAN_EXPORT lean_object* l_Lean_TSyntax_getHygieneInfo(lean_object*); lean_object* l_Lean_MacroScopesView_review(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_simpAllAutoUnfold; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeQuotedChar___boxed__const__1; static lean_object* l_Lean_Syntax_instCoeOutSepArrayArraySyntax___closed__1; -static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__5; LEAN_EXPORT lean_object* l_Lean_Syntax_isNone___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_mkIdentFromRef(lean_object*); static lean_object* l_Lean_Parser_Tactic_dsimpAutoUnfold___closed__10; +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__9; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__62; LEAN_EXPORT uint8_t l_Lean_Meta_DSimp_Config_autoUnfold___default; static lean_object* l_Lean_Syntax_instReprSyntax___closed__1; LEAN_EXPORT lean_object* l_Lean_Name_isInaccessibleUserName___boxed(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__2; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__2; -static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__9; +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__5; lean_object* lean_get_githash(lean_object*); static lean_object* l_Lean_instQuoteProdMkStr1___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_Syntax_isAtom___boxed(lean_object*); static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2377____closed__2; 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_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_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__8; lean_object* l_Array_append___rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____boxed(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__111; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__16; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__5; LEAN_EXPORT lean_object* l_Lean_version_major; static lean_object* l_Lean_Parser_Tactic_dsimpAutoUnfold___closed__8; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__104; -static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__14; static lean_object* l_List_repr_x27___at___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2236____spec__1___closed__9; LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeOutTSepArrayTSyntaxArray___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Syntax_mkScientificLit___closed__2; @@ -878,36 +900,38 @@ static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprTSyntax____x40_Init LEAN_EXPORT lean_object* l_Lean_Syntax_getSubstring_x3f(lean_object*, uint8_t, uint8_t); LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeIdentTerm___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic_tacticErw_______closed__3; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__1; static lean_object* l_Lean_mkHole___closed__2; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__47; 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___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__25; 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_simpAllAutoUnfold___closed__9; -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409_(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_versionString___closed__9; LEAN_EXPORT lean_object* l_Array_mapSepElems(lean_object*, lean_object*); extern lean_object* l_Lean_Parser_Tactic_simpErase; LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeTSyntaxConsSyntaxNodeKindNil___rarg(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__1; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__43; LEAN_EXPORT uint8_t l_Lean_Meta_Simp_Config_arith___default; uint8_t l_Char_isDigit(uint32_t); static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__2; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__6; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__83; -static lean_object* l_List_foldr___at_Lean_Syntax_decodeNameLit___spec__1___closed__4; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__35; static lean_object* l_Lean_Meta_DSimp_instBEqConfig___closed__1; LEAN_EXPORT lean_object* l_Lean_evalOptPrio(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_isLetterLike(uint32_t); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__10; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__10; lean_object* l_Lean_Syntax_node4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Substring_nextn(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__9; lean_object* l_Lean_SourceInfo_getPos_x3f(lean_object*, uint8_t); LEAN_EXPORT lean_object* lean_name_append_index_after(lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__7; +static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__11; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeStrLit___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_instInhabitedConfig; +static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__29; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__82; LEAN_EXPORT lean_object* l_Array_mapSepElemsM___at_Array_mapSepElems___spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__76; @@ -926,14 +950,17 @@ LEAN_EXPORT lean_object* l_Lean_TSyntax_getId___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_isNameLit_x3f(lean_object*); static lean_object* l_Lean_Meta_instBEqEtaStructMode___closed__1; static lean_object* l_Lean_Syntax_getHead_x3f___closed__2; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18519____closed__2; static lean_object* l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__2; static lean_object* l_Lean_Name_reprPrec___closed__10; 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_toolchain___closed__9; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__74; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeOctalLitAux(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____closed__1; LEAN_EXPORT uint8_t l___private_Init_Meta_0__Lean_Name_hasNum(lean_object*); static lean_object* l_Lean_Syntax_mkStrLit___closed__2; +static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__26; 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*); static lean_object* l_Lean_instInhabitedNameGenerator___closed__1; static lean_object* l_Lean_quoteNameMk___closed__2; @@ -946,10 +973,12 @@ static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean LEAN_EXPORT lean_object* l_Lean_mkCIdentFromRef___rarg___lambda__1(lean_object*, lean_object*, uint8_t, lean_object*); static lean_object* l_Lean_Parser_Tactic_dsimpKind___closed__7; static lean_object* l_Lean_toolchain___closed__2; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__6; LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeTSyntaxArrayTSepArray___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_toNat___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_EtaStructMode_toCtorIdx(uint8_t); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__101; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__12; uint8_t l_Lean_Name_hasMacroScopes(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_getTailInfo_x3f(lean_object*); LEAN_EXPORT uint32_t l_Lean_idEndEscape; @@ -962,70 +991,70 @@ LEAN_EXPORT uint8_t l_Lean_isNumericSubscript(uint32_t); static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprTSyntax____x40_Init_Meta___hyg_2642____rarg___closed__13; 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_Name_replacePrefix___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21779_(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__99; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__2; static lean_object* l_Lean_Name_reprPrec___closed__3; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__16; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeOctalLitAux___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__8; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__66; 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_simpAutoUnfold___closed__15; LEAN_EXPORT uint8_t l_Lean_Name_toString_maybePseudoSyntax(lean_object*); static lean_object* l_Lean_instQuoteSubstringMkStr1___closed__3; -static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__11; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__24; +static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__2; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprTSyntax____x40_Init_Meta___hyg_2642____rarg___closed__5; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Internal_isStage0(lean_object*); 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__declareSimpLikeTactic__1___closed__28; LEAN_EXPORT lean_object* l_Lean_Syntax_getSubstring_x3f___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__3; LEAN_EXPORT lean_object* l_Lean_mkIdentFromRef___rarg(lean_object*, lean_object*, lean_object*, uint8_t); static lean_object* l_Lean_Parser_Tactic_simpAllArith___closed__2; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteArray(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20747_(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__14; +static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__30; static lean_object* l_Lean_Parser_Tactic_tacticErw_______closed__1; LEAN_EXPORT lean_object* l_Lean_TSyntax_getScientific(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__4; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__20; -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_beqConfig____x40_Init_Meta___hyg_13366____boxed(lean_object*, lean_object*); lean_object* lean_string_length(lean_object*); static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2377____closed__7; LEAN_EXPORT lean_object* l_Lean_termEval__prio__; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Name_hasNum___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_getTailInfo_x3f___boxed(lean_object*); static lean_object* l_Lean_Name_toStringWithSep___closed__1; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__4; LEAN_EXPORT uint8_t l_Lean_Meta_Simp_Config_dsimp___default; LEAN_EXPORT lean_object* l_Lean_Syntax_getOptionalIdent_x3f(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__declareSimpLikeTactic__1___lambda__1___closed__30; static lean_object* l_Lean_Name_isInaccessibleUserName___closed__1; +static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__20; LEAN_EXPORT uint8_t l_Lean_Syntax_hasArgs(lean_object*); LEAN_EXPORT lean_object* l_Lean_TSyntax_getNat(lean_object*); -static lean_object* l_List_foldr___at_Lean_Syntax_decodeNameLit___spec__1___closed__2; -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__8; static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__15; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__4; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__7; +static lean_object* l_List_foldr___at_Substring_toName___spec__1___closed__3; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeInterpStrLit(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__54; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__10; -static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__16; LEAN_EXPORT lean_object* l_Lean_Syntax_SepArray_ofElems(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__28; static lean_object* l_Lean_termEval__prio_____closed__8; LEAN_EXPORT uint8_t l_Lean_isSubScriptAlnum(uint32_t); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__15; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__68; +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__26; LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeTSyntaxConsSyntaxNodeKind(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__95; LEAN_EXPORT lean_object* l_Lean_Syntax_isLit_x3f(lean_object*, lean_object*); lean_object* lean_nat_mod(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__42; static lean_object* l_Lean_instQuoteBoolMkStr1___closed__6; -static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__19; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__6; static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__14; LEAN_EXPORT lean_object* l_Lean_TSyntax_Compat_instCoeTailSyntaxTSyntax___rarg(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__4; @@ -1040,14 +1069,12 @@ LEAN_EXPORT lean_object* l_Lean_Meta_instBEqEtaStructMode; lean_object* l_id___rarg___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__15; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____closed__1; LEAN_EXPORT lean_object* l_Lean_mkIdentFromRef___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__20; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__120; LEAN_EXPORT lean_object* l_Array_findSomeRevM_x3f_find___at_Lean_Syntax_getTailInfo_x3f___spec__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__79; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2236____closed__6; -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__23; static lean_object* l_Lean_Syntax_isInterpolatedStrLit_x3f___closed__2; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__40; 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*); @@ -1061,38 +1088,33 @@ static lean_object* l_Lean_toolchain___closed__7; LEAN_EXPORT lean_object* l_Lean_Syntax_SepArray_getElems(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___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2236____closed__8; +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteArray___rarg(lean_object*, lean_object*); lean_object* l_Lean_Syntax_node7(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_simpAutoUnfold___closed__17; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__38; -LEAN_EXPORT lean_object* l_List_foldr___at_Lean_Syntax_decodeNameLit___spec__1___boxed(lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__7; static lean_object* l_Lean_Parser_Tactic_simpAllAutoUnfold___closed__6; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__5; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__67; LEAN_EXPORT lean_object* l_Lean_version_patch; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeInterpStrQuotedChar___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Syntax_getHead_x3f___closed__1; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__12; static lean_object* l_Lean_Parser_Tactic_dsimpKind___closed__2; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__60; static uint8_t l_Lean_version_isRelease___closed__1; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__93; -LEAN_EXPORT uint8_t l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_12336_(uint8_t, uint8_t); static lean_object* l_Lean_Parser_Tactic_simpAllAutoUnfold___closed__13; static lean_object* l_Lean_Syntax_instCoeIdentTSyntaxConsSyntaxNodeKindMkStr4Nil___closed__1; lean_object* l_Lean_Syntax_node1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllArith___closed__8; static lean_object* l_Lean_versionString___closed__11; -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__19; LEAN_EXPORT lean_object* l_Lean_Syntax_findAux(lean_object*, lean_object*); static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprTSyntax____x40_Init_Meta___hyg_2642____rarg___closed__3; +static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__24; LEAN_EXPORT uint8_t l_Lean_Meta_Simp_Config_proj___default; LEAN_EXPORT lean_object* l_Lean_mkSepArray(lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__1; uint8_t l_String_isPrefixOf(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Syntax_instCoeTSyntaxArray___spec__1(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); -LEAN_EXPORT uint8_t l___private_Init_Meta_0__Lean_Syntax_beqPreresolved____x40_Init_Meta___hyg_3131_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Name_instToStringName(lean_object*); static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2377____closed__17; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__21; @@ -1110,6 +1132,8 @@ LEAN_EXPORT lean_object* l_Lean_Syntax_instBEqTSyntax___rarg___boxed(lean_object LEAN_EXPORT lean_object* l_Lean_TSyntax_Compat_instCoeTailSyntaxTSyntax___rarg___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeTSyntaxArray___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean___aux__Init__Meta______macroRules__Lean__Parser__Syntax__addPrio__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_String_toName(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__1; static lean_object* l_List_repr_x27___at___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2236____spec__1___closed__11; lean_object* l_Repr_addAppParen(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__13; @@ -1120,8 +1144,10 @@ static lean_object* l_Lean_quoteNameMk___closed__1; LEAN_EXPORT lean_object* l_Lean_NameGenerator_mkChild(lean_object*); LEAN_EXPORT uint8_t l_Lean_Syntax_isNone(lean_object*); LEAN_EXPORT lean_object* l_Lean_NameGenerator_namePrefix___default; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__20; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeScientificLitVal_x3f_decodeExp(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__32; +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__3; LEAN_EXPORT lean_object* l_Lean_Name_getRoot___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__8; LEAN_EXPORT lean_object* l_Lean_TSyntax_Compat_instCoeTailArraySyntaxTSyntaxArray(lean_object*); @@ -1132,17 +1158,18 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1 static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__44; static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__1; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_reprTSyntax____x40_Init_Meta___hyg_2642____boxed(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__8; LEAN_EXPORT uint8_t l_Lean_isIdEndEscape(uint32_t); LEAN_EXPORT lean_object* l_Lean_mkFreshId___rarg___lambda__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllArithAutoUnfold___closed__2; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeInterpStrLit_loop___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_setInfo(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_TSyntax_getName___boxed(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__7; lean_object* lean_nat_sub(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_setHeadInfoAux(lean_object*, lean_object*); static lean_object* l_Lean_Syntax_mkNameLit___closed__1; LEAN_EXPORT lean_object* l_Lean_Syntax_TSepArray_push(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__4; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__66; LEAN_EXPORT lean_object* l_Lean_withHeadRefOnly(lean_object*); static lean_object* l_Lean_mkSepArray___closed__1; @@ -1156,12 +1183,10 @@ LEAN_EXPORT lean_object* l_List_foldl___at___private_Init_Meta_0__Lean_Syntax_re static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__98; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__75; static lean_object* l___private_Init_Meta_0__Lean_quoteArray___rarg___closed__2; -static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__13; -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__28; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__22; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeNumLitPrec___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__18; -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__10; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_simpArith; 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_declareSimpLikeTactic___closed__6; @@ -1184,40 +1209,39 @@ static lean_object* l_Lean_termEval__prec_____closed__5; static lean_object* l_Lean_githash___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_DSimp_instReprConfig; LEAN_EXPORT uint8_t l_Lean_Meta_DSimp_Config_decide___default; -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeNameLitTerm(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_tacticErw____; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__11; +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteList___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Name_reprPrec___closed__4; -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__26; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__18; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Array_filterSepElemsMAux(lean_object*); +lean_object* lean_erase_macro_scopes(lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__13; LEAN_EXPORT uint8_t lean_is_inaccessible_user_name(lean_object*); LEAN_EXPORT lean_object* l_Lean_NameGenerator_next(lean_object*); LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeCharLitTerm(lean_object*); -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__3; static lean_object* l_Lean_versionStringCore___closed__7; LEAN_EXPORT lean_object* l_Lean_Meta_instBEqTransparencyMode; LEAN_EXPORT uint8_t l_Lean_Meta_Simp_Config_eta___default; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__29; LEAN_EXPORT lean_object* l_Lean_instQuote___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_List_reverse___rarg(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__16; +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174_(lean_object*, lean_object*); static lean_object* l_Lean_instQuoteBoolMkStr1___closed__5; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____closed__3; static lean_object* l_Lean_instQuoteNameMkStr1___closed__2; +LEAN_EXPORT lean_object* l_List_beq___at___private_Init_Meta_0__Lean_Syntax_beqPreresolved____x40_Init_Meta___hyg_3135____spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__4; +static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__1; static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__4; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__84; LEAN_EXPORT uint8_t l_Lean_Meta_DSimp_Config_beta___default; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_Config_maxSteps___default; static lean_object* l_Lean_Syntax_unsetTrailing___closed__2; lean_object* l_String_intercalate(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__3; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__19; -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565_(uint8_t, 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 uint8_t l_Lean_Syntax_structEq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Name_beq_match__1_splitter___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1229,7 +1253,6 @@ static lean_object* l___private_Init_Meta_0__Lean_quoteArray___rarg___closed__1; static lean_object* l_Lean_Parser_Tactic_tacticErw_______closed__10; LEAN_EXPORT lean_object* l_Lean_Meta_DSimp_instInhabitedConfig; LEAN_EXPORT lean_object* l_Lean_instQuoteListMkStr1___rarg(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__20; static lean_object* l_Lean_Parser_Tactic_dsimpAutoUnfold___closed__7; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__103; LEAN_EXPORT lean_object* l_Lean_TSyntax_expandInterpolatedStr___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); @@ -1239,50 +1262,45 @@ static lean_object* l_Lean_Meta_instReprEtaStructMode___closed__1; uint8_t l_Substring_beq(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__89; lean_object* l_String_capitalize(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__4; -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__27; size_t lean_usize_add(size_t, size_t); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__17; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__1; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprTSyntax____x40_Init_Meta___hyg_2642____rarg___closed__11; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_defaultMaxSteps; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__8; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__18; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__19; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__29; static lean_object* l_Array_forInUnsafe_loop___at_Lean_TSyntax_expandInterpolatedStrChunks___spec__1___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_TSyntax_getScientific___boxed(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__declareSimpLikeTactic__1___lambda__1___closed__92; -LEAN_EXPORT uint8_t l___private_Init_Meta_0__Lean_Meta_Simp_beqConfig____x40_Init_Meta___hyg_13366_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Option_hasQuote___rarg(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_Simp_Config_iota___default; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__24; lean_object* l_Lean_Macro_expandMacro_x3f(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_mkCIdentFrom___closed__1; -static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__2; static lean_object* l___private_Init_Meta_0__Lean_quoteList___rarg___closed__6; LEAN_EXPORT lean_object* l_Lean_TSyntax_getNat___boxed(lean_object*); static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprTSyntax____x40_Init_Meta___hyg_2642____rarg___closed__10; static lean_object* l_Lean_Meta_instReprTransparencyMode___closed__1; +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739_(uint8_t, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__42; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteList(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__54; static lean_object* l_Lean_Parser_Tactic_simpArith___closed__6; LEAN_EXPORT lean_object* l_List_foldl___at___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2377____spec__3(lean_object*, lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__10; -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__24; LEAN_EXPORT lean_object* l_Lean_version_getSpecialDesc___boxed(lean_object*); static lean_object* l_Lean_mkOptionalNode___closed__3; static lean_object* l_Lean_termEval__prec_____closed__10; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__97; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__53; -static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__23; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__36; LEAN_EXPORT lean_object* l_Array_filterSepElemsM___at_Array_filterSepElems___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__116; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__16; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__16; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__5; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__8; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__12; extern lean_object* l_Lean_Parser_Tactic_config; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_instBEqConfig; LEAN_EXPORT lean_object* l_Lean_mkSepArray___boxed(lean_object*, lean_object*); @@ -1291,36 +1309,40 @@ static lean_object* l_Lean_versionStringCore___closed__1; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__115; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_version_getMinor(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__22; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__13; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2377____closed__1; static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__17; lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_unsetTrailing(lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__24; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__51; LEAN_EXPORT lean_object* l_Lean_Syntax_mkCApp(lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__5; static lean_object* l_Lean_Parser_Tactic_simpAllArith___closed__4; LEAN_EXPORT lean_object* l_Lean_instQuoteProdMkStr1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkCIdent(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__13; static lean_object* l_Lean_quoteNameMk___closed__4; +LEAN_EXPORT lean_object* l_List_foldr___at_Substring_toName___spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__16; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__5; LEAN_EXPORT lean_object* l_Lean_Meta_TransparencyMode_toCtorIdx(uint8_t); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__55; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeCharLit___boxed(lean_object*); static lean_object* l_Lean_toolchain___closed__5; lean_object* l_Lean_Syntax_setKind(lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__2; static lean_object* l_Lean_Parser_Tactic_simpArith___closed__8; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__18; static lean_object* l_Lean_quoteNameMk___closed__3; 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_EXPORT lean_object* l_Lean_TSyntax_getChar___boxed(lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteArray_go___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958_(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_SepArray_ofElemsUsingRef(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_decodeStrLitAux___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Name_reprPrec___closed__8; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_instReprConfig; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__17; lean_object* l_Lean_mkAtomFrom(lean_object*, lean_object*, uint8_t); static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__26; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeNatLitVal_x3f___boxed(lean_object*); @@ -1329,43 +1351,37 @@ static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean LEAN_EXPORT lean_object* l_Lean___aux__Init__Meta______macroRules__Lean__termEval__prec____1(lean_object*, lean_object*, lean_object*); lean_object* l_Nat_pred(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__128; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____closed__6; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__3; 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_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__20; -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__6; lean_object* lean_array_get_size(lean_object*); +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__2; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__17; lean_object* l_panic___at___private_Init_Prelude_0__Lean_assembleParts___spec__1(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__13; LEAN_EXPORT lean_object* l_Lean_Meta_EtaStructMode_toCtorIdx___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__34; -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_isNatLit_x3f(lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__10; static lean_object* l_Lean_instQuoteTermMkStr1___closed__1; static lean_object* l___private_Init_Meta_0__Lean_quoteList___rarg___closed__2; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2377____closed__22; -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__14; LEAN_EXPORT lean_object* l_Array_mapSepElemsM___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_instReprTSyntax___boxed(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_Syntax_splitNameLitAux___closed__1; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__72; -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Name_instDecidableEqName___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllArith___closed__9; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__31; static lean_object* l_Lean_mkGroupNode___closed__1; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__12; LEAN_EXPORT lean_object* l_Lean_Syntax_mkSep___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__49; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__93; uint8_t lean_nat_dec_le(lean_object*, lean_object*); static lean_object* l_Lean_Syntax_mkScientificLit___closed__1; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__9; LEAN_EXPORT lean_object* l_List_foldl___at___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2236____spec__4___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__4; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__6; static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__9; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__33; uint8_t lean_usize_dec_lt(size_t, size_t); @@ -1378,17 +1394,18 @@ static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean static lean_object* l_Lean_Meta_DSimp_instInhabitedConfig___closed__1; LEAN_EXPORT lean_object* l_Lean_TSyntax_getString___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Name_toString_maybePseudoSyntax___boxed(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__16; static lean_object* l_Lean_evalPrio___closed__1; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__14; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__30; LEAN_EXPORT lean_object* l_Lean_Name_toStringWithSep(lean_object*, uint8_t, lean_object*); static lean_object* l___private_Init_Meta_0__Lean_quoteArray_go___rarg___closed__1; static lean_object* l_Lean_Parser_Tactic_simpArith___closed__2; +LEAN_EXPORT lean_object* l_Lean_HygieneInfo_mkIdent(lean_object*, lean_object*, uint8_t); static lean_object* l_Lean_toolchain___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1(lean_object*, lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeBinLitAux(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_beqConfig____x40_Init_Meta___hyg_13540____boxed(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__21; LEAN_EXPORT lean_object* l_Lean_version_getIsRelease___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeTermTSyntaxConsSyntaxNodeKindMkStr4Nil(lean_object*); @@ -1406,21 +1423,21 @@ static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean LEAN_EXPORT lean_object* l_Lean_Syntax_mkScientificLit(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__12; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2377____closed__19; +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_simpAutoUnfold; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2377____closed__14; +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__4; static lean_object* l_Lean_termEval__prio_____closed__4; static lean_object* l_Lean_Option_hasQuote___rarg___closed__3; LEAN_EXPORT lean_object* l_Lean_mkHole___boxed(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_Simp_Config_beta___default; static lean_object* l_Lean_Parser_Tactic_simpAllArithAutoUnfold___closed__5; -static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__2; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__135; LEAN_EXPORT lean_object* l_Lean_expandMacros(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Name_modifyBase(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_Array_forInUnsafe_loop___at_Lean_mkSepArray___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_isIdFirst(uint32_t); -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__6; static lean_object* l_Lean___aux__Init__Meta______macroRules__Lean__Parser__Syntax__addPrio__1___closed__2; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__17; LEAN_EXPORT lean_object* l_Lean_version_specialDesc; @@ -1441,7 +1458,6 @@ static lean_object* l_Lean_Parser_Tactic_simpAllArith___closed__5; LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeOutSepArrayArraySyntax(lean_object*); LEAN_EXPORT lean_object* l_Lean_TSyntax_Compat_instCoeTailArraySyntaxTSepArray___rarg___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__43; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__22; LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeIdentLevel___boxed(lean_object*); lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); lean_object* l_String_drop(lean_object*, lean_object*); @@ -1451,21 +1467,20 @@ static lean_object* l_Lean_termEval__prec_____closed__11; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__38; LEAN_EXPORT lean_object* l_Lean_isIdFirst___boxed(lean_object*); 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*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__1; +static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__16; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__39; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprTSyntax____x40_Init_Meta___hyg_2642____rarg___closed__8; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__17; lean_object* l_Nat_repr(lean_object*); +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____closed__2; LEAN_EXPORT lean_object* l_Lean_Syntax_toNat(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_dsimpAutoUnfold; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__8; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____closed__2; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_isNatLitAux(lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__23; LEAN_EXPORT lean_object* l_List_repr_x27___at___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2236____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_monadNameGeneratorLift(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_dsimpKind___closed__8; static lean_object* l_Lean_toolchain___closed__1; -static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__1; +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__21; lean_object* l_Char_ofNat(lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllAutoUnfold___closed__1; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__133; @@ -1481,22 +1496,22 @@ LEAN_EXPORT lean_object* l_Lean_Syntax_mkNumLit(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpArith___closed__10; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__3; static lean_object* l___private_Init_Meta_0__Lean_quoteList___rarg___closed__5; -static lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__3; static lean_object* l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__1; static lean_object* l_Lean___aux__Init__Meta______macroRules__Lean__Parser__Syntax__subPrio__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_EtaStructMode_noConfusion___rarg(uint8_t, uint8_t, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__45; 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_versionStringCore___closed__6; +static lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__11; +LEAN_EXPORT lean_object* l_Lean_TSyntax_getHygieneInfo___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__14; -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_termEval__prec__; static lean_object* l_Lean_Syntax_mkApp___closed__1; +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716_(lean_object*, lean_object*, lean_object*); uint8_t l_Array_isEmpty___rarg(lean_object*); static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2236____closed__4; LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeTermTSyntaxConsSyntaxNodeKindMkStr4Nil___boxed(lean_object*); static lean_object* l_Lean_versionStringCore___closed__8; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__8; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_version_getMajor___boxed(lean_object* x_1) { _start: { @@ -6210,7 +6225,7 @@ lean_dec(x_1); return x_2; } } -LEAN_EXPORT uint8_t l_List_beq___at___private_Init_Meta_0__Lean_Syntax_beqPreresolved____x40_Init_Meta___hyg_3131____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l_List_beq___at___private_Init_Meta_0__Lean_Syntax_beqPreresolved____x40_Init_Meta___hyg_3135____spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -6260,7 +6275,7 @@ goto _start; } } } -LEAN_EXPORT uint8_t l___private_Init_Meta_0__Lean_Syntax_beqPreresolved____x40_Init_Meta___hyg_3131_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Init_Meta_0__Lean_Syntax_beqPreresolved____x40_Init_Meta___hyg_3135_(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -6305,29 +6320,29 @@ return x_13; else { uint8_t x_14; -x_14 = l_List_beq___at___private_Init_Meta_0__Lean_Syntax_beqPreresolved____x40_Init_Meta___hyg_3131____spec__1(x_9, x_11); +x_14 = l_List_beq___at___private_Init_Meta_0__Lean_Syntax_beqPreresolved____x40_Init_Meta___hyg_3135____spec__1(x_9, x_11); return x_14; } } } } } -LEAN_EXPORT lean_object* l_List_beq___at___private_Init_Meta_0__Lean_Syntax_beqPreresolved____x40_Init_Meta___hyg_3131____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_List_beq___at___private_Init_Meta_0__Lean_Syntax_beqPreresolved____x40_Init_Meta___hyg_3135____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l_List_beq___at___private_Init_Meta_0__Lean_Syntax_beqPreresolved____x40_Init_Meta___hyg_3131____spec__1(x_1, x_2); +x_3 = l_List_beq___at___private_Init_Meta_0__Lean_Syntax_beqPreresolved____x40_Init_Meta___hyg_3135____spec__1(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___private_Init_Meta_0__Lean_Syntax_beqPreresolved____x40_Init_Meta___hyg_3131____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_beqPreresolved____x40_Init_Meta___hyg_3135____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_Syntax_beqPreresolved____x40_Init_Meta___hyg_3131_(x_1, x_2); +x_3 = l___private_Init_Meta_0__Lean_Syntax_beqPreresolved____x40_Init_Meta___hyg_3135_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -6338,7 +6353,7 @@ static lean_object* _init_l_Lean_Syntax_instBEqPreresolved___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Init_Meta_0__Lean_Syntax_beqPreresolved____x40_Init_Meta___hyg_3131____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Init_Meta_0__Lean_Syntax_beqPreresolved____x40_Init_Meta___hyg_3135____boxed), 2, 0); return x_1; } } @@ -6423,7 +6438,7 @@ 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 = l___private_Init_Meta_0__Lean_Syntax_beqPreresolved____x40_Init_Meta___hyg_3131_(x_6, x_8); +x_10 = l___private_Init_Meta_0__Lean_Syntax_beqPreresolved____x40_Init_Meta___hyg_3135_(x_6, x_8); if (x_10 == 0) { uint8_t x_11; @@ -13013,7 +13028,7 @@ x_4 = l_List_reverse___rarg(x_3); return x_4; } } -static lean_object* _init_l_List_foldr___at_Lean_Syntax_decodeNameLit___spec__1___closed__1() { +static lean_object* _init_l_List_foldr___at_Substring_toName___spec__1___closed__1() { _start: { lean_object* x_1; @@ -13021,15 +13036,15 @@ x_1 = lean_mk_string_from_bytes("Init.Meta", 9); return x_1; } } -static lean_object* _init_l_List_foldr___at_Lean_Syntax_decodeNameLit___spec__1___closed__2() { +static lean_object* _init_l_List_foldr___at_Substring_toName___spec__1___closed__2() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Lean.Syntax.decodeNameLit", 25); +x_1 = lean_mk_string_from_bytes("Substring.toName", 16); return x_1; } } -static lean_object* _init_l_List_foldr___at_Lean_Syntax_decodeNameLit___spec__1___closed__3() { +static lean_object* _init_l_List_foldr___at_Substring_toName___spec__1___closed__3() { _start: { lean_object* x_1; @@ -13037,20 +13052,20 @@ x_1 = lean_mk_string_from_bytes("unreachable code has been reached", 33); return x_1; } } -static lean_object* _init_l_List_foldr___at_Lean_Syntax_decodeNameLit___spec__1___closed__4() { +static lean_object* _init_l_List_foldr___at_Substring_toName___spec__1___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_List_foldr___at_Lean_Syntax_decodeNameLit___spec__1___closed__1; -x_2 = l_List_foldr___at_Lean_Syntax_decodeNameLit___spec__1___closed__2; +x_1 = l_List_foldr___at_Substring_toName___spec__1___closed__1; +x_2 = l_List_foldr___at_Substring_toName___spec__1___closed__2; x_3 = lean_unsigned_to_nat(852u); -x_4 = lean_unsigned_to_nat(12u); -x_5 = l_List_foldr___at_Lean_Syntax_decodeNameLit___spec__1___closed__3; +x_4 = lean_unsigned_to_nat(10u); +x_5 = l_List_foldr___at_Substring_toName___spec__1___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_List_foldr___at_Lean_Syntax_decodeNameLit___spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_List_foldr___at_Substring_toName___spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -13063,7 +13078,7 @@ else lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint32_t x_11; uint32_t x_12; uint8_t x_13; x_3 = lean_ctor_get(x_2, 0); x_4 = lean_ctor_get(x_2, 1); -x_5 = l_List_foldr___at_Lean_Syntax_decodeNameLit___spec__1(x_1, x_4); +x_5 = l_List_foldr___at_Substring_toName___spec__1(x_1, x_4); x_6 = lean_ctor_get(x_3, 0); x_7 = lean_ctor_get(x_3, 1); x_8 = lean_ctor_get(x_3, 2); @@ -13091,7 +13106,7 @@ if (lean_obj_tag(x_16) == 0) { lean_object* x_17; lean_object* x_18; lean_dec(x_5); -x_17 = l_List_foldr___at_Lean_Syntax_decodeNameLit___spec__1___closed__4; +x_17 = l_List_foldr___at_Substring_toName___spec__1___closed__4; x_18 = l_panic___at___private_Init_Prelude_0__Lean_assembleParts___spec__1(x_17); return x_18; } @@ -13118,6 +13133,52 @@ return x_24; } } } +LEAN_EXPORT lean_object* l_Substring_toName(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = lean_box(0); +x_3 = l___private_Init_Meta_0__Lean_Syntax_splitNameLitAux(x_1, x_2); +if (lean_obj_tag(x_3) == 0) +{ +lean_object* x_4; +x_4 = lean_box(0); +return x_4; +} +else +{ +lean_object* x_5; lean_object* x_6; +x_5 = lean_box(0); +x_6 = l_List_foldr___at_Substring_toName___spec__1(x_5, x_3); +lean_dec(x_3); +return x_6; +} +} +} +LEAN_EXPORT lean_object* l_List_foldr___at_Substring_toName___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_List_foldr___at_Substring_toName___spec__1(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_String_toName(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_2 = lean_string_utf8_byte_size(x_1); +x_3 = lean_unsigned_to_nat(0u); +x_4 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_3); +lean_ctor_set(x_4, 2, x_2); +x_5 = l_Substring_toName(x_4); +return x_5; +} +} LEAN_EXPORT lean_object* l_Lean_Syntax_decodeNameLit(lean_object* x_1) { _start: { @@ -13135,7 +13196,7 @@ 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; lean_object* x_13; lean_object* x_14; +lean_object* x_7; 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_7 = lean_string_utf8_byte_size(x_1); lean_inc(x_7); lean_inc(x_1); @@ -13152,35 +13213,21 @@ x_12 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_12, 0, x_1); lean_ctor_set(x_12, 1, x_11); lean_ctor_set(x_12, 2, x_7); -x_13 = lean_box(0); -x_14 = l___private_Init_Meta_0__Lean_Syntax_splitNameLitAux(x_12, x_13); -if (lean_obj_tag(x_14) == 0) +x_13 = l_Substring_toName(x_12); +if (lean_obj_tag(x_13) == 0) { -lean_object* x_15; -x_15 = lean_box(0); -return x_15; +lean_object* x_14; +x_14 = lean_box(0); +return x_14; } else { -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_box(0); -x_17 = l_List_foldr___at_Lean_Syntax_decodeNameLit___spec__1(x_16, x_14); -lean_dec(x_14); -x_18 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_18, 0, x_17); -return x_18; -} -} +lean_object* x_15; +x_15 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_15, 0, x_13); +return x_15; } } -LEAN_EXPORT lean_object* l_List_foldr___at_Lean_Syntax_decodeNameLit___spec__1___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = l_List_foldr___at_Lean_Syntax_decodeNameLit___spec__1(x_1, x_2); -lean_dec(x_2); -lean_dec(x_1); -return x_3; } } LEAN_EXPORT lean_object* l_Lean_Syntax_isNameLit_x3f(lean_object* x_1) { @@ -13766,6 +13813,26 @@ lean_dec(x_1); return x_2; } } +LEAN_EXPORT lean_object* l_Lean_TSyntax_getHygieneInfo(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_2 = lean_unsigned_to_nat(0u); +x_3 = l_Lean_Syntax_getArg(x_1, x_2); +x_4 = l_Lean_Syntax_getId(x_3); +lean_dec(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_TSyntax_getHygieneInfo___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_TSyntax_getHygieneInfo(x_1); +lean_dec(x_1); +return x_2; +} +} LEAN_EXPORT lean_object* l_Lean_TSyntax_Compat_instCoeTailArraySyntaxTSepArray___rarg(lean_object* x_1, lean_object* x_2) { _start: { @@ -13804,6 +13871,85 @@ lean_dec(x_1); return x_2; } } +LEAN_EXPORT lean_object* l_Lean_HygieneInfo_mkIdent(lean_object* x_1, lean_object* x_2, uint8_t 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 = l_Lean_Syntax_getId(x_5); +x_7 = l_Lean_extractMacroScopes(x_6); +lean_inc(x_2); +x_8 = lean_erase_macro_scopes(x_2); +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_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_10 = lean_ctor_get(x_7, 0); +lean_dec(x_10); +lean_ctor_set(x_7, 0, x_8); +x_11 = l_Lean_MacroScopesView_review(x_7); +x_12 = l_Lean_SourceInfo_fromRef(x_5, x_3); +x_13 = 1; +x_14 = l_Lean_Name_toString(x_2, x_13); +x_15 = lean_string_utf8_byte_size(x_14); +x_16 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_16, 0, x_14); +lean_ctor_set(x_16, 1, x_4); +lean_ctor_set(x_16, 2, x_15); +x_17 = lean_box(0); +x_18 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_18, 0, x_12); +lean_ctor_set(x_18, 1, x_16); +lean_ctor_set(x_18, 2, x_11); +lean_ctor_set(x_18, 3, x_17); +return x_18; +} +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; uint8_t 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_ctor_get(x_7, 1); +x_20 = lean_ctor_get(x_7, 2); +x_21 = lean_ctor_get(x_7, 3); +lean_inc(x_21); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_7); +x_22 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_22, 0, x_8); +lean_ctor_set(x_22, 1, x_19); +lean_ctor_set(x_22, 2, x_20); +lean_ctor_set(x_22, 3, x_21); +x_23 = l_Lean_MacroScopesView_review(x_22); +x_24 = l_Lean_SourceInfo_fromRef(x_5, x_3); +x_25 = 1; +x_26 = l_Lean_Name_toString(x_2, x_25); +x_27 = lean_string_utf8_byte_size(x_26); +x_28 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 1, x_4); +lean_ctor_set(x_28, 2, x_27); +x_29 = lean_box(0); +x_30 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_30, 0, x_24); +lean_ctor_set(x_30, 1, x_28); +lean_ctor_set(x_30, 2, x_23); +lean_ctor_set(x_30, 3, x_29); +return x_30; +} +} +} +LEAN_EXPORT lean_object* l_Lean_HygieneInfo_mkIdent___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_HygieneInfo_mkIdent(x_1, x_2, x_4); +lean_dec(x_1); +return x_5; +} +} LEAN_EXPORT lean_object* l_Lean_instQuote___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { @@ -18458,7 +18604,7 @@ x_1 = 0; return x_1; } } -LEAN_EXPORT uint8_t l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_12336_(uint8_t x_1, uint8_t x_2) { +LEAN_EXPORT uint8_t l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_12510_(uint8_t x_1, uint8_t x_2) { _start: { lean_object* x_3; lean_object* x_4; uint8_t x_5; @@ -18470,7 +18616,7 @@ lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_12336____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_12510____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; @@ -18478,7 +18624,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_12336_(x_3, x_4); +x_5 = l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_12510_(x_3, x_4); x_6 = lean_box(x_5); return x_6; } @@ -18487,7 +18633,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_12336____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_12510____boxed), 2, 0); return x_1; } } @@ -18499,7 +18645,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_12360____closed__1() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__1() { _start: { lean_object* x_1; @@ -18507,33 +18653,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_12360____closed__2() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____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_12360____closed__1; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__1; x_2 = lean_alloc_ctor(3, 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_12360____closed__3() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2236____closed__4; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__2; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____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_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__4() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____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_12360____closed__3; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__3; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -18541,23 +18687,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_12360____closed__5() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2236____closed__5; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__2; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____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_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__6() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____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_12360____closed__5; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__5; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -18565,7 +18711,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_12360____closed__7() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__7() { _start: { lean_object* x_1; @@ -18573,33 +18719,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_12360____closed__8() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____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_12360____closed__7; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__7; x_2 = lean_alloc_ctor(3, 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_12360____closed__9() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2236____closed__4; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__8; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__8; 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_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__10() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____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_12360____closed__9; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__9; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -18607,23 +18753,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_12360____closed__11() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2236____closed__5; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__8; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__8; 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_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__12() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____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_12360____closed__11; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__11; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -18631,7 +18777,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_12360____closed__13() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__13() { _start: { lean_object* x_1; @@ -18639,33 +18785,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_12360____closed__14() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____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_12360____closed__13; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__13; x_2 = lean_alloc_ctor(3, 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_12360____closed__15() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2236____closed__4; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__14; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__14; 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_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__16() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____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_12360____closed__15; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__15; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -18673,23 +18819,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_12360____closed__17() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2236____closed__5; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__14; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__14; 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_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__18() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____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_12360____closed__17; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__17; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -18697,7 +18843,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_12360____closed__19() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__19() { _start: { lean_object* x_1; @@ -18705,33 +18851,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_12360____closed__20() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____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_12360____closed__19; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__19; x_2 = lean_alloc_ctor(3, 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_12360____closed__21() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2236____closed__4; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__20; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__20; 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_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__22() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____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_12360____closed__21; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__21; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -18739,23 +18885,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_12360____closed__23() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2236____closed__5; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__20; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__20; 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_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__24() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____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_12360____closed__23; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__23; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -18763,7 +18909,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_12360_(uint8_t x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534_(uint8_t x_1, lean_object* x_2) { _start: { switch (x_1) { @@ -18775,14 +18921,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_12360____closed__4; +x_5 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____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_12360____closed__6; +x_7 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__6; x_8 = l_Repr_addAppParen(x_7, x_2); return x_8; } @@ -18795,14 +18941,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_12360____closed__10; +x_11 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____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_12360____closed__12; +x_13 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__12; x_14 = l_Repr_addAppParen(x_13, x_2); return x_14; } @@ -18815,14 +18961,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_12360____closed__16; +x_17 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____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_12360____closed__18; +x_19 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__18; x_20 = l_Repr_addAppParen(x_19, x_2); return x_20; } @@ -18835,14 +18981,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_12360____closed__22; +x_23 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____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_12360____closed__24; +x_25 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__24; x_26 = l_Repr_addAppParen(x_25, x_2); return x_26; } @@ -18850,13 +18996,13 @@ return x_26; } } } -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____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_12534____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_12360_(x_3, x_2); +x_4 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534_(x_3, x_2); lean_dec(x_2); return x_4; } @@ -18865,7 +19011,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_12360____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____boxed), 2, 0); return x_1; } } @@ -18948,7 +19094,7 @@ x_1 = 0; return x_1; } } -LEAN_EXPORT uint8_t l___private_Init_Meta_0__Lean_Meta_beqEtaStructMode____x40_Init_Meta___hyg_12541_(uint8_t x_1, uint8_t x_2) { +LEAN_EXPORT uint8_t l___private_Init_Meta_0__Lean_Meta_beqEtaStructMode____x40_Init_Meta___hyg_12715_(uint8_t x_1, uint8_t x_2) { _start: { lean_object* x_3; lean_object* x_4; uint8_t x_5; @@ -18960,7 +19106,7 @@ lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_beqEtaStructMode____x40_Init_Meta___hyg_12541____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_12715____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; @@ -18968,7 +19114,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_12541_(x_3, x_4); +x_5 = l___private_Init_Meta_0__Lean_Meta_beqEtaStructMode____x40_Init_Meta___hyg_12715_(x_3, x_4); x_6 = lean_box(x_5); return x_6; } @@ -18977,7 +19123,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_12541____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Init_Meta_0__Lean_Meta_beqEtaStructMode____x40_Init_Meta___hyg_12715____boxed), 2, 0); return x_1; } } @@ -18989,7 +19135,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_12565____closed__1() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__1() { _start: { lean_object* x_1; @@ -18997,33 +19143,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_12565____closed__2() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____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_12565____closed__1; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__1; x_2 = lean_alloc_ctor(3, 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_12565____closed__3() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2236____closed__4; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__2; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____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_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__4() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____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_12565____closed__3; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__3; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -19031,23 +19177,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_12565____closed__5() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2236____closed__5; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__2; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____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_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__6() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____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_12565____closed__5; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__5; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -19055,7 +19201,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_12565____closed__7() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__7() { _start: { lean_object* x_1; @@ -19063,33 +19209,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_12565____closed__8() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____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_12565____closed__7; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__7; x_2 = lean_alloc_ctor(3, 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_12565____closed__9() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2236____closed__4; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__8; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__8; 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_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__10() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____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_12565____closed__9; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__9; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -19097,23 +19243,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_12565____closed__11() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2236____closed__5; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__8; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__8; 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_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__12() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____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_12565____closed__11; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__11; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -19121,7 +19267,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_12565____closed__13() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__13() { _start: { lean_object* x_1; @@ -19129,33 +19275,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_12565____closed__14() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____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_12565____closed__13; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__13; x_2 = lean_alloc_ctor(3, 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_12565____closed__15() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2236____closed__4; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__14; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__14; 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_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__16() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____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_12565____closed__15; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__15; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -19163,23 +19309,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_12565____closed__17() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_2236____closed__5; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__14; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__14; 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_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__18() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____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_12565____closed__17; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__17; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -19187,7 +19333,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_12565_(uint8_t x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739_(uint8_t x_1, lean_object* x_2) { _start: { switch (x_1) { @@ -19199,14 +19345,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_12565____closed__4; +x_5 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____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_12565____closed__6; +x_7 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__6; x_8 = l_Repr_addAppParen(x_7, x_2); return x_8; } @@ -19219,14 +19365,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_12565____closed__10; +x_11 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____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_12565____closed__12; +x_13 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__12; x_14 = l_Repr_addAppParen(x_13, x_2); return x_14; } @@ -19239,14 +19385,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_12565____closed__16; +x_17 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____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_12565____closed__18; +x_19 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__18; x_20 = l_Repr_addAppParen(x_19, x_2); return x_20; } @@ -19254,13 +19400,13 @@ return x_20; } } } -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____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_12739____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_12565_(x_3, x_2); +x_4 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739_(x_3, x_2); lean_dec(x_2); return x_4; } @@ -19269,7 +19415,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_12565____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____boxed), 2, 0); return x_1; } } @@ -19371,7 +19517,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_12794_(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_12968_(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; @@ -19524,7 +19670,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_12541_(x_6, x_14); +x_36 = l___private_Init_Meta_0__Lean_Meta_beqEtaStructMode____x40_Init_Meta___hyg_12715_(x_6, x_14); if (x_36 == 0) { uint8_t x_37; @@ -19649,11 +19795,11 @@ goto block_48; } } } -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_beqConfig____x40_Init_Meta___hyg_12794____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_12968____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_12794_(x_1, x_2); +x_3 = l___private_Init_Meta_0__Lean_Meta_DSimp_beqConfig____x40_Init_Meta___hyg_12968_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -19664,7 +19810,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_12794____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Init_Meta_0__Lean_Meta_DSimp_beqConfig____x40_Init_Meta___hyg_12968____boxed), 2, 0); return x_1; } } @@ -19676,7 +19822,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_13000____closed__1() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__1() { _start: { lean_object* x_1; @@ -19684,33 +19830,33 @@ 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_13000____closed__2() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____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_13000____closed__1; +x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__1; x_2 = lean_alloc_ctor(3, 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_13000____closed__3() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____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_13000____closed__2; +x_2 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__2; x_3 = lean_alloc_ctor(5, 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_13000____closed__4() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__4() { _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_13000____closed__3; +x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__3; x_2 = l___private_Init_Meta_0__Lean_Syntax_reprTSyntax____x40_Init_Meta___hyg_2642____rarg___closed__5; x_3 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -19718,7 +19864,7 @@ 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_13000____closed__5() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__5() { _start: { lean_object* x_1; lean_object* x_2; @@ -19727,7 +19873,7 @@ 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_13000____closed__6() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__6() { _start: { lean_object* x_1; @@ -19735,17 +19881,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_13000____closed__7() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__7() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__6; +x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__6; x_2 = lean_alloc_ctor(3, 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_13000____closed__8() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__8() { _start: { lean_object* x_1; @@ -19753,17 +19899,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_13000____closed__9() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__9() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__8; +x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__8; x_2 = lean_alloc_ctor(3, 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_13000____closed__10() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__10() { _start: { lean_object* x_1; @@ -19771,17 +19917,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_13000____closed__11() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__11() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__10; +x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__10; x_2 = lean_alloc_ctor(3, 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_13000____closed__12() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__12() { _start: { lean_object* x_1; lean_object* x_2; @@ -19790,7 +19936,7 @@ 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_13000____closed__13() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__13() { _start: { lean_object* x_1; @@ -19798,17 +19944,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_13000____closed__14() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____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_13000____closed__13; +x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__13; x_2 = lean_alloc_ctor(3, 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_13000____closed__15() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__15() { _start: { lean_object* x_1; @@ -19816,17 +19962,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_13000____closed__16() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____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_13000____closed__15; +x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__15; x_2 = lean_alloc_ctor(3, 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_13000____closed__17() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__17() { _start: { lean_object* x_1; @@ -19834,17 +19980,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_13000____closed__18() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____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_13000____closed__17; +x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__17; x_2 = lean_alloc_ctor(3, 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_13000____closed__19() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__19() { _start: { lean_object* x_1; lean_object* x_2; @@ -19853,7 +19999,7 @@ 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_13000____closed__20() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__20() { _start: { lean_object* x_1; @@ -19861,17 +20007,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_13000____closed__21() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__21() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__20; +x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__20; x_2 = lean_alloc_ctor(3, 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_13000____closed__22() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__22() { _start: { lean_object* x_1; lean_object* x_2; @@ -19880,7 +20026,7 @@ 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_13000____closed__23() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__23() { _start: { lean_object* x_1; lean_object* x_2; @@ -19890,23 +20036,23 @@ 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_13000____closed__24() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__24() { _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_13000____closed__22; -x_2 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__23; +x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__22; +x_2 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____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___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__25() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__25() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__24; +x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__24; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -19914,7 +20060,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_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__26() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__26() { _start: { lean_object* x_1; lean_object* x_2; @@ -19924,23 +20070,23 @@ 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_13000____closed__27() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__27() { _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_13000____closed__22; -x_2 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__26; +x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__22; +x_2 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__26; 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_13000____closed__28() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__28() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__27; +x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__27; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -19948,7 +20094,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_DSimp_reprConfig____x40_Init_Meta___hyg_13000_(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_13174_(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; lean_object* x_9; lean_object* x_10; uint8_t x_11; lean_object* x_12; uint8_t x_13; uint8_t x_14; uint8_t x_15; uint8_t x_16; lean_object* x_17; @@ -19957,8 +20103,8 @@ 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_12565_(x_6, x_7); -x_9 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__12; +x_8 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739_(x_6, x_7); +x_9 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__12; x_10 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_10, 0, x_9); lean_ctor_set(x_10, 1, x_8); @@ -19973,28 +20119,28 @@ x_16 = lean_ctor_get_uint8(x_1, 7); if (x_3 == 0) { lean_object* x_118; -x_118 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__23; +x_118 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__23; x_17 = x_118; goto block_117; } else { lean_object* x_119; -x_119 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__26; +x_119 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__26; x_17 = x_119; goto block_117; } block_117: { 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_18 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__5; +x_18 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__5; x_19 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_19, 0, x_18); lean_ctor_set(x_19, 1, x_17); x_20 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_20, 0, x_19); lean_ctor_set_uint8(x_20, sizeof(void*)*1, x_11); -x_21 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__4; +x_21 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__4; x_22 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_22, 0, x_21); lean_ctor_set(x_22, 1, x_20); @@ -20006,7 +20152,7 @@ x_25 = lean_box(1); x_26 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_26, 0, x_24); lean_ctor_set(x_26, 1, x_25); -x_27 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__7; +x_27 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__7; x_28 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_28, 0, x_26); lean_ctor_set(x_28, 1, x_27); @@ -20017,14 +20163,14 @@ lean_ctor_set(x_30, 1, x_29); if (x_4 == 0) { lean_object* x_115; -x_115 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__23; +x_115 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__23; x_31 = x_115; goto block_114; } else { lean_object* x_116; -x_116 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__26; +x_116 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__26; x_31 = x_116; goto block_114; } @@ -20046,7 +20192,7 @@ lean_ctor_set(x_35, 1, x_23); x_36 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_36, 0, x_35); lean_ctor_set(x_36, 1, x_25); -x_37 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__9; +x_37 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__9; x_38 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_38, 0, x_36); lean_ctor_set(x_38, 1, x_37); @@ -20056,14 +20202,14 @@ lean_ctor_set(x_39, 1, x_29); if (x_5 == 0) { lean_object* x_112; -x_112 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__23; +x_112 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__23; x_40 = x_112; goto block_111; } else { lean_object* x_113; -x_113 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__26; +x_113 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__26; x_40 = x_113; goto block_111; } @@ -20086,7 +20232,7 @@ lean_ctor_set(x_45, 1, x_23); x_46 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_46, 0, x_45); lean_ctor_set(x_46, 1, x_25); -x_47 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__11; +x_47 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__11; x_48 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_48, 0, x_46); lean_ctor_set(x_48, 1, x_47); @@ -20102,7 +20248,7 @@ lean_ctor_set(x_51, 1, x_23); x_52 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_52, 0, x_51); lean_ctor_set(x_52, 1, x_25); -x_53 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__14; +x_53 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__14; x_54 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_54, 0, x_52); lean_ctor_set(x_54, 1, x_53); @@ -20112,14 +20258,14 @@ lean_ctor_set(x_55, 1, x_29); if (x_13 == 0) { lean_object* x_109; -x_109 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__23; +x_109 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__23; x_56 = x_109; goto block_108; } else { lean_object* x_110; -x_110 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__26; +x_110 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__26; x_56 = x_110; goto block_108; } @@ -20141,7 +20287,7 @@ lean_ctor_set(x_60, 1, x_23); x_61 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_61, 0, x_60); lean_ctor_set(x_61, 1, x_25); -x_62 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__16; +x_62 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__16; x_63 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_63, 0, x_61); lean_ctor_set(x_63, 1, x_62); @@ -20151,14 +20297,14 @@ lean_ctor_set(x_64, 1, x_29); if (x_14 == 0) { lean_object* x_106; -x_106 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__23; +x_106 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__23; x_65 = x_106; goto block_105; } else { lean_object* x_107; -x_107 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__26; +x_107 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__26; x_65 = x_107; goto block_105; } @@ -20180,7 +20326,7 @@ lean_ctor_set(x_69, 1, x_23); x_70 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_70, 0, x_69); lean_ctor_set(x_70, 1, x_25); -x_71 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__18; +x_71 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__18; x_72 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_72, 0, x_70); lean_ctor_set(x_72, 1, x_71); @@ -20190,21 +20336,21 @@ lean_ctor_set(x_73, 1, x_29); if (x_15 == 0) { lean_object* x_103; -x_103 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__23; +x_103 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__23; x_74 = x_103; goto block_102; } else { lean_object* x_104; -x_104 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__26; +x_104 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__26; x_74 = x_104; goto block_102; } block_102: { 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_75 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__19; +x_75 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__19; x_76 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_76, 0, x_75); lean_ctor_set(x_76, 1, x_74); @@ -20220,7 +20366,7 @@ lean_ctor_set(x_79, 1, x_23); x_80 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_80, 0, x_79); lean_ctor_set(x_80, 1, x_25); -x_81 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__21; +x_81 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__21; x_82 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_82, 0, x_80); lean_ctor_set(x_82, 1, x_81); @@ -20230,7 +20376,7 @@ lean_ctor_set(x_83, 1, x_29); if (x_16 == 0) { 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_84 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__25; +x_84 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__25; x_85 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_85, 0, x_83); lean_ctor_set(x_85, 1, x_84); @@ -20254,7 +20400,7 @@ return x_92; 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; -x_93 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__28; +x_93 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__28; x_94 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_94, 0, x_83); lean_ctor_set(x_94, 1, x_93); @@ -20283,11 +20429,11 @@ return x_101; } } } -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____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_13174____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_13000_(x_1, x_2); +x_3 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); return x_3; @@ -20297,7 +20443,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_13000____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____boxed), 2, 0); return x_1; } } @@ -20471,7 +20617,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_13366_(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_13540_(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; uint8_t x_17; lean_object* x_18; lean_object* 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; uint8_t x_31; uint8_t x_32; lean_object* x_33; lean_object* x_37; lean_object* x_43; lean_object* x_49; lean_object* x_55; lean_object* x_61; lean_object* x_69; lean_object* x_75; lean_object* x_81; lean_object* x_87; lean_object* x_93; uint8_t x_99; @@ -20729,7 +20875,7 @@ goto block_54; { uint8_t x_62; lean_dec(x_61); -x_62 = l___private_Init_Meta_0__Lean_Meta_beqEtaStructMode____x40_Init_Meta___hyg_12541_(x_11, x_26); +x_62 = l___private_Init_Meta_0__Lean_Meta_beqEtaStructMode____x40_Init_Meta___hyg_12715_(x_11, x_26); if (x_62 == 0) { uint8_t x_63; @@ -20954,11 +21100,11 @@ goto block_92; } } } -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_beqConfig____x40_Init_Meta___hyg_13366____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_13540____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_13366_(x_1, x_2); +x_3 = l___private_Init_Meta_0__Lean_Meta_Simp_beqConfig____x40_Init_Meta___hyg_13540_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -20969,7 +21115,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_13366____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Init_Meta_0__Lean_Meta_Simp_beqConfig____x40_Init_Meta___hyg_13540____boxed), 2, 0); return x_1; } } @@ -20981,7 +21127,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_13703____closed__1() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__1() { _start: { lean_object* x_1; @@ -20989,33 +21135,33 @@ 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_13703____closed__2() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____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_13703____closed__1; +x_1 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__1; x_2 = lean_alloc_ctor(3, 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_13703____closed__3() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____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_13703____closed__2; +x_2 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__2; x_3 = lean_alloc_ctor(5, 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_13703____closed__4() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____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_13703____closed__3; +x_1 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__3; x_2 = l___private_Init_Meta_0__Lean_Syntax_reprTSyntax____x40_Init_Meta___hyg_2642____rarg___closed__5; x_3 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -21023,7 +21169,7 @@ 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_13703____closed__5() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__5() { _start: { lean_object* x_1; lean_object* x_2; @@ -21032,7 +21178,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__6() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__6() { _start: { lean_object* x_1; @@ -21040,17 +21186,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_13703____closed__7() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__7() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__6; +x_1 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__6; x_2 = lean_alloc_ctor(3, 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_13703____closed__8() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__8() { _start: { lean_object* x_1; lean_object* x_2; @@ -21059,7 +21205,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__9() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__9() { _start: { lean_object* x_1; @@ -21067,17 +21213,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_13703____closed__10() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____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_13703____closed__9; +x_1 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__9; x_2 = lean_alloc_ctor(3, 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_13703____closed__11() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__11() { _start: { lean_object* x_1; @@ -21085,17 +21231,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_13703____closed__12() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____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_13703____closed__11; +x_1 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__11; x_2 = lean_alloc_ctor(3, 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_13703____closed__13() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__13() { _start: { lean_object* x_1; lean_object* x_2; @@ -21104,7 +21250,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__14() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__14() { _start: { lean_object* x_1; @@ -21112,17 +21258,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_13703____closed__15() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__15() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__14; +x_1 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__14; x_2 = lean_alloc_ctor(3, 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_13703____closed__16() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__16() { _start: { lean_object* x_1; @@ -21130,17 +21276,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_13703____closed__17() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__17() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__16; +x_1 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__16; x_2 = lean_alloc_ctor(3, 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_13703____closed__18() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__18() { _start: { lean_object* x_1; lean_object* x_2; @@ -21149,7 +21295,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__19() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__19() { _start: { lean_object* x_1; @@ -21157,33 +21303,33 @@ x_1 = lean_mk_string_from_bytes("dsimp", 5); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__20() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__20() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__19; +x_1 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__19; x_2 = lean_alloc_ctor(3, 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_13703____closed__21() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__21() { _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_13703____closed__18; -x_2 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__23; +x_1 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__18; +x_2 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____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___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__22() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__22() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__21; +x_1 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__21; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -21191,23 +21337,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_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__23() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__23() { _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_13703____closed__18; -x_2 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__26; +x_1 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__18; +x_2 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__26; 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_13703____closed__24() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__24() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__23; +x_1 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__23; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -21215,7 +21361,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_Simp_reprConfig____x40_Init_Meta___hyg_13703_(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_13877_(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; uint8_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_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; uint8_t x_32; uint8_t 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; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; uint8_t x_45; uint8_t x_46; uint8_t x_47; uint8_t x_48; uint8_t x_49; lean_object* x_50; @@ -21224,7 +21370,7 @@ lean_inc(x_3); x_4 = l_Nat_repr(x_3); x_5 = lean_alloc_ctor(3, 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_13703____closed__5; +x_6 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__5; x_7 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_5); @@ -21232,7 +21378,7 @@ x_8 = 0; x_9 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_9, 0, x_7); lean_ctor_set_uint8(x_9, sizeof(void*)*1, x_8); -x_10 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__4; +x_10 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__4; x_11 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); @@ -21244,7 +21390,7 @@ x_14 = lean_box(1); x_15 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_15, 0, x_13); lean_ctor_set(x_15, 1, x_14); -x_16 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__7; +x_16 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__7; x_17 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_17, 0, x_15); lean_ctor_set(x_17, 1, x_16); @@ -21257,7 +21403,7 @@ lean_inc(x_20); x_21 = l_Nat_repr(x_20); x_22 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_22, 0, x_21); -x_23 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__8; +x_23 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__8; x_24 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_24, 0, x_23); lean_ctor_set(x_24, 1, x_22); @@ -21273,7 +21419,7 @@ lean_ctor_set(x_27, 1, x_12); x_28 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_28, 0, x_27); lean_ctor_set(x_28, 1, x_14); -x_29 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__10; +x_29 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__10; x_30 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_30, 0, x_28); lean_ctor_set(x_30, 1, x_29); @@ -21288,8 +21434,8 @@ x_36 = lean_ctor_get_uint8(x_1, sizeof(void*)*2 + 4); x_37 = lean_ctor_get_uint8(x_1, sizeof(void*)*2 + 5); x_38 = lean_ctor_get_uint8(x_1, sizeof(void*)*2 + 6); x_39 = lean_unsigned_to_nat(0u); -x_40 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565_(x_38, x_39); -x_41 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__12; +x_40 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739_(x_38, x_39); +x_41 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__12; x_42 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_42, 0, x_41); lean_ctor_set(x_42, 1, x_40); @@ -21306,21 +21452,21 @@ lean_dec(x_1); if (x_32 == 0) { lean_object* x_210; -x_210 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__23; +x_210 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__23; x_50 = x_210; goto block_209; } else { lean_object* x_211; -x_211 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__26; +x_211 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__26; x_50 = x_211; goto block_209; } block_209: { 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_51 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__22; +x_51 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__22; x_52 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_52, 0, x_51); lean_ctor_set(x_52, 1, x_50); @@ -21336,7 +21482,7 @@ lean_ctor_set(x_55, 1, x_12); x_56 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_56, 0, x_55); lean_ctor_set(x_56, 1, x_14); -x_57 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__12; +x_57 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__12; x_58 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_58, 0, x_56); lean_ctor_set(x_58, 1, x_57); @@ -21346,21 +21492,21 @@ lean_ctor_set(x_59, 1, x_18); if (x_33 == 0) { lean_object* x_207; -x_207 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__23; +x_207 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__23; x_60 = x_207; goto block_206; } else { lean_object* x_208; -x_208 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__26; +x_208 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__26; x_60 = x_208; goto block_206; } block_206: { 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 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__13; +x_61 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__13; x_62 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_62, 0, x_61); lean_ctor_set(x_62, 1, x_60); @@ -21376,7 +21522,7 @@ lean_ctor_set(x_65, 1, x_12); x_66 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_66, 0, x_65); lean_ctor_set(x_66, 1, x_14); -x_67 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__15; +x_67 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__15; x_68 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_68, 0, x_66); lean_ctor_set(x_68, 1, x_67); @@ -21386,14 +21532,14 @@ lean_ctor_set(x_69, 1, x_18); if (x_34 == 0) { lean_object* x_204; -x_204 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__23; +x_204 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__23; x_70 = x_204; goto block_203; } else { lean_object* x_205; -x_205 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__26; +x_205 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__26; x_70 = x_205; goto block_203; } @@ -21415,7 +21561,7 @@ lean_ctor_set(x_74, 1, x_12); x_75 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_75, 0, x_74); lean_ctor_set(x_75, 1, x_14); -x_76 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__2; +x_76 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__2; x_77 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_77, 0, x_75); lean_ctor_set(x_77, 1, x_76); @@ -21425,21 +21571,21 @@ lean_ctor_set(x_78, 1, x_18); if (x_35 == 0) { lean_object* x_201; -x_201 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__23; +x_201 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__23; x_79 = x_201; goto block_200; } else { lean_object* x_202; -x_202 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__26; +x_202 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__26; x_79 = x_202; goto block_200; } block_200: { 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___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__5; +x_80 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__5; x_81 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_81, 0, x_80); lean_ctor_set(x_81, 1, x_79); @@ -21455,7 +21601,7 @@ lean_ctor_set(x_84, 1, x_12); x_85 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_85, 0, x_84); lean_ctor_set(x_85, 1, x_14); -x_86 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__7; +x_86 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__7; x_87 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_87, 0, x_85); lean_ctor_set(x_87, 1, x_86); @@ -21465,14 +21611,14 @@ lean_ctor_set(x_88, 1, x_18); if (x_36 == 0) { lean_object* x_198; -x_198 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__23; +x_198 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__23; x_89 = x_198; goto block_197; } else { lean_object* x_199; -x_199 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__26; +x_199 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__26; x_89 = x_199; goto block_197; } @@ -21494,7 +21640,7 @@ lean_ctor_set(x_93, 1, x_12); x_94 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_94, 0, x_93); lean_ctor_set(x_94, 1, x_14); -x_95 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__9; +x_95 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__9; x_96 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_96, 0, x_94); lean_ctor_set(x_96, 1, x_95); @@ -21504,14 +21650,14 @@ lean_ctor_set(x_97, 1, x_18); if (x_37 == 0) { lean_object* x_195; -x_195 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__23; +x_195 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__23; x_98 = x_195; goto block_194; } else { lean_object* x_196; -x_196 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__26; +x_196 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__26; x_98 = x_196; goto block_194; } @@ -21534,7 +21680,7 @@ lean_ctor_set(x_103, 1, x_12); x_104 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_104, 0, x_103); lean_ctor_set(x_104, 1, x_14); -x_105 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__11; +x_105 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__11; x_106 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_106, 0, x_104); lean_ctor_set(x_106, 1, x_105); @@ -21550,7 +21696,7 @@ lean_ctor_set(x_109, 1, x_12); x_110 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_110, 0, x_109); lean_ctor_set(x_110, 1, x_14); -x_111 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__14; +x_111 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__14; x_112 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_112, 0, x_110); lean_ctor_set(x_112, 1, x_111); @@ -21560,14 +21706,14 @@ lean_ctor_set(x_113, 1, x_18); if (x_44 == 0) { lean_object* x_192; -x_192 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__23; +x_192 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__23; x_114 = x_192; goto block_191; } else { lean_object* x_193; -x_193 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__26; +x_193 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__26; x_114 = x_193; goto block_191; } @@ -21589,7 +21735,7 @@ lean_ctor_set(x_118, 1, x_12); x_119 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_119, 0, x_118); lean_ctor_set(x_119, 1, x_14); -x_120 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__16; +x_120 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__16; x_121 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_121, 0, x_119); lean_ctor_set(x_121, 1, x_120); @@ -21599,14 +21745,14 @@ lean_ctor_set(x_122, 1, x_18); if (x_45 == 0) { lean_object* x_189; -x_189 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__23; +x_189 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__23; x_123 = x_189; goto block_188; } else { lean_object* x_190; -x_190 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__26; +x_190 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__26; x_123 = x_190; goto block_188; } @@ -21628,7 +21774,7 @@ lean_ctor_set(x_127, 1, x_12); x_128 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_128, 0, x_127); lean_ctor_set(x_128, 1, x_14); -x_129 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__18; +x_129 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__18; x_130 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_130, 0, x_128); lean_ctor_set(x_130, 1, x_129); @@ -21638,21 +21784,21 @@ lean_ctor_set(x_131, 1, x_18); if (x_46 == 0) { lean_object* x_186; -x_186 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__23; +x_186 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__23; x_132 = x_186; goto block_185; } else { lean_object* x_187; -x_187 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__26; +x_187 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__26; x_132 = x_187; goto block_185; } block_185: { 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_133 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__19; +x_133 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__19; x_134 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_134, 0, x_133); lean_ctor_set(x_134, 1, x_132); @@ -21668,7 +21814,7 @@ lean_ctor_set(x_137, 1, x_12); x_138 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_138, 0, x_137); lean_ctor_set(x_138, 1, x_14); -x_139 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__17; +x_139 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__17; x_140 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_140, 0, x_138); lean_ctor_set(x_140, 1, x_139); @@ -21678,21 +21824,21 @@ lean_ctor_set(x_141, 1, x_18); if (x_47 == 0) { lean_object* x_183; -x_183 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__23; +x_183 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__23; x_142 = x_183; goto block_182; } else { lean_object* x_184; -x_184 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__26; +x_184 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__26; x_142 = x_184; goto block_182; } block_182: { 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; -x_143 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__18; +x_143 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__18; x_144 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_144, 0, x_143); lean_ctor_set(x_144, 1, x_142); @@ -21708,7 +21854,7 @@ lean_ctor_set(x_147, 1, x_12); x_148 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_148, 0, x_147); lean_ctor_set(x_148, 1, x_14); -x_149 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__21; +x_149 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__21; x_150 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_150, 0, x_148); lean_ctor_set(x_150, 1, x_149); @@ -21718,14 +21864,14 @@ lean_ctor_set(x_151, 1, x_18); if (x_48 == 0) { lean_object* x_180; -x_180 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__23; +x_180 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__23; x_152 = x_180; goto block_179; } else { lean_object* x_181; -x_181 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__26; +x_181 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__26; x_152 = x_181; goto block_179; } @@ -21747,7 +21893,7 @@ lean_ctor_set(x_156, 1, x_12); x_157 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_157, 0, x_156); lean_ctor_set(x_157, 1, x_14); -x_158 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__20; +x_158 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__20; x_159 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_159, 0, x_157); lean_ctor_set(x_159, 1, x_158); @@ -21757,7 +21903,7 @@ lean_ctor_set(x_160, 1, x_18); if (x_49 == 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; -x_161 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__22; +x_161 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__22; x_162 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_162, 0, x_160); lean_ctor_set(x_162, 1, x_161); @@ -21781,7 +21927,7 @@ return x_169; 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; -x_170 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__24; +x_170 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__24; x_171 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_171, 0, x_160); lean_ctor_set(x_171, 1, x_170); @@ -21815,11 +21961,11 @@ return x_178; } } } -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____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_13877____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_13703_(x_1, x_2); +x_3 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877_(x_1, x_2); lean_dec(x_2); return x_3; } @@ -21828,7 +21974,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_13703____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____boxed), 2, 0); return x_1; } } @@ -21932,7 +22078,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_tacticErw_______closed__4() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("erw ", 4); +x_1 = lean_mk_string_from_bytes("erw", 3); return x_1; } } @@ -22571,14 +22717,22 @@ return x_3; static lean_object* _init_l_Lean_Parser_Tactic_simpAllKind___closed__5() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(" (", 2); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_simpAllKind___closed__6() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_TSyntax_expandInterpolatedStr___closed__4; +x_1 = l_Lean_Parser_Tactic_simpAllKind___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_Tactic_simpAllKind___closed__6() { +static lean_object* _init_l_Lean_Parser_Tactic_simpAllKind___closed__7() { _start: { lean_object* x_1; @@ -22586,11 +22740,11 @@ x_1 = lean_mk_string_from_bytes("all", 3); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_simpAllKind___closed__7() { +static lean_object* _init_l_Lean_Parser_Tactic_simpAllKind___closed__8() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_simpAllKind___closed__6; +x_1 = l_Lean_Parser_Tactic_simpAllKind___closed__7; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -22598,13 +22752,13 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_simpAllKind___closed__8() { +static lean_object* _init_l_Lean_Parser_Tactic_simpAllKind___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_simpAllKind___closed__5; -x_3 = l_Lean_Parser_Tactic_simpAllKind___closed__7; +x_2 = l_Lean_Parser_Tactic_simpAllKind___closed__6; +x_3 = l_Lean_Parser_Tactic_simpAllKind___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); @@ -22612,19 +22766,19 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_simpAllKind___closed__9() { +static lean_object* _init_l_Lean_Parser_Tactic_simpAllKind___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Tactic_simpAllKind___closed__4; -x_2 = l_Lean_Parser_Tactic_simpAllKind___closed__8; +x_2 = l_Lean_Parser_Tactic_simpAllKind___closed__9; 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_simpAllKind___closed__10() { +static lean_object* _init_l_Lean_Parser_Tactic_simpAllKind___closed__11() { _start: { lean_object* x_1; lean_object* x_2; @@ -22634,13 +22788,13 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic_simpAllKind___closed__11() { +static lean_object* _init_l_Lean_Parser_Tactic_simpAllKind___closed__12() { _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_simpAllKind___closed__9; -x_3 = l_Lean_Parser_Tactic_simpAllKind___closed__10; +x_2 = l_Lean_Parser_Tactic_simpAllKind___closed__10; +x_3 = l_Lean_Parser_Tactic_simpAllKind___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); @@ -22648,7 +22802,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_simpAllKind___closed__12() { +static lean_object* _init_l_Lean_Parser_Tactic_simpAllKind___closed__13() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; @@ -22660,13 +22814,13 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_simpAllKind___closed__13() { +static lean_object* _init_l_Lean_Parser_Tactic_simpAllKind___closed__14() { _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_simpAllKind___closed__11; -x_3 = l_Lean_Parser_Tactic_simpAllKind___closed__12; +x_2 = l_Lean_Parser_Tactic_simpAllKind___closed__12; +x_3 = l_Lean_Parser_Tactic_simpAllKind___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); @@ -22674,7 +22828,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_simpAllKind___closed__14() { +static lean_object* _init_l_Lean_Parser_Tactic_simpAllKind___closed__15() { _start: { lean_object* x_1; lean_object* x_2; @@ -22684,13 +22838,13 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic_simpAllKind___closed__15() { +static lean_object* _init_l_Lean_Parser_Tactic_simpAllKind___closed__16() { _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_simpAllKind___closed__13; -x_3 = l_Lean_Parser_Tactic_simpAllKind___closed__14; +x_2 = l_Lean_Parser_Tactic_simpAllKind___closed__14; +x_3 = l_Lean_Parser_Tactic_simpAllKind___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); @@ -22698,13 +22852,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_simpAllKind___closed__16() { +static lean_object* _init_l_Lean_Parser_Tactic_simpAllKind___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_simpAllKind___closed__1; x_2 = l_Lean_Parser_Tactic_simpAllKind___closed__2; -x_3 = l_Lean_Parser_Tactic_simpAllKind___closed__15; +x_3 = l_Lean_Parser_Tactic_simpAllKind___closed__16; x_4 = lean_alloc_ctor(9, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -22716,7 +22870,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_simpAllKind() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Tactic_simpAllKind___closed__16; +x_1 = l_Lean_Parser_Tactic_simpAllKind___closed__17; return x_1; } } @@ -22744,7 +22898,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_dsimpKind___closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__19; +x_1 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__19; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -22757,7 +22911,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_dsimpKind___closed__4() { { 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_simpAllKind___closed__5; +x_2 = l_Lean_Parser_Tactic_simpAllKind___closed__6; x_3 = l_Lean_Parser_Tactic_dsimpKind___closed__3; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -22784,7 +22938,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_dsimpKind___closed__6() { 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_dsimpKind___closed__5; -x_3 = l_Lean_Parser_Tactic_simpAllKind___closed__10; +x_3 = l_Lean_Parser_Tactic_simpAllKind___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); @@ -22798,7 +22952,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_dsimpKind___closed__7() { 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_dsimpKind___closed__6; -x_3 = l_Lean_Parser_Tactic_simpAllKind___closed__12; +x_3 = l_Lean_Parser_Tactic_simpAllKind___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); @@ -22812,7 +22966,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_dsimpKind___closed__8() { 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_dsimpKind___closed__7; -x_3 = l_Lean_Parser_Tactic_simpAllKind___closed__14; +x_3 = l_Lean_Parser_Tactic_simpAllKind___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); @@ -22996,7 +23150,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__1 _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("ident", 5); +x_1 = lean_mk_string_from_bytes("ppSpace", 7); return x_1; } } @@ -23023,10 +23177,22 @@ return x_2; static lean_object* _init_l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__18() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_mkGroupNode___closed__2; +x_2 = l_Lean_Parser_Tactic_declareSimpLikeTactic___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_Tactic_declareSimpLikeTactic___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_declareSimpLikeTactic___closed__14; -x_3 = l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__17; +x_3 = l_Lean_Parser_Tactic_declareSimpLikeTactic___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); @@ -23034,7 +23200,63 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__19() { +static lean_object* _init_l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__20() { +_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_declareSimpLikeTactic___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_Tactic_declareSimpLikeTactic___closed__20; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__22() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Tactic_declareSimpLikeTactic___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_Tactic_declareSimpLikeTactic___closed__23() { +_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_declareSimpLikeTactic___closed__19; +x_3 = l_Lean_Parser_Tactic_declareSimpLikeTactic___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); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__24() { +_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_declareSimpLikeTactic___closed__23; +x_3 = l_Lean_Parser_Tactic_declareSimpLikeTactic___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); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__25() { _start: { lean_object* x_1; lean_object* x_2; @@ -23044,13 +23266,13 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__20() { +static lean_object* _init_l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__26() { _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_declareSimpLikeTactic___closed__18; -x_3 = l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__19; +x_2 = l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__24; +x_3 = l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__25; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -23058,7 +23280,21 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__21() { +static lean_object* _init_l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__27() { +_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_declareSimpLikeTactic___closed__26; +x_3 = l_Lean_Parser_Tactic_declareSimpLikeTactic___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); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__28() { _start: { lean_object* x_1; @@ -23066,21 +23302,21 @@ x_1 = lean_mk_string_from_bytes("term", 4); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__22() { +static lean_object* _init_l_Lean_Parser_Tactic_declareSimpLikeTactic___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_declareSimpLikeTactic___closed__21; +x_2 = l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__28; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__23() { +static lean_object* _init_l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__30() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__22; +x_1 = l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__29; x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23088,13 +23324,13 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__24() { +static lean_object* _init_l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__31() { _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_declareSimpLikeTactic___closed__20; -x_3 = l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__23; +x_2 = l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__27; +x_3 = l_Lean_Parser_Tactic_declareSimpLikeTactic___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); @@ -23102,13 +23338,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__25() { +static lean_object* _init_l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__32() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__2; x_2 = lean_unsigned_to_nat(1022u); -x_3 = l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__24; +x_3 = l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__31; x_4 = lean_alloc_ctor(3, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -23120,7 +23356,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_declareSimpLikeTactic() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__25; +x_1 = l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__32; return x_1; } } @@ -23919,7 +24155,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_Lean_Parser_Tactic_declareSimpLikeTactic___closed__21; +x_1 = l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__28; x_2 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__81; x_3 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__82; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); @@ -25111,7 +25347,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_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__19; +x_1 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__19; x_2 = l_String_toSubstring_x27(x_1); return x_2; } @@ -25121,7 +25357,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; x_1 = lean_box(0); -x_2 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__19; +x_2 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__19; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -25133,7 +25369,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__1; x_2 = l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__2; x_3 = l_Lean_Parser_Tactic_tacticErw_______closed__1; -x_4 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__19; +x_4 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__19; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } @@ -26078,11 +26314,11 @@ 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); -x_12 = lean_unsigned_to_nat(3u); +x_12 = lean_unsigned_to_nat(4u); x_13 = l_Lean_Syntax_getArg(x_1, x_12); -x_14 = lean_unsigned_to_nat(4u); +x_14 = lean_unsigned_to_nat(6u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); -x_16 = lean_unsigned_to_nat(5u); +x_16 = lean_unsigned_to_nat(8u); x_17 = l_Lean_Syntax_getArg(x_1, x_16); lean_dec(x_1); x_18 = l_Lean_Syntax_getOptional_x3f(x_9); @@ -27648,7 +27884,7 @@ x_1 = l_Lean_Parser_Tactic_simpAutoUnfold___closed__26; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____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_17716____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; 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; lean_object* x_19; lean_object* x_20; @@ -27676,7 +27912,7 @@ lean_ctor_set(x_20, 1, x_4); return x_20; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__1() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__1() { _start: { lean_object* x_1; @@ -27684,16 +27920,16 @@ 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_17473____closed__2() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__1; +x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__1; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__3() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__3() { _start: { lean_object* x_1; @@ -27701,7 +27937,7 @@ x_1 = lean_mk_string_from_bytes("Meta", 4); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__4() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__4() { _start: { lean_object* x_1; @@ -27709,7 +27945,7 @@ 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_17473____closed__5() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__5() { _start: { lean_object* x_1; @@ -27717,84 +27953,84 @@ 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_17473____closed__6() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__1; -x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__3; -x_3 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__4; -x_4 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__5; +x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__3; +x_3 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__4; +x_4 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__5; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__7() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____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_17473____closed__6; +x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____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_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__8() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__6; +x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____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_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__9() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__9() { _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_17473____closed__8; +x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____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_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__10() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__7; -x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__9; +x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__7; +x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__9; 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_expandSimp____x40_Init_Meta___hyg_17473____closed__11() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__11() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__20; +x_1 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__20; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__12() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__12() { _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_13000____closed__20; +x_2 = l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__20; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473_(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_17716_(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; @@ -27855,12 +28091,12 @@ 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_expandSimp____x40_Init_Meta___hyg_17473____closed__6; +x_30 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____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_17473____closed__2; -x_33 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__10; +x_32 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__2; +x_33 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__10; lean_inc(x_12); x_34 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_34, 0, x_12); @@ -27908,11 +28144,11 @@ lean_ctor_set(x_50, 0, x_12); lean_ctor_set(x_50, 1, x_49); lean_inc(x_12); x_51 = l_Lean_Syntax_node2(x_12, x_35, x_48, x_50); -x_52 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__12; +x_52 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__12; lean_inc(x_13); lean_inc(x_14); x_53 = l_Lean_addMacroScope(x_14, x_52, x_13); -x_54 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__11; +x_54 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__11; lean_inc(x_12); x_55 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_55, 0, x_12); @@ -27979,7 +28215,7 @@ x_82 = l_Lean_Syntax_mkApp___closed__2; lean_inc(x_12); x_83 = l_Lean_Syntax_node2(x_12, x_82, x_75, x_81); x_84 = l_Lean_Syntax_node5(x_12, x_8, x_16, x_18, x_20, x_83, x_38); -x_85 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____lambda__1(x_1, x_84, x_2, x_3); +x_85 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____lambda__1(x_1, x_84, x_2, x_3); lean_dec(x_2); return x_85; } @@ -28034,12 +28270,12 @@ lean_inc(x_90); x_107 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_107, 0, x_90); lean_ctor_set(x_107, 1, x_106); -x_108 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__6; +x_108 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__6; lean_inc(x_91); lean_inc(x_92); x_109 = l_Lean_addMacroScope(x_92, x_108, x_91); -x_110 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__2; -x_111 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__10; +x_110 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__2; +x_111 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__10; lean_inc(x_90); x_112 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_112, 0, x_90); @@ -28087,11 +28323,11 @@ lean_ctor_set(x_128, 0, x_90); lean_ctor_set(x_128, 1, x_127); lean_inc(x_90); x_129 = l_Lean_Syntax_node2(x_90, x_113, x_126, x_128); -x_130 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__12; +x_130 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__12; lean_inc(x_91); lean_inc(x_92); x_131 = l_Lean_addMacroScope(x_92, x_130, x_91); -x_132 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__11; +x_132 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__11; lean_inc(x_90); x_133 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_133, 0, x_90); @@ -28143,17 +28379,17 @@ x_155 = l_Lean_Syntax_mkApp___closed__2; lean_inc(x_90); x_156 = l_Lean_Syntax_node2(x_90, x_155, x_153, x_154); x_157 = l_Lean_Syntax_node5(x_90, x_8, x_94, x_96, x_98, x_156, x_116); -x_158 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____lambda__1(x_1, x_157, x_2, x_3); +x_158 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____lambda__1(x_1, x_157, x_2, x_3); lean_dec(x_2); return x_158; } } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____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_17716____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_17473____lambda__1(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____lambda__1(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } @@ -28290,26 +28526,26 @@ 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_18276____closed__1() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18519____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_13703____closed__16; +x_1 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__16; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18276____closed__2() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18519____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_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__16; +x_2 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__16; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18276_(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_18519_(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; @@ -28370,12 +28606,12 @@ 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_expandSimp____x40_Init_Meta___hyg_17473____closed__6; +x_30 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____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_17473____closed__2; -x_33 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__10; +x_32 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__2; +x_33 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__10; lean_inc(x_12); x_34 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_34, 0, x_12); @@ -28423,11 +28659,11 @@ lean_ctor_set(x_50, 0, x_12); lean_ctor_set(x_50, 1, x_49); lean_inc(x_12); x_51 = l_Lean_Syntax_node2(x_12, x_35, x_48, x_50); -x_52 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18276____closed__2; +x_52 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18519____closed__2; lean_inc(x_13); lean_inc(x_14); x_53 = l_Lean_addMacroScope(x_14, x_52, x_13); -x_54 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18276____closed__1; +x_54 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18519____closed__1; lean_inc(x_12); x_55 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_55, 0, x_12); @@ -28494,7 +28730,7 @@ x_82 = l_Lean_Syntax_mkApp___closed__2; lean_inc(x_12); x_83 = l_Lean_Syntax_node2(x_12, x_82, x_75, x_81); x_84 = l_Lean_Syntax_node5(x_12, x_8, x_16, x_18, x_20, x_83, x_38); -x_85 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____lambda__1(x_1, x_84, x_2, x_3); +x_85 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____lambda__1(x_1, x_84, x_2, x_3); lean_dec(x_2); return x_85; } @@ -28549,12 +28785,12 @@ lean_inc(x_90); x_107 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_107, 0, x_90); lean_ctor_set(x_107, 1, x_106); -x_108 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__6; +x_108 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__6; lean_inc(x_91); lean_inc(x_92); x_109 = l_Lean_addMacroScope(x_92, x_108, x_91); -x_110 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__2; -x_111 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__10; +x_110 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__2; +x_111 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__10; lean_inc(x_90); x_112 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_112, 0, x_90); @@ -28602,11 +28838,11 @@ lean_ctor_set(x_128, 0, x_90); lean_ctor_set(x_128, 1, x_127); lean_inc(x_90); x_129 = l_Lean_Syntax_node2(x_90, x_113, x_126, x_128); -x_130 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18276____closed__2; +x_130 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18519____closed__2; lean_inc(x_91); lean_inc(x_92); x_131 = l_Lean_addMacroScope(x_92, x_130, x_91); -x_132 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18276____closed__1; +x_132 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18519____closed__1; lean_inc(x_90); x_133 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_133, 0, x_90); @@ -28658,7 +28894,7 @@ x_155 = l_Lean_Syntax_mkApp___closed__2; lean_inc(x_90); x_156 = l_Lean_Syntax_node2(x_90, x_155, x_153, x_154); x_157 = l_Lean_Syntax_node5(x_90, x_8, x_94, x_96, x_98, x_156, x_116); -x_158 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____lambda__1(x_1, x_157, x_2, x_3); +x_158 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____lambda__1(x_1, x_157, x_2, x_3); lean_dec(x_2); return x_158; } @@ -28796,7 +29032,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_19075_(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_19318_(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; @@ -28857,12 +29093,12 @@ 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_expandSimp____x40_Init_Meta___hyg_17473____closed__6; +x_30 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____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_17473____closed__2; -x_33 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__10; +x_32 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__2; +x_33 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__10; lean_inc(x_12); x_34 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_34, 0, x_12); @@ -28910,11 +29146,11 @@ lean_ctor_set(x_50, 0, x_12); lean_ctor_set(x_50, 1, x_49); lean_inc(x_12); x_51 = l_Lean_Syntax_node2(x_12, x_35, x_48, x_50); -x_52 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18276____closed__2; +x_52 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18519____closed__2; lean_inc(x_13); lean_inc(x_14); x_53 = l_Lean_addMacroScope(x_14, x_52, x_13); -x_54 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18276____closed__1; +x_54 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18519____closed__1; lean_inc(x_12); x_55 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_55, 0, x_12); @@ -28947,9 +29183,9 @@ lean_inc(x_12); x_66 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_66, 0, x_12); lean_ctor_set(x_66, 1, x_65); -x_67 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__12; +x_67 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__12; x_68 = l_Lean_addMacroScope(x_14, x_67, x_13); -x_69 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__11; +x_69 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__11; lean_inc(x_12); x_70 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_70, 0, x_12); @@ -29004,7 +29240,7 @@ x_90 = l_Lean_Syntax_mkApp___closed__2; lean_inc(x_12); x_91 = l_Lean_Syntax_node2(x_12, x_90, x_83, x_89); x_92 = l_Lean_Syntax_node5(x_12, x_8, x_16, x_18, x_20, x_91, x_38); -x_93 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____lambda__1(x_1, x_92, x_2, x_3); +x_93 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____lambda__1(x_1, x_92, x_2, x_3); lean_dec(x_2); return x_93; } @@ -29059,12 +29295,12 @@ lean_inc(x_98); x_115 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_115, 0, x_98); lean_ctor_set(x_115, 1, x_114); -x_116 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__6; +x_116 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__6; lean_inc(x_99); lean_inc(x_100); x_117 = l_Lean_addMacroScope(x_100, x_116, x_99); -x_118 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__2; -x_119 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__10; +x_118 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__2; +x_119 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__10; lean_inc(x_98); x_120 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_120, 0, x_98); @@ -29112,11 +29348,11 @@ lean_ctor_set(x_136, 0, x_98); lean_ctor_set(x_136, 1, x_135); lean_inc(x_98); x_137 = l_Lean_Syntax_node2(x_98, x_121, x_134, x_136); -x_138 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18276____closed__2; +x_138 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18519____closed__2; lean_inc(x_99); lean_inc(x_100); x_139 = l_Lean_addMacroScope(x_100, x_138, x_99); -x_140 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18276____closed__1; +x_140 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18519____closed__1; lean_inc(x_98); x_141 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_141, 0, x_98); @@ -29149,9 +29385,9 @@ lean_inc(x_98); x_152 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_152, 0, x_98); lean_ctor_set(x_152, 1, x_151); -x_153 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__12; +x_153 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__12; x_154 = l_Lean_addMacroScope(x_100, x_153, x_99); -x_155 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__11; +x_155 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__11; lean_inc(x_98); x_156 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_156, 0, x_98); @@ -29191,7 +29427,7 @@ x_171 = l_Lean_Syntax_mkApp___closed__2; lean_inc(x_98); x_172 = l_Lean_Syntax_node2(x_98, x_171, x_169, x_170); x_173 = l_Lean_Syntax_node5(x_98, x_8, x_102, x_104, x_106, x_172, x_124); -x_174 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____lambda__1(x_1, x_173, x_2, x_3); +x_174 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____lambda__1(x_1, x_173, x_2, x_3); lean_dec(x_2); return x_174; } @@ -29371,7 +29607,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_19958____lambda__1___closed__1() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____lambda__1___closed__1() { _start: { lean_object* x_1; @@ -29379,7 +29615,7 @@ x_1 = lean_mk_string_from_bytes("simp_all", 8); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____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_20201____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; 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; lean_object* x_19; lean_object* x_20; @@ -29387,7 +29623,7 @@ x_5 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tact 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_19958____lambda__1___closed__1; +x_9 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____lambda__1___closed__1; x_10 = 1; x_11 = l_Lean_mkAtomFrom(x_8, x_9, x_10); x_12 = l_Lean_Syntax_setArg(x_6, x_7, x_11); @@ -29407,7 +29643,7 @@ lean_ctor_set(x_20, 1, x_4); return x_20; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__1() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__1() { _start: { lean_object* x_1; @@ -29415,16 +29651,16 @@ 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_19958____closed__2() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__1; +x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__1; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__3() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__3() { _start: { lean_object* x_1; @@ -29432,65 +29668,65 @@ 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_19958____closed__4() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____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_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__1; -x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__3; -x_3 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__4; -x_4 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__3; +x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__3; +x_3 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__4; +x_4 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__3; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__5() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__5() { _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_19958____closed__4; +x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____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_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__6() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__4; +x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____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_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__7() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____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_19958____closed__6; +x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____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_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__8() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__5; -x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__7; +x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__5; +x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____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; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958_(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_20201_(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; @@ -29551,12 +29787,12 @@ 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_expandSimp____x40_Init_Meta___hyg_19958____closed__4; +x_30 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__4; 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_19958____closed__2; -x_33 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__8; +x_32 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__2; +x_33 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__8; lean_inc(x_12); x_34 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_34, 0, x_12); @@ -29604,11 +29840,11 @@ lean_ctor_set(x_50, 0, x_12); lean_ctor_set(x_50, 1, x_49); lean_inc(x_12); x_51 = l_Lean_Syntax_node2(x_12, x_35, x_48, x_50); -x_52 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__12; +x_52 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__12; lean_inc(x_13); lean_inc(x_14); x_53 = l_Lean_addMacroScope(x_14, x_52, x_13); -x_54 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__11; +x_54 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__11; lean_inc(x_12); x_55 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_55, 0, x_12); @@ -29675,7 +29911,7 @@ x_82 = l_Lean_Syntax_mkApp___closed__2; lean_inc(x_12); x_83 = l_Lean_Syntax_node2(x_12, x_82, x_75, x_81); x_84 = l_Lean_Syntax_node5(x_12, x_8, x_16, x_18, x_20, x_83, x_38); -x_85 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____lambda__1(x_1, x_84, x_2, x_3); +x_85 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____lambda__1(x_1, x_84, x_2, x_3); lean_dec(x_2); return x_85; } @@ -29730,12 +29966,12 @@ lean_inc(x_90); x_107 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_107, 0, x_90); lean_ctor_set(x_107, 1, x_106); -x_108 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__4; +x_108 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__4; lean_inc(x_91); lean_inc(x_92); x_109 = l_Lean_addMacroScope(x_92, x_108, x_91); -x_110 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__2; -x_111 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__8; +x_110 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__2; +x_111 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__8; lean_inc(x_90); x_112 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_112, 0, x_90); @@ -29783,11 +30019,11 @@ lean_ctor_set(x_128, 0, x_90); lean_ctor_set(x_128, 1, x_127); lean_inc(x_90); x_129 = l_Lean_Syntax_node2(x_90, x_113, x_126, x_128); -x_130 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__12; +x_130 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__12; lean_inc(x_91); lean_inc(x_92); x_131 = l_Lean_addMacroScope(x_92, x_130, x_91); -x_132 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__11; +x_132 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__11; lean_inc(x_90); x_133 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_133, 0, x_90); @@ -29839,17 +30075,17 @@ x_155 = l_Lean_Syntax_mkApp___closed__2; lean_inc(x_90); x_156 = l_Lean_Syntax_node2(x_90, x_155, x_153, x_154); x_157 = l_Lean_Syntax_node5(x_90, x_8, x_94, x_96, x_98, x_156, x_116); -x_158 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____lambda__1(x_1, x_157, x_2, x_3); +x_158 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____lambda__1(x_1, x_157, x_2, x_3); lean_dec(x_2); return x_158; } } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____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_20201____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_19958____lambda__1(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____lambda__1(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } @@ -29972,7 +30208,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_20747_(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_20990_(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; @@ -30033,12 +30269,12 @@ 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_expandSimp____x40_Init_Meta___hyg_19958____closed__4; +x_30 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__4; 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_19958____closed__2; -x_33 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__8; +x_32 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__2; +x_33 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__8; lean_inc(x_12); x_34 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_34, 0, x_12); @@ -30086,11 +30322,11 @@ lean_ctor_set(x_50, 0, x_12); lean_ctor_set(x_50, 1, x_49); lean_inc(x_12); x_51 = l_Lean_Syntax_node2(x_12, x_35, x_48, x_50); -x_52 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18276____closed__2; +x_52 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18519____closed__2; lean_inc(x_13); lean_inc(x_14); x_53 = l_Lean_addMacroScope(x_14, x_52, x_13); -x_54 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18276____closed__1; +x_54 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18519____closed__1; lean_inc(x_12); x_55 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_55, 0, x_12); @@ -30157,7 +30393,7 @@ x_82 = l_Lean_Syntax_mkApp___closed__2; lean_inc(x_12); x_83 = l_Lean_Syntax_node2(x_12, x_82, x_75, x_81); x_84 = l_Lean_Syntax_node5(x_12, x_8, x_16, x_18, x_20, x_83, x_38); -x_85 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____lambda__1(x_1, x_84, x_2, x_3); +x_85 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____lambda__1(x_1, x_84, x_2, x_3); lean_dec(x_2); return x_85; } @@ -30212,12 +30448,12 @@ lean_inc(x_90); x_107 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_107, 0, x_90); lean_ctor_set(x_107, 1, x_106); -x_108 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__4; +x_108 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__4; lean_inc(x_91); lean_inc(x_92); x_109 = l_Lean_addMacroScope(x_92, x_108, x_91); -x_110 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__2; -x_111 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__8; +x_110 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__2; +x_111 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__8; lean_inc(x_90); x_112 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_112, 0, x_90); @@ -30265,11 +30501,11 @@ lean_ctor_set(x_128, 0, x_90); lean_ctor_set(x_128, 1, x_127); lean_inc(x_90); x_129 = l_Lean_Syntax_node2(x_90, x_113, x_126, x_128); -x_130 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18276____closed__2; +x_130 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18519____closed__2; lean_inc(x_91); lean_inc(x_92); x_131 = l_Lean_addMacroScope(x_92, x_130, x_91); -x_132 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18276____closed__1; +x_132 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18519____closed__1; lean_inc(x_90); x_133 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_133, 0, x_90); @@ -30321,7 +30557,7 @@ x_155 = l_Lean_Syntax_mkApp___closed__2; lean_inc(x_90); x_156 = l_Lean_Syntax_node2(x_90, x_155, x_153, x_154); x_157 = l_Lean_Syntax_node5(x_90, x_8, x_94, x_96, x_98, x_156, x_116); -x_158 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____lambda__1(x_1, x_157, x_2, x_3); +x_158 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____lambda__1(x_1, x_157, x_2, x_3); lean_dec(x_2); return x_158; } @@ -30445,7 +30681,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_21536_(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_21779_(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; @@ -30506,12 +30742,12 @@ 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_expandSimp____x40_Init_Meta___hyg_19958____closed__4; +x_30 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__4; 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_19958____closed__2; -x_33 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__8; +x_32 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__2; +x_33 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__8; lean_inc(x_12); x_34 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_34, 0, x_12); @@ -30559,11 +30795,11 @@ lean_ctor_set(x_50, 0, x_12); lean_ctor_set(x_50, 1, x_49); lean_inc(x_12); x_51 = l_Lean_Syntax_node2(x_12, x_35, x_48, x_50); -x_52 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18276____closed__2; +x_52 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18519____closed__2; lean_inc(x_13); lean_inc(x_14); x_53 = l_Lean_addMacroScope(x_14, x_52, x_13); -x_54 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18276____closed__1; +x_54 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18519____closed__1; lean_inc(x_12); x_55 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_55, 0, x_12); @@ -30596,9 +30832,9 @@ lean_inc(x_12); x_66 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_66, 0, x_12); lean_ctor_set(x_66, 1, x_65); -x_67 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__12; +x_67 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__12; x_68 = l_Lean_addMacroScope(x_14, x_67, x_13); -x_69 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__11; +x_69 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__11; lean_inc(x_12); x_70 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_70, 0, x_12); @@ -30653,7 +30889,7 @@ x_90 = l_Lean_Syntax_mkApp___closed__2; lean_inc(x_12); x_91 = l_Lean_Syntax_node2(x_12, x_90, x_83, x_89); x_92 = l_Lean_Syntax_node5(x_12, x_8, x_16, x_18, x_20, x_91, x_38); -x_93 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____lambda__1(x_1, x_92, x_2, x_3); +x_93 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____lambda__1(x_1, x_92, x_2, x_3); lean_dec(x_2); return x_93; } @@ -30708,12 +30944,12 @@ lean_inc(x_98); x_115 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_115, 0, x_98); lean_ctor_set(x_115, 1, x_114); -x_116 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__4; +x_116 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__4; lean_inc(x_99); lean_inc(x_100); x_117 = l_Lean_addMacroScope(x_100, x_116, x_99); -x_118 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__2; -x_119 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__8; +x_118 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__2; +x_119 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__8; lean_inc(x_98); x_120 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_120, 0, x_98); @@ -30761,11 +30997,11 @@ lean_ctor_set(x_136, 0, x_98); lean_ctor_set(x_136, 1, x_135); lean_inc(x_98); x_137 = l_Lean_Syntax_node2(x_98, x_121, x_134, x_136); -x_138 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18276____closed__2; +x_138 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18519____closed__2; lean_inc(x_99); lean_inc(x_100); x_139 = l_Lean_addMacroScope(x_100, x_138, x_99); -x_140 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18276____closed__1; +x_140 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18519____closed__1; lean_inc(x_98); x_141 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_141, 0, x_98); @@ -30798,9 +31034,9 @@ lean_inc(x_98); x_152 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_152, 0, x_98); lean_ctor_set(x_152, 1, x_151); -x_153 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__12; +x_153 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__12; x_154 = l_Lean_addMacroScope(x_100, x_153, x_99); -x_155 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__11; +x_155 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__11; lean_inc(x_98); x_156 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_156, 0, x_98); @@ -30840,7 +31076,7 @@ x_171 = l_Lean_Syntax_mkApp___closed__2; lean_inc(x_98); x_172 = l_Lean_Syntax_node2(x_98, x_171, x_169, x_170); x_173 = l_Lean_Syntax_node5(x_98, x_8, x_102, x_104, x_106, x_172, x_124); -x_174 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____lambda__1(x_1, x_173, x_2, x_3); +x_174 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____lambda__1(x_1, x_173, x_2, x_3); lean_dec(x_2); return x_174; } @@ -30978,7 +31214,7 @@ x_1 = l_Lean_Parser_Tactic_dsimpAutoUnfold___closed__10; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____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_22652____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; 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; lean_object* x_19; lean_object* x_20; @@ -30986,7 +31222,7 @@ x_5 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tact 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___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__19; +x_9 = l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__19; x_10 = 1; x_11 = l_Lean_mkAtomFrom(x_8, x_9, x_10); x_12 = l_Lean_Syntax_setArg(x_6, x_7, x_11); @@ -31006,7 +31242,7 @@ lean_ctor_set(x_20, 1, x_4); return x_20; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____closed__1() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____closed__1() { _start: { lean_object* x_1; @@ -31014,16 +31250,16 @@ 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_22409____closed__2() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____closed__1; +x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____closed__1; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____closed__3() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____closed__3() { _start: { lean_object* x_1; @@ -31031,65 +31267,65 @@ 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_22409____closed__4() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____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_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__1; -x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__3; -x_3 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____closed__3; -x_4 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__5; +x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__3; +x_3 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____closed__3; +x_4 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__5; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____closed__5() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____closed__5() { _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_22409____closed__4; +x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____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_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____closed__6() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____closed__4; +x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____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_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____closed__7() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____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_22409____closed__6; +x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____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_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____closed__8() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____closed__5; -x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____closed__7; +x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____closed__5; +x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____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; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409_(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_22652_(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; @@ -31150,12 +31386,12 @@ 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_expandSimp____x40_Init_Meta___hyg_22409____closed__4; +x_30 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____closed__4; 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_22409____closed__2; -x_33 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____closed__8; +x_32 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____closed__2; +x_33 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____closed__8; lean_inc(x_12); x_34 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_34, 0, x_12); @@ -31203,11 +31439,11 @@ lean_ctor_set(x_50, 0, x_12); lean_ctor_set(x_50, 1, x_49); lean_inc(x_12); x_51 = l_Lean_Syntax_node2(x_12, x_35, x_48, x_50); -x_52 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__12; +x_52 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__12; lean_inc(x_13); lean_inc(x_14); x_53 = l_Lean_addMacroScope(x_14, x_52, x_13); -x_54 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__11; +x_54 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__11; lean_inc(x_12); x_55 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_55, 0, x_12); @@ -31274,7 +31510,7 @@ x_82 = l_Lean_Syntax_mkApp___closed__2; lean_inc(x_12); x_83 = l_Lean_Syntax_node2(x_12, x_82, x_75, x_81); x_84 = l_Lean_Syntax_node5(x_12, x_8, x_16, x_18, x_20, x_83, x_38); -x_85 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____lambda__1(x_1, x_84, x_2, x_3); +x_85 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____lambda__1(x_1, x_84, x_2, x_3); lean_dec(x_2); return x_85; } @@ -31329,12 +31565,12 @@ lean_inc(x_90); x_107 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_107, 0, x_90); lean_ctor_set(x_107, 1, x_106); -x_108 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____closed__4; +x_108 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____closed__4; lean_inc(x_91); lean_inc(x_92); x_109 = l_Lean_addMacroScope(x_92, x_108, x_91); -x_110 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____closed__2; -x_111 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____closed__8; +x_110 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____closed__2; +x_111 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____closed__8; lean_inc(x_90); x_112 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_112, 0, x_90); @@ -31382,11 +31618,11 @@ lean_ctor_set(x_128, 0, x_90); lean_ctor_set(x_128, 1, x_127); lean_inc(x_90); x_129 = l_Lean_Syntax_node2(x_90, x_113, x_126, x_128); -x_130 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__12; +x_130 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__12; lean_inc(x_91); lean_inc(x_92); x_131 = l_Lean_addMacroScope(x_92, x_130, x_91); -x_132 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__11; +x_132 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__11; lean_inc(x_90); x_133 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_133, 0, x_90); @@ -31438,17 +31674,17 @@ x_155 = l_Lean_Syntax_mkApp___closed__2; lean_inc(x_90); x_156 = l_Lean_Syntax_node2(x_90, x_155, x_153, x_154); x_157 = l_Lean_Syntax_node5(x_90, x_8, x_94, x_96, x_98, x_156, x_116); -x_158 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____lambda__1(x_1, x_157, x_2, x_3); +x_158 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____lambda__1(x_1, x_157, x_2, x_3); lean_dec(x_2); return x_158; } } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____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_22652____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_22409____lambda__1(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____lambda__1(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } @@ -31830,14 +32066,14 @@ l_Lean_Syntax_isCharLit_x3f___closed__2 = _init_l_Lean_Syntax_isCharLit_x3f___cl 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(); -lean_mark_persistent(l_List_foldr___at_Lean_Syntax_decodeNameLit___spec__1___closed__1); -l_List_foldr___at_Lean_Syntax_decodeNameLit___spec__1___closed__2 = _init_l_List_foldr___at_Lean_Syntax_decodeNameLit___spec__1___closed__2(); -lean_mark_persistent(l_List_foldr___at_Lean_Syntax_decodeNameLit___spec__1___closed__2); -l_List_foldr___at_Lean_Syntax_decodeNameLit___spec__1___closed__3 = _init_l_List_foldr___at_Lean_Syntax_decodeNameLit___spec__1___closed__3(); -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_List_foldr___at_Substring_toName___spec__1___closed__1 = _init_l_List_foldr___at_Substring_toName___spec__1___closed__1(); +lean_mark_persistent(l_List_foldr___at_Substring_toName___spec__1___closed__1); +l_List_foldr___at_Substring_toName___spec__1___closed__2 = _init_l_List_foldr___at_Substring_toName___spec__1___closed__2(); +lean_mark_persistent(l_List_foldr___at_Substring_toName___spec__1___closed__2); +l_List_foldr___at_Substring_toName___spec__1___closed__3 = _init_l_List_foldr___at_Substring_toName___spec__1___closed__3(); +lean_mark_persistent(l_List_foldr___at_Substring_toName___spec__1___closed__3); +l_List_foldr___at_Substring_toName___spec__1___closed__4 = _init_l_List_foldr___at_Substring_toName___spec__1___closed__4(); +lean_mark_persistent(l_List_foldr___at_Substring_toName___spec__1___closed__4); l_Lean_TSyntax_getScientific___closed__1 = _init_l_Lean_TSyntax_getScientific___closed__1(); lean_mark_persistent(l_Lean_TSyntax_getScientific___closed__1); l_Lean_TSyntax_getScientific___closed__2 = _init_l_Lean_TSyntax_getScientific___closed__2(); @@ -32037,54 +32273,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_12360____closed__1 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__1(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__1); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__2 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__2(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__2); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__3 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__3(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__3); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__4 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__4(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__4); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__5 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__5(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__5); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__6 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__6(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__6); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__7 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__7(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__7); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__8 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__8(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__8); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__9 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__9(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__9); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__10 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__10(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__10); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__11 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__11(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__11); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__12 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__12(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__12); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__13 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__13(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__13); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__14 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__14(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__14); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__15 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__15(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__15); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__16 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__16(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__16); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__17 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__17(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__17); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__18 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__18(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__18); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__19 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__19(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__19); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__20 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__20(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__20); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__21 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__21(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__21); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__22 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__22(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__22); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__23 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__23(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__23); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__24 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__24(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12360____closed__24); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__1 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__1(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__1); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__2 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__2(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__2); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__3 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__3(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__3); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__4 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__4(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__4); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__5 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__5(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__5); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__6 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__6(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__6); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__7 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__7(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__7); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__8 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__8(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__8); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__9 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__9(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__9); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__10 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__10(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__10); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__11 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__11(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__11); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__12 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__12(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__12); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__13 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__13(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__13); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__14 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__14(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__14); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__15 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__15(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__15); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__16 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__16(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__16); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__17 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__17(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__17); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__18 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__18(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__18); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__19 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__19(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__19); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__20 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__20(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__20); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__21 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__21(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__21); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__22 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__22(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__22); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__23 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__23(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__23); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__24 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____closed__24(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_12534____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(); @@ -32094,42 +32330,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_12565____closed__1 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__1(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__1); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__2 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__2(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__2); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__3 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__3(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__3); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__4 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__4(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__4); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__5 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__5(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__5); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__6 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__6(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__6); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__7 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__7(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__7); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__8 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__8(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__8); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__9 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__9(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__9); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__10 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__10(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__10); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__11 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__11(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__11); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__12 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__12(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__12); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__13 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__13(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__13); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__14 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__14(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__14); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__15 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__15(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__15); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__16 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__16(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__16); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__17 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__17(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__17); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__18 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__18(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12565____closed__18); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__1 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__1(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__1); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__2 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__2(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__2); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__3 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__3(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__3); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__4 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__4(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__4); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__5 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__5(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__5); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__6 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__6(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__6); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__7 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__7(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__7); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__8 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__8(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__8); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__9 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__9(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__9); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__10 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__10(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__10); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__11 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__11(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__11); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__12 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__12(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__12); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__13 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__13(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__13); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__14 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__14(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__14); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__15 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__15(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__15); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__16 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__16(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__16); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__17 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__17(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__17); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__18 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____closed__18(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_12739____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(); @@ -32150,62 +32386,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_13000____closed__1 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__1(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__1); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__2 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__2(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__2); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__3 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__3(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__3); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__4 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__4(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__4); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__5 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__5(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__5); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__6 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__6(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__6); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__7 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__7(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__7); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__8 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__8(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__8); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__9 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__9(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__9); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__10 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__10(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__10); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__11 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__11(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__11); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__12 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__12(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__12); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__13 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__13(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__13); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__14 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__14(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__14); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__15 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__15(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__15); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__16 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__16(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__16); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__17 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__17(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__17); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__18 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__18(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__18); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__19 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__19(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__19); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__20 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__20(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__20); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__21 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__21(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__21); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__22 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__22(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__22); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__23 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__23(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__23); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__24 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__24(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__24); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__25 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__25(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__25); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__26 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__26(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__26); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__27 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__27(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__27); -l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__28 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__28(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13000____closed__28); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__1 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__1(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__1); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__2 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__2(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__2); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__3 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__3(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__3); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__4 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__4(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__4); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__5 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__5(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__5); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__6 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__6(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__6); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__7 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__7(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__7); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__8 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__8(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__8); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__9 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__9(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__9); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__10 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__10(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__10); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__11 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__11(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__11); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__12 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__12(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__12); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__13 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__13(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__13); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__14 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__14(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__14); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__15 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__15(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__15); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__16 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__16(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__16); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__17 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__17(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__17); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__18 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__18(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__18); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__19 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__19(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__19); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__20 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__20(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__20); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__21 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__21(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__21); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__22 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__22(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__22); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__23 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__23(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__23); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__24 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__24(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__24); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__25 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__25(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__25); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__26 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__26(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__26); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__27 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__27(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__27); +l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__28 = _init_l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____closed__28(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_DSimp_reprConfig____x40_Init_Meta___hyg_13174____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(); @@ -32237,54 +32473,54 @@ 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_13703____closed__1 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__1(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__1); -l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__2 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__2(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__2); -l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__3 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__3(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__3); -l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__4 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__4(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__4); -l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__5 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__5(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__5); -l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__6 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__6(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__6); -l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__7 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__7(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__7); -l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__8 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__8(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__8); -l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__9 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__9(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__9); -l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__10 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__10(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__10); -l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__11 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__11(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__11); -l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__12 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__12(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__12); -l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__13 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__13(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__13); -l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__14 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__14(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__14); -l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__15 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__15(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__15); -l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__16 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__16(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__16); -l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__17 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__17(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__17); -l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__18 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__18(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__18); -l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__19 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__19(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__19); -l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__20 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__20(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__20); -l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__21 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__21(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__21); -l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__22 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__22(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__22); -l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__23 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__23(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__23); -l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__24 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__24(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13703____closed__24); +l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__1 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__1(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__1); +l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__2 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__2(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__2); +l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__3 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__3(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__3); +l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__4 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__4(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__4); +l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__5 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__5(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__5); +l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__6 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__6(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__6); +l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__7 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__7(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__7); +l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__8 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__8(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__8); +l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__9 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__9(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__9); +l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__10 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__10(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__10); +l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__11 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__11(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__11); +l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__12 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__12(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__12); +l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__13 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__13(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__13); +l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__14 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__14(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__14); +l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__15 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__15(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__15); +l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__16 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__16(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__16); +l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__17 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__17(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__17); +l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__18 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__18(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__18); +l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__19 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__19(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__19); +l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__20 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__20(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__20); +l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__21 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__21(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__21); +l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__22 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__22(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__22); +l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__23 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__23(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__23); +l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__24 = _init_l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__24(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_Simp_reprConfig____x40_Init_Meta___hyg_13877____closed__24); 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(); @@ -32414,6 +32650,8 @@ l_Lean_Parser_Tactic_simpAllKind___closed__15 = _init_l_Lean_Parser_Tactic_simpA lean_mark_persistent(l_Lean_Parser_Tactic_simpAllKind___closed__15); l_Lean_Parser_Tactic_simpAllKind___closed__16 = _init_l_Lean_Parser_Tactic_simpAllKind___closed__16(); lean_mark_persistent(l_Lean_Parser_Tactic_simpAllKind___closed__16); +l_Lean_Parser_Tactic_simpAllKind___closed__17 = _init_l_Lean_Parser_Tactic_simpAllKind___closed__17(); +lean_mark_persistent(l_Lean_Parser_Tactic_simpAllKind___closed__17); l_Lean_Parser_Tactic_simpAllKind = _init_l_Lean_Parser_Tactic_simpAllKind(); lean_mark_persistent(l_Lean_Parser_Tactic_simpAllKind); l_Lean_Parser_Tactic_dsimpKind___closed__1 = _init_l_Lean_Parser_Tactic_dsimpKind___closed__1(); @@ -32486,6 +32724,20 @@ l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__24 = _init_l_Lean_Parser_Ta lean_mark_persistent(l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__24); l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__25 = _init_l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__25(); lean_mark_persistent(l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__25); +l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__26 = _init_l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__26(); +lean_mark_persistent(l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__26); +l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__27 = _init_l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__27(); +lean_mark_persistent(l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__27); +l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__28 = _init_l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__28(); +lean_mark_persistent(l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__28); +l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__29 = _init_l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__29(); +lean_mark_persistent(l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__29); +l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__30 = _init_l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__30(); +lean_mark_persistent(l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__30); +l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__31 = _init_l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__31(); +lean_mark_persistent(l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__31); +l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__32 = _init_l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__32(); +lean_mark_persistent(l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__32); l_Lean_Parser_Tactic_declareSimpLikeTactic = _init_l_Lean_Parser_Tactic_declareSimpLikeTactic(); lean_mark_persistent(l_Lean_Parser_Tactic_declareSimpLikeTactic); l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__1 = _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__1(); @@ -33010,30 +33262,30 @@ l_Lean_Parser_Tactic_simpAutoUnfold___closed__26 = _init_l_Lean_Parser_Tactic_si lean_mark_persistent(l_Lean_Parser_Tactic_simpAutoUnfold___closed__26); 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_17473____closed__1 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__1(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__1); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__2 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__2(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__2); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__3 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__3(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__3); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__4 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__4(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__4); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__5 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__5(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__5); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__6 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__6(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__6); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__7 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__7(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__7); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__8 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__8(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__8); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__9 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__9(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__9); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__10 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__10(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__10); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__11 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__11(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__11); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__12 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__12(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17473____closed__12); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__1 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__1(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__1); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__2 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__2(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__2); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__3 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__3(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__3); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__4 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__4(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__4); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__5 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__5(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__5); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__6 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__6(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__6); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__7 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__7(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__7); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__8 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__8(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__8); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__9 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__9(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__9); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__10 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__10(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__10); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__11 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__11(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__11); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__12 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__12(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17716____closed__12); 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(); @@ -33056,10 +33308,10 @@ 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_18276____closed__1 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18276____closed__1(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18276____closed__1); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18276____closed__2 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18276____closed__2(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18276____closed__2); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18519____closed__1 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18519____closed__1(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18519____closed__1); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18519____closed__2 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18519____closed__2(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18519____closed__2); 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(); @@ -33110,24 +33362,24 @@ 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_19958____lambda__1___closed__1 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____lambda__1___closed__1(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____lambda__1___closed__1); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__1 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__1(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__1); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__2 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__2(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__2); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__3 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__3(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__3); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__4 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__4(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__4); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__5 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__5(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__5); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__6 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__6(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__6); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__7 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__7(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__7); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__8 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__8(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19958____closed__8); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____lambda__1___closed__1 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____lambda__1___closed__1); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__1 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__1(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__1); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__2 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__2(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__2); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__3 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__3(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__3); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__4 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__4(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__4); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__5 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__5(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__5); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__6 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__6(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__6); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__7 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__7(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__7); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__8 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__8(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20201____closed__8); 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(); @@ -33190,22 +33442,22 @@ 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_22409____closed__1 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____closed__1(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____closed__1); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____closed__2 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____closed__2(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____closed__2); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____closed__3 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____closed__3(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____closed__3); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____closed__4 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____closed__4(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____closed__4); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____closed__5 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____closed__5(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____closed__5); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____closed__6 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____closed__6(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____closed__6); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____closed__7 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____closed__7(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____closed__7); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____closed__8 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____closed__8(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22409____closed__8); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____closed__1 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____closed__1(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____closed__1); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____closed__2 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____closed__2(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____closed__2); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____closed__3 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____closed__3(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____closed__3); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____closed__4 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____closed__4(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____closed__4); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____closed__5 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____closed__5(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____closed__5); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____closed__6 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____closed__6(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____closed__6); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____closed__7 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____closed__7(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____closed__7); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____closed__8 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____closed__8(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_22652____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 76d0909591f1..adc53cd7c40a 100644 --- a/stage0/stdlib/Init/Notation.c +++ b/stage0/stdlib/Init/Notation.c @@ -67,7 +67,6 @@ LEAN_EXPORT lean_object* l_term___x3a_x3a__; LEAN_EXPORT lean_object* l_prioMid; LEAN_EXPORT lean_object* l___aux__Init__Notation______unexpand__Seq__seq__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_term_x7b___x3a___x2f_x2f___x7d___closed__19; -static lean_object* l_Lean_termThis___closed__3; LEAN_EXPORT lean_object* l___aux__Init__Notation______unexpand__HMod__hMod__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__16; LEAN_EXPORT lean_object* l___aux__Init__Notation______unexpand__HShiftLeft__hShiftLeft__1(lean_object*, lean_object*, lean_object*); @@ -120,7 +119,6 @@ static lean_object* l_termDepIfThenElse___closed__21; static lean_object* l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__4; LEAN_EXPORT lean_object* l_term___x26_x26__; static lean_object* l_termIfLet___closed__17; -lean_object* l_Lean_Syntax_getHeadInfo(lean_object*); static lean_object* l_term_x21_____closed__5; static lean_object* l_term___x3e_____closed__1; static lean_object* l_prioLow___closed__1; @@ -162,10 +160,10 @@ LEAN_EXPORT lean_object* l___aux__Init__Notation______macroRules__stx___x2c_x2a_ LEAN_EXPORT lean_object* l___aux__Init__Notation______macroRules__term___u2218____1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___aux__Init__Notation______unexpand__HAnd__hAnd__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_term_x7b___x3a___x2f_x2f___x7d___closed__17; -static lean_object* l_Lean_termThis___closed__1; static lean_object* l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__6; static lean_object* l_Lean_withAnnotateTerm___closed__2; LEAN_EXPORT lean_object* l___aux__Init__Notation______macroRules__termIfLet__1(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic_caseArg___closed__9; LEAN_EXPORT lean_object* l___aux__Init__Notation______unexpand__Eq__1(lean_object*, lean_object*, lean_object*); static lean_object* l_term___x3e_x3e_____closed__5; static lean_object* l_termDepIfThenElse___closed__16; @@ -297,7 +295,6 @@ static lean_object* l_precArg___closed__1; static lean_object* l_term___x7c_x7c_____closed__7; static lean_object* l_term___x7c_x3e_____closed__4; LEAN_EXPORT lean_object* l___aux__Init__Notation______macroRules__prioHigh__1(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_termThis___closed__4; lean_object* l_Lean_Syntax_node5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_term___x3c_x2a_____closed__5; static lean_object* l_rawNatLit___closed__7; @@ -358,7 +355,6 @@ static lean_object* l_termDepIfThenElse___closed__3; static lean_object* l___aux__Init__Notation______macroRules__term___x2d____1___closed__4; static lean_object* l_term_x7b___x3a___x2f_x2f___x7d___closed__7; static lean_object* l_term___u2208_____closed__3; -lean_object* lean_string_utf8_byte_size(lean_object*); static lean_object* l___aux__Init__Notation______macroRules__termWithout__expected__type____1___closed__2; static lean_object* l_term___x7c_x7c_____closed__2; static lean_object* l_Lean_withAnnotateTerm___closed__1; @@ -416,7 +412,6 @@ static lean_object* l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___ static lean_object* l___aux__Init__Notation______macroRules__term___x26_x26____1___closed__5; static lean_object* l_term_x25_x5b___x7c___x5d___closed__9; static lean_object* l_Lean_Parser_Syntax_addPrec___closed__8; -static lean_object* l_termDepIfThenElse___closed__35; static lean_object* l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__4; static lean_object* l___aux__Init__Notation______macroRules__term___x3c____1___closed__3; static lean_object* l_stx___x3f___closed__4; @@ -426,12 +421,10 @@ static lean_object* l_termIfThenElse___closed__12; static lean_object* l_termDepIfThenElse___closed__28; static lean_object* l_term_x25_x5b___x7c___x5d___closed__10; static lean_object* l_term_x25_x5b___x7c___x5d___closed__3; -static lean_object* l_Lean_termThis___closed__5; LEAN_EXPORT lean_object* l___aux__Init__Notation______unexpand__GT__gt__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_term___x3d__; static lean_object* l_term___x3c_x7c_x3e_____closed__4; static lean_object* l_stx___x2c_x2b___closed__2; -LEAN_EXPORT lean_object* l_Lean___aux__Init__Notation______macroRules__Lean__termThis__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_stx___x2c_x2b_x2c_x3f; static lean_object* l___aux__Init__Notation______unexpand__Function__comp__1___closed__3; static lean_object* l_term___x24_______closed__11; @@ -583,7 +576,6 @@ static lean_object* l_rawNatLit___closed__1; LEAN_EXPORT lean_object* l___aux__Init__Notation______macroRules__term___x26_x26_x26____1(lean_object*, lean_object*, lean_object*); static lean_object* l_termIfLet___closed__14; static lean_object* l_term___u2264_____closed__5; -static lean_object* l_Lean_termThis___closed__2; static lean_object* l___aux__Init__Notation______macroRules__term___x3c_x3d____2___closed__1; static lean_object* l___aux__Init__Notation______macroRules__term___u2218____1___closed__11; static lean_object* l_term___x3a_x3a_____closed__3; @@ -655,6 +647,7 @@ lean_object* l_Lean_Syntax_node3(lean_object*, lean_object*, lean_object*, lean_ static lean_object* l___aux__Init__Notation______macroRules__term___x7c_x7c____1___closed__6; static lean_object* l_Lean_Parser_Syntax_addPrec___closed__10; static lean_object* l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__2; +static lean_object* l_term_x25_x5b___x7c___x5d___closed__12; LEAN_EXPORT lean_object* l_term___x2b__; static lean_object* l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__11; static lean_object* l___aux__Init__Notation______macroRules__term___x26_x26_x26____1___closed__1; @@ -665,7 +658,6 @@ static lean_object* l_term_x21_____closed__2; static lean_object* l_term___xd7_____closed__4; static lean_object* l___aux__Init__Notation______macroRules__term___x3c_x7c_x3e____1___closed__9; LEAN_EXPORT lean_object* l_term___u2264__; -LEAN_EXPORT lean_object* l_Lean___aux__Init__Notation______macroRules__Lean__termThis__1___boxed(lean_object*, lean_object*, lean_object*); 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___x3e_x3e_x3d_____closed__6; LEAN_EXPORT lean_object* l___aux__Init__Notation______unexpand__HMul__hMul__1(lean_object*, lean_object*, lean_object*); @@ -725,6 +717,7 @@ static lean_object* l___aux__Init__Notation______macroRules__term___x7c_x7c____1 LEAN_EXPORT lean_object* l___aux__Init__Notation______macroRules__precArg__1(lean_object*, lean_object*, lean_object*); static lean_object* l___aux__Init__Notation______macroRules__term___x2f_x5c____1___closed__4; static lean_object* l_termDepIfThenElse___closed__2; +static lean_object* l_Lean_Parser_Tactic_caseArg___closed__7; LEAN_EXPORT lean_object* l_prec_x28___x29; static lean_object* l___aux__Init__Notation______macroRules__term___x3c_x24_x3e____1___closed__1; static lean_object* l_termDepIfThenElse___closed__19; @@ -789,6 +782,7 @@ static lean_object* l_term_x5b___x5d___closed__8; LEAN_EXPORT lean_object* l_Lean_deprecated; static lean_object* l_term_xac_____closed__3; static lean_object* l_termIfLet___closed__19; +static lean_object* l_Lean_Parser_Tactic_caseArg___closed__8; static lean_object* l___aux__Init__Notation______macroRules__term___x7c_x7c_x7c____1___closed__5; static lean_object* l___aux__Init__Notation______macroRules__boolIfThenElse__1___closed__6; static lean_object* l_Lean_Parser_Syntax_subPrio___closed__3; @@ -1026,7 +1020,6 @@ static lean_object* l___aux__Init__Notation______macroRules__term___x5e_x5e_x5e_ lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); static lean_object* l_term___x26_x26_x26_____closed__5; static lean_object* l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__9; -static lean_object* l_termDepIfThenElse___closed__37; static lean_object* l_term___x3c_____closed__6; LEAN_EXPORT lean_object* l___aux__Init__Notation______unexpand__Function__comp__1(lean_object*, lean_object*, lean_object*); static lean_object* l___aux__Init__Notation______macroRules__term___x3d_x3d____1___closed__1; @@ -1076,7 +1069,6 @@ static lean_object* l_stx___x3c_x7c_x3e_____closed__5; LEAN_EXPORT lean_object* l___aux__Init__Notation______unexpand__HAdd__hAdd__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___aux__Init__Notation______unexpand__Function__comp__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___aux__Init__Notation______macroRules__term___x5e_x5e_x5e____1(lean_object*, lean_object*, lean_object*); -static lean_object* l_termDepIfThenElse___closed__36; 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___x2d____1___closed__3; LEAN_EXPORT lean_object* l___aux__Init__Notation______macroRules__precLead__1(lean_object*, lean_object*, lean_object*); @@ -1107,7 +1099,6 @@ static lean_object* l___aux__Init__Notation______macroRules__term___x7c_x7c____1 static lean_object* l_term___x3e_x3e_x3e_____closed__1; static lean_object* l_term___x5e_x5e_x5e_____closed__1; static lean_object* l_stx___x2a___closed__4; -static lean_object* l_Lean___aux__Init__Notation______macroRules__Lean__termThis__1___closed__1; static lean_object* l___aux__Init__Notation______macroRules__term___x2f____1___closed__6; static lean_object* l_term___u2227_____closed__4; static lean_object* l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__3; @@ -1131,7 +1122,6 @@ LEAN_EXPORT lean_object* l___aux__Init__Notation______macroRules__term___x7c_x7c static lean_object* l_term___x2a_____closed__5; static lean_object* l_stx___x2c_x2a_x2c_x3f___closed__3; static lean_object* l___aux__Init__Notation______macroRules__term___x3e_x3e_x3e____1___closed__4; -static lean_object* l_Lean___aux__Init__Notation______macroRules__Lean__termThis__1___closed__2; LEAN_EXPORT lean_object* l___aux__Init__Notation______unexpand__HAnd__hAnd__1(lean_object*, lean_object*, lean_object*); static lean_object* l___aux__Init__Notation______macroRules__term___x3a_x3a____1___closed__4; static lean_object* l_stx___x3f___closed__3; @@ -1224,6 +1214,7 @@ static lean_object* l_term_x21_____closed__1; static lean_object* l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__6; static lean_object* l_Lean_Parser_Syntax_addPrec___closed__6; LEAN_EXPORT lean_object* l___aux__Init__Notation______unexpand__Not__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic_caseArg___closed__10; static lean_object* l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__7; LEAN_EXPORT lean_object* l___aux__Init__Notation______unexpand__HMul__hMul__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_term_x7e_x7e_x7e_____closed__6; @@ -1275,7 +1266,6 @@ static lean_object* l___aux__Init__Notation______macroRules__term___u2209____1__ static lean_object* l_Lean_deprecated___closed__2; static lean_object* l_termIfLet___closed__13; static lean_object* l_term___x24_______closed__2; -static lean_object* l_Lean___aux__Init__Notation______macroRules__Lean__termThis__1___closed__3; static lean_object* l_term___x2b_____closed__1; static lean_object* l_term___u2228_____closed__2; static lean_object* l_term___x3c_x3c_x3c_____closed__2; @@ -1291,7 +1281,6 @@ static lean_object* l_termIfLet___closed__1; static lean_object* l___aux__Init__Notation______macroRules__term___x5e____1___closed__2; static lean_object* l_termDepIfThenElse___closed__7; static lean_object* l_Lean_deprecated___closed__5; -LEAN_EXPORT lean_object* l_Lean_termThis; static lean_object* l_termWithout__expected__type_____closed__1; static lean_object* l_term_x7b___x3a___x2f_x2f___x7d___closed__1; uint8_t lean_nat_dec_le(lean_object*, lean_object*); @@ -22131,22 +22120,64 @@ return x_5; static lean_object* _init_l_Lean_Parser_Tactic_caseArg___closed__4() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("ppSpace", 7); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_caseArg___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Parser_Tactic_caseArg___closed__4; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_caseArg___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Tactic_caseArg___closed__5; +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_caseArg___closed__7() { +_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__7; +x_2 = l_Lean_Parser_Tactic_caseArg___closed__6; +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); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_caseArg___closed__8() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___aux__Init__Notation______macroRules__stx___x2a__1___closed__3; -x_2 = l_Lean_binderIdent; +x_2 = l_Lean_Parser_Tactic_caseArg___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_caseArg___closed__5() { +static lean_object* _init_l_Lean_Parser_Tactic_caseArg___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__7; x_2 = l_Lean_binderIdent; -x_3 = l_Lean_Parser_Tactic_caseArg___closed__4; +x_3 = l_Lean_Parser_Tactic_caseArg___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); @@ -22154,13 +22185,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_caseArg___closed__6() { +static lean_object* _init_l_Lean_Parser_Tactic_caseArg___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_caseArg___closed__2; x_2 = l_Lean_Parser_Tactic_caseArg___closed__3; -x_3 = l_Lean_Parser_Tactic_caseArg___closed__5; +x_3 = l_Lean_Parser_Tactic_caseArg___closed__9; x_4 = lean_alloc_ctor(9, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -22172,7 +22203,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_caseArg() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Tactic_caseArg___closed__6; +x_1 = l_Lean_Parser_Tactic_caseArg___closed__10; return x_1; } } @@ -22385,38 +22416,10 @@ return x_3; static lean_object* _init_l_termDepIfThenElse___closed__21() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("ppSpace", 7); -return x_1; -} -} -static lean_object* _init_l_termDepIfThenElse___closed__22() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_termDepIfThenElse___closed__21; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_termDepIfThenElse___closed__23() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_termDepIfThenElse___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_termDepIfThenElse___closed__24() { -_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__7; x_2 = l_termDepIfThenElse___closed__20; -x_3 = l_termDepIfThenElse___closed__23; +x_3 = l_Lean_Parser_Tactic_caseArg___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); @@ -22424,12 +22427,12 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_termDepIfThenElse___closed__25() { +static lean_object* _init_l_termDepIfThenElse___closed__22() { _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__7; -x_2 = l_termDepIfThenElse___closed__24; +x_2 = l_termDepIfThenElse___closed__21; x_3 = l_termDepIfThenElse___closed__15; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -22438,19 +22441,19 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_termDepIfThenElse___closed__26() { +static lean_object* _init_l_termDepIfThenElse___closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_termDepIfThenElse___closed__6; -x_2 = l_termDepIfThenElse___closed__25; +x_2 = l_termDepIfThenElse___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_termDepIfThenElse___closed__27() { +static lean_object* _init_l_termDepIfThenElse___closed__24() { _start: { lean_object* x_1; @@ -22458,35 +22461,35 @@ x_1 = lean_mk_string_from_bytes("ppDedent", 8); return x_1; } } -static lean_object* _init_l_termDepIfThenElse___closed__28() { +static lean_object* _init_l_termDepIfThenElse___closed__25() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_termDepIfThenElse___closed__27; +x_2 = l_termDepIfThenElse___closed__24; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_termDepIfThenElse___closed__29() { +static lean_object* _init_l_termDepIfThenElse___closed__26() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_termDepIfThenElse___closed__28; -x_2 = l_termDepIfThenElse___closed__23; +x_1 = l_termDepIfThenElse___closed__25; +x_2 = l_Lean_Parser_Tactic_caseArg___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_termDepIfThenElse___closed__30() { +static lean_object* _init_l_termDepIfThenElse___closed__27() { _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__7; -x_2 = l_termDepIfThenElse___closed__26; -x_3 = l_termDepIfThenElse___closed__29; +x_2 = l_termDepIfThenElse___closed__23; +x_3 = l_termDepIfThenElse___closed__26; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -22494,7 +22497,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_termDepIfThenElse___closed__31() { +static lean_object* _init_l_termDepIfThenElse___closed__28() { _start: { lean_object* x_1; @@ -22502,22 +22505,22 @@ x_1 = lean_mk_string_from_bytes("else ", 5); return x_1; } } -static lean_object* _init_l_termDepIfThenElse___closed__32() { +static lean_object* _init_l_termDepIfThenElse___closed__29() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_termDepIfThenElse___closed__31; +x_1 = l_termDepIfThenElse___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_termDepIfThenElse___closed__33() { +static lean_object* _init_l_termDepIfThenElse___closed__30() { _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__7; -x_2 = l_termDepIfThenElse___closed__32; +x_2 = l_termDepIfThenElse___closed__29; x_3 = l_termDepIfThenElse___closed__15; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -22526,25 +22529,25 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_termDepIfThenElse___closed__34() { +static lean_object* _init_l_termDepIfThenElse___closed__31() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_termDepIfThenElse___closed__6; -x_2 = l_termDepIfThenElse___closed__33; +x_2 = l_termDepIfThenElse___closed__30; 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_termDepIfThenElse___closed__35() { +static lean_object* _init_l_termDepIfThenElse___closed__32() { _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__7; -x_2 = l_termDepIfThenElse___closed__30; -x_3 = l_termDepIfThenElse___closed__34; +x_2 = l_termDepIfThenElse___closed__27; +x_3 = l_termDepIfThenElse___closed__31; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -22552,25 +22555,25 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_termDepIfThenElse___closed__36() { +static lean_object* _init_l_termDepIfThenElse___closed__33() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_termDepIfThenElse___closed__4; -x_2 = l_termDepIfThenElse___closed__35; +x_2 = l_termDepIfThenElse___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_termDepIfThenElse___closed__37() { +static lean_object* _init_l_termDepIfThenElse___closed__34() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_termDepIfThenElse___closed__2; x_2 = lean_unsigned_to_nat(1022u); -x_3 = l_termDepIfThenElse___closed__36; +x_3 = l_termDepIfThenElse___closed__33; x_4 = lean_alloc_ctor(3, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -22582,7 +22585,7 @@ static lean_object* _init_l_termDepIfThenElse() { _start: { lean_object* x_1; -x_1 = l_termDepIfThenElse___closed__37; +x_1 = l_termDepIfThenElse___closed__34; return x_1; } } @@ -23220,7 +23223,7 @@ static lean_object* _init_l_termIfThenElse___closed__6() { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Syntax_addPrec___closed__7; x_2 = l_termIfThenElse___closed__5; -x_3 = l_termDepIfThenElse___closed__23; +x_3 = l_Lean_Parser_Tactic_caseArg___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); @@ -23260,7 +23263,7 @@ static lean_object* _init_l_termIfThenElse___closed__9() { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Syntax_addPrec___closed__7; x_2 = l_termIfThenElse___closed__8; -x_3 = l_termDepIfThenElse___closed__29; +x_3 = l_termDepIfThenElse___closed__26; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -23274,7 +23277,7 @@ static lean_object* _init_l_termIfThenElse___closed__10() { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Syntax_addPrec___closed__7; x_2 = l_termIfThenElse___closed__9; -x_3 = l_termDepIfThenElse___closed__34; +x_3 = l_termDepIfThenElse___closed__31; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -23651,7 +23654,7 @@ static lean_object* _init_l_termIfLet___closed__13() { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Syntax_addPrec___closed__7; x_2 = l_termIfLet___closed__12; -x_3 = l_termDepIfThenElse___closed__23; +x_3 = l_Lean_Parser_Tactic_caseArg___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); @@ -23691,7 +23694,7 @@ static lean_object* _init_l_termIfLet___closed__16() { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Syntax_addPrec___closed__7; x_2 = l_termIfLet___closed__15; -x_3 = l_termDepIfThenElse___closed__29; +x_3 = l_termDepIfThenElse___closed__26; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -23705,7 +23708,7 @@ static lean_object* _init_l_termIfLet___closed__17() { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Syntax_addPrec___closed__7; x_2 = l_termIfLet___closed__16; -x_3 = l_termDepIfThenElse___closed__34; +x_3 = l_termDepIfThenElse___closed__31; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -24031,7 +24034,7 @@ static lean_object* _init_l_boolIfThenElse___closed__8() { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Syntax_addPrec___closed__7; x_2 = l_boolIfThenElse___closed__7; -x_3 = l_termDepIfThenElse___closed__23; +x_3 = l_Lean_Parser_Tactic_caseArg___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); @@ -24071,7 +24074,7 @@ static lean_object* _init_l_boolIfThenElse___closed__11() { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Syntax_addPrec___closed__7; x_2 = l_boolIfThenElse___closed__10; -x_3 = l_termDepIfThenElse___closed__29; +x_3 = l_termDepIfThenElse___closed__26; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -24085,7 +24088,7 @@ static lean_object* _init_l_boolIfThenElse___closed__12() { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Syntax_addPrec___closed__7; x_2 = l_boolIfThenElse___closed__11; -x_3 = l_termDepIfThenElse___closed__34; +x_3 = l_termDepIfThenElse___closed__31; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -25789,20 +25792,28 @@ return x_2; static lean_object* _init_l_term_x25_x5b___x7c___x5d___closed__5() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(" | ", 3); +return x_1; +} +} +static lean_object* _init_l_term_x25_x5b___x7c___x5d___closed__6() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Init__Notation______macroRules__termIfLet__1___closed__10; +x_1 = l_term_x25_x5b___x7c___x5d___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_term_x25_x5b___x7c___x5d___closed__6() { +static lean_object* _init_l_term_x25_x5b___x7c___x5d___closed__7() { _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__7; x_2 = l_term_x5b___x5d___closed__7; -x_3 = l_term_x25_x5b___x7c___x5d___closed__5; +x_3 = l_term_x25_x5b___x7c___x5d___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); @@ -25810,12 +25821,12 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_term_x25_x5b___x7c___x5d___closed__7() { +static lean_object* _init_l_term_x25_x5b___x7c___x5d___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__7; -x_2 = l_term_x25_x5b___x7c___x5d___closed__6; +x_2 = l_term_x25_x5b___x7c___x5d___closed__7; x_3 = l_termDepIfThenElse___closed__15; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -25824,25 +25835,25 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_term_x25_x5b___x7c___x5d___closed__8() { +static lean_object* _init_l_term_x25_x5b___x7c___x5d___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_term_x7b___x3a___x2f_x2f___x7d___closed__6; -x_2 = l_term_x25_x5b___x7c___x5d___closed__7; +x_2 = l_term_x25_x5b___x7c___x5d___closed__8; 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_x25_x5b___x7c___x5d___closed__9() { +static lean_object* _init_l_term_x25_x5b___x7c___x5d___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__7; x_2 = l_term_x25_x5b___x7c___x5d___closed__4; -x_3 = l_term_x25_x5b___x7c___x5d___closed__8; +x_3 = l_term_x25_x5b___x7c___x5d___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); @@ -25850,12 +25861,12 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_term_x25_x5b___x7c___x5d___closed__10() { +static lean_object* _init_l_term_x25_x5b___x7c___x5d___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__7; -x_2 = l_term_x25_x5b___x7c___x5d___closed__9; +x_2 = l_term_x25_x5b___x7c___x5d___closed__10; x_3 = l_term_x5b___x5d___closed__11; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -25864,13 +25875,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_term_x25_x5b___x7c___x5d___closed__11() { +static lean_object* _init_l_term_x25_x5b___x7c___x5d___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_term_x25_x5b___x7c___x5d___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_term_x25_x5b___x7c___x5d___closed__10; +x_3 = l_term_x25_x5b___x7c___x5d___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); @@ -25882,7 +25893,7 @@ static lean_object* _init_l_term_x25_x5b___x7c___x5d() { _start: { lean_object* x_1; -x_1 = l_term_x25_x5b___x7c___x5d___closed__11; +x_1 = l_term_x25_x5b___x7c___x5d___closed__12; return x_1; } } @@ -26133,146 +26144,6 @@ return x_47; } } } -static lean_object* _init_l_Lean_termThis___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("termThis", 8); -return x_1; -} -} -static lean_object* _init_l_Lean_termThis___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Syntax_addPrec___closed__1; -x_2 = l_Lean_termThis___closed__1; -x_3 = l_Lean_Name_mkStr2(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_termThis___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("this", 4); -return x_1; -} -} -static lean_object* _init_l_Lean_termThis___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_termThis___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_termThis___closed__5() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termThis___closed__2; -x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_termThis___closed__4; -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_termThis() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_termThis___closed__5; -return x_1; -} -} -static lean_object* _init_l_Lean___aux__Init__Notation______macroRules__Lean__termThis__1___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_termThis___closed__3; -x_2 = lean_string_utf8_byte_size(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean___aux__Init__Notation______macroRules__Lean__termThis__1___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termThis___closed__3; -x_2 = lean_unsigned_to_nat(0u); -x_3 = l_Lean___aux__Init__Notation______macroRules__Lean__termThis__1___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); -lean_ctor_set(x_4, 2, x_3); -return x_4; -} -} -static lean_object* _init_l_Lean___aux__Init__Notation______macroRules__Lean__termThis__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_termThis___closed__3; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Lean___aux__Init__Notation______macroRules__Lean__termThis__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_termThis___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_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; -x_8 = lean_unsigned_to_nat(0u); -x_9 = l_Lean_Syntax_getArg(x_1, x_8); -lean_dec(x_1); -x_10 = l_Lean_Syntax_getHeadInfo(x_9); -lean_dec(x_9); -x_11 = lean_box(0); -x_12 = l_Lean___aux__Init__Notation______macroRules__Lean__termThis__1___closed__2; -x_13 = l_Lean___aux__Init__Notation______macroRules__Lean__termThis__1___closed__3; -x_14 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_14, 0, x_10); -lean_ctor_set(x_14, 1, x_12); -lean_ctor_set(x_14, 2, x_13); -lean_ctor_set(x_14, 3, x_11); -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; -} -} -} -LEAN_EXPORT lean_object* l_Lean___aux__Init__Notation______macroRules__Lean__termThis__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; -x_4 = l_Lean___aux__Init__Notation______macroRules__Lean__termThis__1(x_1, x_2, x_3); -lean_dec(x_2); -return x_4; -} -} static lean_object* _init_l_Lean_rawStx_quot___closed__1() { _start: { @@ -26495,7 +26366,7 @@ static lean_object* _init_l_Lean_withAnnotateTerm___closed__6() { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Syntax_addPrec___closed__7; x_2 = l_Lean_withAnnotateTerm___closed__5; -x_3 = l_termDepIfThenElse___closed__23; +x_3 = l_Lean_Parser_Tactic_caseArg___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); @@ -26560,16 +26431,8 @@ return x_3; static lean_object* _init_l_Lean_deprecated___closed__3() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("deprecated ", 11); -return x_1; -} -} -static lean_object* _init_l_Lean_deprecated___closed__4() { -_start: -{ lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l_Lean_deprecated___closed__3; +x_1 = l_Lean_deprecated___closed__1; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -26577,12 +26440,26 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } +static lean_object* _init_l_Lean_deprecated___closed__4() { +_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__7; +x_2 = l_Lean_Parser_Tactic_caseArg___closed__6; +x_3 = l_Lean_binderIdent___closed__3; +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_deprecated___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___aux__Init__Notation______macroRules__stx___x3f__1___closed__3; -x_2 = l_Lean_binderIdent___closed__3; +x_2 = l_Lean_deprecated___closed__4; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -26594,7 +26471,7 @@ static lean_object* _init_l_Lean_deprecated___closed__6() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Syntax_addPrec___closed__7; -x_2 = l_Lean_deprecated___closed__4; +x_2 = l_Lean_deprecated___closed__3; x_3 = l_Lean_deprecated___closed__5; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -26647,7 +26524,7 @@ static lean_object* _init_l_Lean_includeStr___closed__3() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("include_str", 11); +x_1 = lean_mk_string_from_bytes("include_str ", 12); return x_1; } } @@ -28432,6 +28309,14 @@ l_Lean_Parser_Tactic_caseArg___closed__5 = _init_l_Lean_Parser_Tactic_caseArg___ lean_mark_persistent(l_Lean_Parser_Tactic_caseArg___closed__5); l_Lean_Parser_Tactic_caseArg___closed__6 = _init_l_Lean_Parser_Tactic_caseArg___closed__6(); lean_mark_persistent(l_Lean_Parser_Tactic_caseArg___closed__6); +l_Lean_Parser_Tactic_caseArg___closed__7 = _init_l_Lean_Parser_Tactic_caseArg___closed__7(); +lean_mark_persistent(l_Lean_Parser_Tactic_caseArg___closed__7); +l_Lean_Parser_Tactic_caseArg___closed__8 = _init_l_Lean_Parser_Tactic_caseArg___closed__8(); +lean_mark_persistent(l_Lean_Parser_Tactic_caseArg___closed__8); +l_Lean_Parser_Tactic_caseArg___closed__9 = _init_l_Lean_Parser_Tactic_caseArg___closed__9(); +lean_mark_persistent(l_Lean_Parser_Tactic_caseArg___closed__9); +l_Lean_Parser_Tactic_caseArg___closed__10 = _init_l_Lean_Parser_Tactic_caseArg___closed__10(); +lean_mark_persistent(l_Lean_Parser_Tactic_caseArg___closed__10); l_Lean_Parser_Tactic_caseArg = _init_l_Lean_Parser_Tactic_caseArg(); lean_mark_persistent(l_Lean_Parser_Tactic_caseArg); l_termDepIfThenElse___closed__1 = _init_l_termDepIfThenElse___closed__1(); @@ -28502,12 +28387,6 @@ l_termDepIfThenElse___closed__33 = _init_l_termDepIfThenElse___closed__33(); lean_mark_persistent(l_termDepIfThenElse___closed__33); l_termDepIfThenElse___closed__34 = _init_l_termDepIfThenElse___closed__34(); lean_mark_persistent(l_termDepIfThenElse___closed__34); -l_termDepIfThenElse___closed__35 = _init_l_termDepIfThenElse___closed__35(); -lean_mark_persistent(l_termDepIfThenElse___closed__35); -l_termDepIfThenElse___closed__36 = _init_l_termDepIfThenElse___closed__36(); -lean_mark_persistent(l_termDepIfThenElse___closed__36); -l_termDepIfThenElse___closed__37 = _init_l_termDepIfThenElse___closed__37(); -lean_mark_persistent(l_termDepIfThenElse___closed__37); l_termDepIfThenElse = _init_l_termDepIfThenElse(); lean_mark_persistent(l_termDepIfThenElse); l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__1 = _init_l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__1(); @@ -28894,6 +28773,8 @@ l_term_x25_x5b___x7c___x5d___closed__10 = _init_l_term_x25_x5b___x7c___x5d___clo lean_mark_persistent(l_term_x25_x5b___x7c___x5d___closed__10); l_term_x25_x5b___x7c___x5d___closed__11 = _init_l_term_x25_x5b___x7c___x5d___closed__11(); lean_mark_persistent(l_term_x25_x5b___x7c___x5d___closed__11); +l_term_x25_x5b___x7c___x5d___closed__12 = _init_l_term_x25_x5b___x7c___x5d___closed__12(); +lean_mark_persistent(l_term_x25_x5b___x7c___x5d___closed__12); l_term_x25_x5b___x7c___x5d = _init_l_term_x25_x5b___x7c___x5d(); lean_mark_persistent(l_term_x25_x5b___x7c___x5d); l_Lean___aux__Init__Notation______macroRules__term_x5b___x5d__1___closed__1 = _init_l_Lean___aux__Init__Notation______macroRules__term_x5b___x5d__1___closed__1(); @@ -28908,24 +28789,6 @@ l_Lean___aux__Init__Notation______macroRules__term_x5b___x5d__1___closed__5 = _i lean_mark_persistent(l_Lean___aux__Init__Notation______macroRules__term_x5b___x5d__1___closed__5); l_Lean___aux__Init__Notation______macroRules__term_x5b___x5d__1___closed__6 = _init_l_Lean___aux__Init__Notation______macroRules__term_x5b___x5d__1___closed__6(); lean_mark_persistent(l_Lean___aux__Init__Notation______macroRules__term_x5b___x5d__1___closed__6); -l_Lean_termThis___closed__1 = _init_l_Lean_termThis___closed__1(); -lean_mark_persistent(l_Lean_termThis___closed__1); -l_Lean_termThis___closed__2 = _init_l_Lean_termThis___closed__2(); -lean_mark_persistent(l_Lean_termThis___closed__2); -l_Lean_termThis___closed__3 = _init_l_Lean_termThis___closed__3(); -lean_mark_persistent(l_Lean_termThis___closed__3); -l_Lean_termThis___closed__4 = _init_l_Lean_termThis___closed__4(); -lean_mark_persistent(l_Lean_termThis___closed__4); -l_Lean_termThis___closed__5 = _init_l_Lean_termThis___closed__5(); -lean_mark_persistent(l_Lean_termThis___closed__5); -l_Lean_termThis = _init_l_Lean_termThis(); -lean_mark_persistent(l_Lean_termThis); -l_Lean___aux__Init__Notation______macroRules__Lean__termThis__1___closed__1 = _init_l_Lean___aux__Init__Notation______macroRules__Lean__termThis__1___closed__1(); -lean_mark_persistent(l_Lean___aux__Init__Notation______macroRules__Lean__termThis__1___closed__1); -l_Lean___aux__Init__Notation______macroRules__Lean__termThis__1___closed__2 = _init_l_Lean___aux__Init__Notation______macroRules__Lean__termThis__1___closed__2(); -lean_mark_persistent(l_Lean___aux__Init__Notation______macroRules__Lean__termThis__1___closed__2); -l_Lean___aux__Init__Notation______macroRules__Lean__termThis__1___closed__3 = _init_l_Lean___aux__Init__Notation______macroRules__Lean__termThis__1___closed__3(); -lean_mark_persistent(l_Lean___aux__Init__Notation______macroRules__Lean__termThis__1___closed__3); l_Lean_rawStx_quot___closed__1 = _init_l_Lean_rawStx_quot___closed__1(); lean_mark_persistent(l_Lean_rawStx_quot___closed__1); l_Lean_rawStx_quot___closed__2 = _init_l_Lean_rawStx_quot___closed__2(); diff --git a/stage0/stdlib/Init/NotationExtra.c b/stage0/stdlib/Init/NotationExtra.c index 45a22554d032..b8ef77575aac 100644 --- a/stage0/stdlib/Init/NotationExtra.c +++ b/stage0/stdlib/Init/NotationExtra.c @@ -14,18 +14,22 @@ extern "C" { #endif static lean_object* l_Lean_unbracketedExplicitBinders___closed__11; -static lean_object* l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__21; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__23; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__18; static lean_object* l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___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__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_termExists___x2c_____closed__4; static lean_object* l_term_u2203___x2c_____closed__6; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemWhile___x3a__Do____1___closed__5; +static lean_object* l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__43; +static lean_object* l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__26; static lean_object* l_unexpandEqNDRec___closed__2; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__16; static lean_object* l_solve___closed__2; static lean_object* l_Lean_unbracketedExplicitBinders___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Command_classAbbrev; lean_object* l___private_Init_Util_0__outOfBounds___rarg(lean_object*); +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__22; +static lean_object* l_solve___closed__18; static lean_object* l_cdotTk___closed__3; static lean_object* l_Lean_explicitBinders___closed__4; static lean_object* l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__8; @@ -33,49 +37,48 @@ LEAN_EXPORT lean_object* l_unexpandMkArray5(lean_object*, lean_object*, lean_obj static lean_object* l_term___xd7____1___closed__2; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__11; 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___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__2; static lean_object* l_calcTactic___closed__3; static lean_object* l_Lean_doElemWhile___x3a__Do_____closed__3; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____1___closed__15; lean_object* l_Lean_Syntax_mkNameLit(lean_object*, lean_object*); -static lean_object* l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__26; static lean_object* l_term_u03a3___x2c_____closed__6; static lean_object* l_Lean_unbracketedExplicitBinders___closed__7; static lean_object* l_Lean_unbracketedExplicitBinders___closed__3; static lean_object* l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__8; static lean_object* l_Lean_doElemWhile___x3a__Do_____closed__10; LEAN_EXPORT lean_object* l_unexpandMkStr4(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__8; LEAN_EXPORT lean_object* l_Lean_Loop_toCtorIdx___boxed(lean_object*); LEAN_EXPORT lean_object* l_unexpandEqNDRec(lean_object*, lean_object*, lean_object*); static lean_object* l_termExists___x2c_____closed__7; +LEAN_EXPORT lean_object* l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2__; +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*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__25; -static lean_object* l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__29; static lean_object* l_calcFirstStep___closed__11; static lean_object* l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__4; static lean_object* l_term_u2203___x2c_____closed__3; -LEAN_EXPORT lean_object* l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2__; -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_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__13; LEAN_EXPORT lean_object* l_unexpandListToArray(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__37; static lean_object* l_Lean_bracketedExplicitBinders___closed__7; 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__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__24; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__44; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__21; static lean_object* l_unexpandUnit___rarg___closed__2; static lean_object* l_Lean_expandExplicitBindersAux_loop___closed__10; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__33; -static lean_object* l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__32; +LEAN_EXPORT lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__1(lean_object*); static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__5; static lean_object* l_cdotTk___closed__8; static lean_object* l_term___xd7____1___closed__6; static lean_object* l_term___xd7_x27_____closed__6; static lean_object* l_unexpandMkArray6___closed__1; LEAN_EXPORT lean_object* l_unexpandMkArray2(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__18; +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_termExists___x2c_____closed__3; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__19; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__20; LEAN_EXPORT lean_object* l_unexpandArrayEmpty(lean_object*); static lean_object* l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__2; -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___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__16; static lean_object* l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__7; LEAN_EXPORT lean_object* l_unexpandMkStr6___boxed(lean_object*, lean_object*, lean_object*); @@ -83,17 +86,17 @@ static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__ter static lean_object* l_Lean_termMacro_x2etrace_x5b___x5d_____closed__12; static lean_object* l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__1; static lean_object* l_Lean_unbracketedExplicitBinders___closed__13; -static lean_object* l_Lean_doElemRepeat__Until_____closed__7; LEAN_EXPORT lean_object* l_term_u03a3_x27___x2c__; 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___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__28; static lean_object* l_unexpandGetElem___closed__2; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__15; static lean_object* l_calcSteps___closed__5; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__23; static lean_object* l_cdot___closed__6; +static lean_object* l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__45; static lean_object* l_unexpandIte___closed__2; LEAN_EXPORT lean_object* l_termExists___x2c__; static lean_object* l_term___xd7____1___closed__8; +static lean_object* l_Lean_doElemRepeat____Until_____closed__8; lean_object* l_Lean_TSyntax_getString(lean_object*); static lean_object* l_unexpandMkStr1___closed__3; LEAN_EXPORT lean_object* l_calcStep; @@ -101,35 +104,34 @@ static lean_object* l___aux__Init__NotationExtra______macroRules__Lean__Parser__ static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__13; static lean_object* l_Lean_bracketedExplicitBinders___closed__11; LEAN_EXPORT lean_object* l_unexpandMkStr1___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_doElemRepeat__Until_____closed__1; 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___aux__Init__NotationExtra______macroRules__term_u03a3___x2c____1___closed__2; LEAN_EXPORT lean_object* l_unexpandSigma(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__9; static lean_object* l_Lean_doElemWhile___x3a__Do_____closed__1; static lean_object* l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__solve__1___spec__3___closed__2; +static lean_object* l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__41; lean_object* l_Lean_Syntax_getId(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*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); static lean_object* l_termExists___x2c_____closed__2; static lean_object* l_unexpandMkArray5___closed__1; +static lean_object* l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__10; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____1___closed__6; LEAN_EXPORT lean_object* l_unexpandMkArray6(lean_object*, lean_object*, lean_object*); static lean_object* l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__1; static lean_object* l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__3; lean_object* lean_array_push(lean_object*, lean_object*); static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemWhile___x3a__Do____1___closed__2; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__18; lean_object* l_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_unexpandIte___closed__1; static lean_object* l_tacticFunext_________closed__7; static lean_object* l_Lean_Parser_Command_classAbbrev___closed__25; static lean_object* l_term___xd7____1___closed__5; static lean_object* l_tacticFunext_________closed__13; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__4; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__16; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__11; static lean_object* l_unexpandGetElem_x21___closed__4; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__37; +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_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__solve__1___spec__3___closed__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___boxed(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_unexpandMkArray1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_explicitBinders___closed__1; @@ -138,13 +140,15 @@ static lean_object* l_calcFirstStep___closed__12; lean_object* l_Lean_Syntax_getArgs(lean_object*); static lean_object* l_calcTactic___closed__5; LEAN_EXPORT lean_object* l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1(lean_object*, lean_object*, lean_object*); +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*); 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_termExists___x2c_____closed__8; +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___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__34; static lean_object* l_unexpandGetElem___closed__1; -static lean_object* l_Lean_doElemRepeat__Until_____closed__5; static lean_object* l___aux__Init__NotationExtra______macroRules__Lean__Parser__Command__classAbbrev__1___lambda__3___closed__3; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__22; -static lean_object* l_tacticFunext_________closed__15; +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*); lean_object* lean_array_fget(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_classAbbrev___closed__7; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__6; @@ -153,6 +157,7 @@ static lean_object* l_unexpandSigma___closed__1; static lean_object* l_term___xd7_x27_____closed__7; static lean_object* l_Lean_unifConstraint___closed__10; static lean_object* l_Lean_expandExplicitBindersAux_loop___closed__11; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__38; static lean_object* l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__5; LEAN_EXPORT lean_object* l_Lean_expandExplicitBindersAux___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Loop_forIn___rarg(lean_object*, lean_object*, lean_object*, lean_object*); @@ -165,13 +170,11 @@ static lean_object* l_tacticFunext_________closed__5; static lean_object* l_unexpandIte___closed__3; lean_object* l_Lean_Macro_throwError___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_term_u03a3___x2c_____closed__2; -static lean_object* l_Lean_doElemRepeat__Until_____closed__3; LEAN_EXPORT lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____1(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__24; static lean_object* l_Lean_Parser_Command_classAbbrev___closed__20; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__35; 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*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_unifConstraintElem; -static lean_object* l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__34; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____1___closed__17; static lean_object* l_unexpandGetElem_x3f___closed__3; static lean_object* l_unexpandGetElem_x3f___closed__2; @@ -188,12 +191,15 @@ static lean_object* l_solve___closed__9; static lean_object* l_unexpandListToArray___closed__1; LEAN_EXPORT lean_object* l___aux__Init__NotationExtra______macroRules__solve__1(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkIdentFrom(lean_object*, lean_object*, uint8_t); +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__33; static lean_object* l_Lean_bracketedExplicitBinders___closed__9; static lean_object* l_Lean_termMacro_x2etrace_x5b___x5d_____closed__22; +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__4; +static lean_object* l_Lean_bracketedExplicitBinders___closed__13; 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_tacticFunext_________closed__14; +static lean_object* l_Lean_doElemRepeat____Until_____closed__5; static lean_object* l_term_u03a3_x27___x2c_____closed__4; -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; lean_object* l_Lean_Syntax_node5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_cdotTk___closed__6; uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); @@ -204,15 +210,14 @@ static lean_object* l_Lean_Parser_Command_classAbbrev___closed__27; static lean_object* l_Lean_termMacro_x2etrace_x5b___x5d_____closed__5; static lean_object* l_unexpandIte___closed__5; LEAN_EXPORT lean_object* l_Lean_unifConstraint; -static lean_object* l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__7; static lean_object* l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4___closed__4; LEAN_EXPORT lean_object* l_Lean_instForInLoopUnit___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_unexpandListToArray___closed__2; static lean_object* l_cdotTk___closed__9; -static lean_object* l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__5; static lean_object* l_tacticFunext_________closed__3; static lean_object* l_Lean_unifConstraint___closed__16; static lean_object* l_Lean_doElemWhile___x3a__Do_____closed__2; +static lean_object* l_solve___closed__17; static lean_object* l_Lean_Parser_Command_classAbbrev___closed__2; 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_expandExplicitBindersAux_loop___closed__8; @@ -221,16 +226,19 @@ static lean_object* l___aux__Init__NotationExtra______macroRules__tacticFunext__ static lean_object* l_term___xd7_x27_____closed__2; static lean_object* l_Lean_term__Matches___x7c___closed__8; static lean_object* l_cdot___closed__2; -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_Lean_doElemWhile___x3a__Do_____closed__14; static lean_object* l_term___xd7____1___closed__4; static lean_object* l_term_u03a3___x2c_____closed__8; static lean_object* l_term_u03a3_x27___x2c_____closed__6; static lean_object* l_Lean_explicitBinders___closed__5; +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__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__6; static lean_object* l_calcSteps___closed__9; static lean_object* l___aux__Init__NotationExtra______macroRules__Lean__Parser__Command__classAbbrev__1___lambda__3___closed__5; +static lean_object* l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__33; static lean_object* l_calcSteps___closed__12; LEAN_EXPORT lean_object* l_Lean_instForInLoopUnit(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__termMacro_x2etrace_x5b___x5d____1___closed__1; static lean_object* l___aux__Init__NotationExtra______macroRules__Lean__Parser__Command__classAbbrev__1___lambda__3___closed__7; static lean_object* l_Lean_unifConstraint___closed__7; @@ -239,60 +247,59 @@ static lean_object* l_Lean_unifConstraint___closed__17; static lean_object* l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__10; static lean_object* l_term___xd7_x27_____closed__4; lean_object* l_Lean_Syntax_getNumArgs(lean_object*); -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_Lean_termMacro_x2etrace_x5b___x5d_____closed__7; -static lean_object* l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__25; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__9; static lean_object* l_term_u2203___x2c_____closed__2; LEAN_EXPORT lean_object* l_unexpandMkStr3(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__39; +LEAN_EXPORT lean_object* l_Lean_doElemRepeat____Until__; +static lean_object* l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__34; LEAN_EXPORT lean_object* l_unexpandMkStr5(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__31; 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*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___spec__7(lean_object*, size_t, size_t, lean_object*); +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__41; static lean_object* l_calcStep___closed__2; static lean_object* l_unexpandExists___closed__3; 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___lambda__2___closed__29; static lean_object* l_termExists___x2c_____closed__5; static lean_object* l_calcFirstStep___closed__1; static lean_object* l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__7; 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___lambda__2___closed__27; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____1___closed__1; static lean_object* l_tacticFunext_________closed__8; +static lean_object* l_Lean_term__Matches___x7c___closed__9; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__3; lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__1; static lean_object* l_cdotTk___closed__4; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__22; static lean_object* l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__6; LEAN_EXPORT lean_object* l_unexpandSubtype(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_doElemRepeat__Until_____closed__4; static lean_object* l_calcSteps___closed__2; -static lean_object* l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__28; static lean_object* l_calcFirstStep___closed__8; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemWhile___x3a__Do____1___closed__9; size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_expandExplicitBinders___spec__1(lean_object*, size_t, size_t); -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__32; +static lean_object* l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__42; LEAN_EXPORT lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemWhile___x3a__Do____1(lean_object*, lean_object*, lean_object*); static lean_object* l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__15; static lean_object* l___aux__Init__NotationExtra______macroRules__Lean__Parser__Command__classAbbrev__1___lambda__3___closed__1; -static lean_object* l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__35; static lean_object* l_term_u03a3___x2c_____closed__4; LEAN_EXPORT lean_object* l_Lean_bracketedExplicitBinders; static lean_object* l_Lean_bracketedExplicitBinders___closed__3; -LEAN_EXPORT lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2(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___lambda__2___closed__19; 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*); -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__1; static lean_object* l_Lean_unbracketedExplicitBinders___closed__6; -static lean_object* l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__13; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat__Until____1___closed__1; static lean_object* l_Lean_unbracketedExplicitBinders___closed__4; static lean_object* l_tacticFunext_________closed__1; +static lean_object* l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__8; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__17; lean_object* l_Array_reverse___rarg(lean_object*); static lean_object* l_Lean_bracketedExplicitBinders___closed__12; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__14; -static lean_object* l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__40; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemWhile___x3a__Do____1___closed__8; static lean_object* l_Lean_doElemRepeat_____closed__8; +static lean_object* l_Lean_Parser_Command_classAbbrev___closed__30; static lean_object* l_Lean_expandExplicitBindersAux_loop___closed__7; static lean_object* l_Lean_termMacro_x2etrace_x5b___x5d_____closed__11; static lean_object* l_calcFirstStep___closed__4; @@ -306,10 +313,8 @@ LEAN_EXPORT lean_object* l___aux__Init__NotationExtra______macroRules__Lean__Par static lean_object* l_Lean_doElemWhile___x3a__Do_____closed__11; static lean_object* l_tacticFunext_________closed__6; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____1___closed__14; -static lean_object* l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__17; 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_unexpandSubtype___closed__2; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__11; static lean_object* l_Lean_Parser_Command_classAbbrev___closed__23; static lean_object* l_Lean_termMacro_x2etrace_x5b___x5d_____closed__16; static lean_object* l_Lean_expandExplicitBindersAux_loop___closed__12; @@ -318,15 +323,16 @@ static lean_object* l_unexpandListCons___closed__1; static lean_object* l_termExists___x2c_____closed__1; 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__Lean__Parser__Command__classAbbrev__1___lambda__3___closed__6; -static lean_object* l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__3; lean_object* l_Array_zip___rarg(lean_object*, lean_object*); +static lean_object* l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__1; +static lean_object* l_solve___closed__16; static lean_object* l_Lean_expandExplicitBindersAux_loop___closed__3; static lean_object* l_Lean_unifConstraint___closed__5; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__38; lean_object* l_Lean_SourceInfo_fromRef(lean_object*, uint8_t); +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__31; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__2; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__15; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__20; +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*, size_t, size_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_term__Matches___x7c___closed__7; static lean_object* l_tacticFunext_________closed__9; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__28; @@ -347,11 +353,11 @@ LEAN_EXPORT lean_object* l_unexpandTSepArray(lean_object*, lean_object*, lean_ob static lean_object* l_Lean_Parser_Command_classAbbrev___closed__6; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__21; static lean_object* l___aux__Init__NotationExtra______macroRules__tacticFunext________1___closed__3; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__36; LEAN_EXPORT lean_object* l___aux__Init__NotationExtra______macroRules__term_u03a3_x27___x2c____1(lean_object*, lean_object*, lean_object*); static lean_object* l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__11; -static lean_object* l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__30; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____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*); +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__30; static lean_object* l_unexpandListNil___rarg___closed__1; static lean_object* l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__11; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__22; @@ -361,10 +367,10 @@ static lean_object* l___aux__Init__NotationExtra______macroRules__solve__1___clo static lean_object* l_term_u03a3_x27___x2c_____closed__1; static lean_object* l_Lean_expandExplicitBindersAux_loop___closed__9; lean_object* l_Lean_Syntax_getKind(lean_object*); +static lean_object* l_Lean_unbracketedExplicitBinders___closed__17; static lean_object* l_Lean_termMacro_x2etrace_x5b___x5d_____closed__9; static lean_object* l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__8; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____1___closed__16; -static lean_object* l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__19; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__26; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____1___closed__12; lean_object* l_Lean_mkCIdentFrom(lean_object*, lean_object*, uint8_t); @@ -372,7 +378,8 @@ static lean_object* l___aux__Init__NotationExtra______macroRules__Lean__Parser__ static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__20; static lean_object* l_Lean_Parser_Command_classAbbrev___closed__10; static lean_object* l_Lean_unifConstraintElem___closed__11; -static lean_object* l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__10; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____Until____1___closed__2; +static lean_object* l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__20; 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_unexpandListNil___rarg___closed__3; static lean_object* l_calcStep___closed__5; @@ -381,11 +388,12 @@ static lean_object* l_Lean_Parser_Command_classAbbrev___closed__5; static lean_object* l_term_u2203___x2c_____closed__1; LEAN_EXPORT lean_object* l_unexpandUnit___boxed(lean_object*); static lean_object* l_unexpandSubtype___closed__5; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__10; static lean_object* l_calcFirstStep___closed__10; LEAN_EXPORT lean_object* l_Lean_Loop_noConfusion___rarg(lean_object*); lean_object* l_Lean_quoteNameMk(lean_object*); LEAN_EXPORT lean_object* l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__36; +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*); LEAN_EXPORT lean_object* l_unexpandGetElem_x3f(lean_object*, lean_object*, lean_object*); static lean_object* l_calcFirstStep___closed__6; static lean_object* l_solve___closed__5; @@ -394,14 +402,18 @@ static lean_object* l_Lean_termMacro_x2etrace_x5b___x5d_____closed__19; static lean_object* l_Lean_bracketedExplicitBinders___closed__2; static lean_object* l___aux__Init__NotationExtra______macroRules__Lean__Parser__Command__classAbbrev__1___lambda__3___closed__16; static lean_object* l_termExists___x2c_____closed__6; +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__5; static lean_object* l_Lean_explicitBinders___closed__3; +static lean_object* l_Lean_term__Matches___x7c___closed__10; LEAN_EXPORT lean_object* l_unexpandMkStr4___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__33; +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_unexpandMkArray8___closed__1; lean_object* l_Lean_Syntax_getOptional_x3f(lean_object*); static lean_object* l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__12; +static lean_object* l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__21; static lean_object* l_Lean_doElemWhile___x3a__Do_____closed__8; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__7; static lean_object* l_Lean_termMacro_x2etrace_x5b___x5d_____closed__4; LEAN_EXPORT lean_object* l_calcTactic; static lean_object* l_calcSteps___closed__3; @@ -416,8 +428,7 @@ LEAN_EXPORT lean_object* l_Lean_termMacro_x2etrace_x5b___x5d__; static lean_object* l_Lean_termMacro_x2etrace_x5b___x5d_____closed__18; static lean_object* l_Lean_Parser_Command_classAbbrev___closed__22; LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___rarg___lambda__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___lambda__2___closed__5; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__39; +static lean_object* l_Lean_doElemRepeat____Until_____closed__4; static lean_object* l_cdot___closed__4; static lean_object* l_Lean_doElemWhile__Do_____closed__2; static lean_object* l_calcStep___closed__6; @@ -425,28 +436,35 @@ static lean_object* l_Lean_Parser_Command_classAbbrev___closed__18; LEAN_EXPORT lean_object* l_term_u03a3___x2c__; static lean_object* l_unexpandSubtype___closed__4; static lean_object* l_solve___closed__13; +static lean_object* l_Lean_explicitBinders___closed__8; static lean_object* l_Lean_unbracketedExplicitBinders___closed__2; -static lean_object* l_unexpandExists___closed__5; LEAN_EXPORT lean_object* l_Lean_expandBrackedBindersAux_loop(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_solve___closed__12; static lean_object* l___aux__Init__NotationExtra______macroRules__solve__1___closed__1; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__8; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__40; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemWhile___x3a__Do____1___closed__4; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_unexpandEqNDRec___closed__1; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__13; LEAN_EXPORT lean_object* l_unexpandMkArray0(lean_object*); static lean_object* l_Lean_termMacro_x2etrace_x5b___x5d_____closed__15; static lean_object* l_Lean_explicitBinders___closed__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*, 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__27; +static lean_object* l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__7; static lean_object* l_calcSteps___closed__10; +static lean_object* l_Lean_bracketedExplicitBinders___closed__15; +LEAN_EXPORT lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____Until____1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____Until____1___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_Lean_unbracketedExplicitBinders___closed__8; +static lean_object* l_Lean_doElemRepeat____Until_____closed__6; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__23; LEAN_EXPORT lean_object* l_unexpandUnit(lean_object*); +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__14; lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____1___closed__8; 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_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__2; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____1___closed__2; lean_object* l_Lean_extractMacroScopes(lean_object*); static lean_object* l_solve___closed__3; @@ -459,28 +477,30 @@ static lean_object* l___aux__Init__NotationExtra______macroRules__Lean__Parser__ static lean_object* l_Lean_expandExplicitBindersAux_loop___closed__6; LEAN_EXPORT lean_object* l___aux__Init__NotationExtra______macroRules__term_u03a3___x2c____1(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__17; static lean_object* l_calcSteps___closed__7; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__25; +static lean_object* l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__18; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__27; uint8_t lean_name_eq(lean_object*, lean_object*); static lean_object* l_term___xd7____1___closed__3; LEAN_EXPORT lean_object* l_unexpandMkStr2___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); static lean_object* l_calcFirstStep___closed__3; -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*, size_t, size_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_doElemWhile__Do_____closed__1; 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*); static lean_object* l_Lean_unifConstraintElem___closed__9; static lean_object* l_calcFirstStep___closed__9; static lean_object* l_term___xd7_x27_____closed__3; lean_object* l_Lean_Syntax_node2(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__23; LEAN_EXPORT lean_object* l_unexpandMkStr7(lean_object*, lean_object*, lean_object*); static lean_object* l___aux__Init__NotationExtra______macroRules__tacticFunext________1___closed__1; static lean_object* l_cdotTk___closed__13; static lean_object* l_Lean_explicitBinders___closed__7; +static lean_object* l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__27; static lean_object* l___aux__Init__NotationExtra______macroRules__solve__1___closed__3; lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); +static lean_object* l_Lean_unbracketedExplicitBinders___closed__14; LEAN_EXPORT lean_object* l_Lean_Loop_noConfusion___rarg___boxed(lean_object*); static lean_object* l___aux__Init__NotationExtra______macroRules__tacticFunext________1___closed__6; static lean_object* l_unexpandListNil___rarg___closed__2; @@ -489,21 +509,21 @@ static lean_object* l_Lean_unifConstraint___closed__9; static lean_object* l_Lean_Parser_Command_classAbbrev___closed__24; static lean_object* l_Lean_unifConstraint___closed__1; static lean_object* l_unexpandListToArray___closed__3; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__12; static lean_object* l_term_u2203___x2c_____closed__4; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__4; static lean_object* l_unexpandExists___closed__2; uint8_t l_Lean_Syntax_isIdent(lean_object*); static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__3; -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_Lean_unifConstraintElem___closed__5; static lean_object* l_Lean_doElemRepeat_____closed__6; LEAN_EXPORT lean_object* l_unexpandMkStr8___boxed(lean_object*, lean_object*, lean_object*); 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*); static lean_object* l_Lean_term__Matches___x7c___closed__3; static lean_object* l_unexpandGetElem_x21___closed__5; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__45; static lean_object* l___aux__Init__NotationExtra______macroRules__Lean__Parser__Command__classAbbrev__1___lambda__3___closed__2; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__40; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__10; +static lean_object* l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__23; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemWhile___x3a__Do____1___closed__6; static lean_object* l___aux__Init__NotationExtra______macroRules__solve__1___closed__2; LEAN_EXPORT lean_object* l_unexpandArrayEmpty___rarg(lean_object*, lean_object*); @@ -513,15 +533,13 @@ static lean_object* l_Lean_unifConstraint___closed__6; LEAN_EXPORT lean_object* l_Lean_term__Matches___x7c; static lean_object* l_unexpandSubtype___closed__1; static lean_object* l_Lean_termMacro_x2etrace_x5b___x5d_____closed__21; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__44; lean_object* l_Lean_MacroScopesView_review(lean_object*); static lean_object* l_solve___closed__11; 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*); -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__42; static lean_object* l_cdotTk___closed__5; static lean_object* l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__2; LEAN_EXPORT lean_object* l_calcSteps; -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_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__22; LEAN_EXPORT lean_object* l_unexpandTSyntaxArray(lean_object*, lean_object*, lean_object*); static lean_object* l_cdotTk___closed__7; 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*); @@ -531,7 +549,7 @@ static lean_object* l_calcSteps___closed__11; static lean_object* l_Lean_Parser_Command_classAbbrev___closed__1; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__15; static lean_object* l_unexpandSubtype___closed__3; -static lean_object* l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__20; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___spec__6(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l___aux__Init__NotationExtra______macroRules__Lean__Parser__Command__classAbbrev__1___lambda__2(lean_object*); static lean_object* l_Lean_bracketedExplicitBinders___closed__4; lean_object* l_Array_append___rarg(lean_object*, lean_object*); @@ -542,22 +560,28 @@ static lean_object* l_term_u03a3_x27___x2c_____closed__2; static lean_object* l_Lean_unifConstraint___closed__13; LEAN_EXPORT lean_object* l_Lean_Loop_noConfusion(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__7; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__9; static lean_object* l_Lean_term__Matches___x7c___closed__6; +static lean_object* l_Lean_doElemRepeat____Until_____closed__7; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__5; static lean_object* l_Lean_unifConstraint___closed__4; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__14; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__13; 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__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__17; lean_object* l_Array_extract___rarg(lean_object*, lean_object*, lean_object*); +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__4; lean_object* l_Lean_Syntax_node4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_doElemRepeat__Until_____closed__6; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__28; static lean_object* l_Lean_Parser_Command_classAbbrev___closed__9; LEAN_EXPORT lean_object* l_term_u2203___x2c__; static lean_object* l_Lean_doElemRepeat_____closed__5; static lean_object* l_Lean_bracketedExplicitBinders___closed__5; LEAN_EXPORT lean_object* l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_term_u2203___x2c_____closed__5; -static lean_object* l_calcSteps___closed__13; static lean_object* l_Lean_unifConstraintElem___closed__3; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__26; +static lean_object* l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__37; +static lean_object* l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__9; static lean_object* l_Lean_unifConstraint___closed__18; static lean_object* l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__13; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__13; @@ -568,13 +592,16 @@ static lean_object* l_Lean_bracketedExplicitBinders___closed__1; static lean_object* l_Lean_unifConstraintElem___closed__7; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__16; static lean_object* l_calc___closed__3; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__36; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__8; +static lean_object* l_Lean_bracketedExplicitBinders___closed__14; +static lean_object* l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__12; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__24; +static lean_object* l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__44; LEAN_EXPORT lean_object* l_Lean_expandExplicitBindersAux(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_unexpandTSyntax(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_unexpandMkStr1(lean_object*, lean_object*, 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*); uint8_t l_Lean_Name_hasMacroScopes(lean_object*); -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__17; static lean_object* l_term_u2203___x2c_____closed__7; static lean_object* l_Lean_Parser_Command_classAbbrev___closed__15; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__2; @@ -585,6 +612,7 @@ LEAN_EXPORT lean_object* l___aux__Init__NotationExtra______macroRules__Lean__Par LEAN_EXPORT lean_object* l___aux__Init__NotationExtra______macroRules__Lean__Parser__Command__classAbbrev__1___lambda__1___boxed(lean_object*, lean_object*); lean_object* l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(lean_object*, lean_object*); static lean_object* l_Lean_unifConstraint___closed__14; +static lean_object* l_Lean_doElemRepeat____Until_____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___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__5; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__18; @@ -592,10 +620,8 @@ static lean_object* l_unexpandGetElem_x21___closed__1; static lean_object* l_term___xd7_x27_____closed__1; static lean_object* l_Lean_expandExplicitBindersAux_loop___closed__4; static lean_object* l_calc___closed__1; -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_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__5; static lean_object* l_Lean_doElemWhile___x3a__Do_____closed__4; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__25; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__35; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__28; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_unexpandListCons(lean_object*, lean_object*, lean_object*); @@ -606,91 +632,88 @@ LEAN_EXPORT lean_object* l_unexpandEqRec(lean_object*, lean_object*, lean_object static lean_object* l_Lean_Parser_Command_classAbbrev___closed__19; static lean_object* l_cdot___closed__3; static lean_object* l_term_u03a3___x2c_____closed__1; -static lean_object* l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__39; static lean_object* l_solve___closed__1; static lean_object* l_cdot___closed__1; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_unexpandMkStr8(lean_object*, lean_object*, lean_object*); +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_doElemWhile___x3a__Do_____closed__6; LEAN_EXPORT lean_object* l_unexpandProdMk(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__12; static lean_object* l_Lean_Parser_Command_classAbbrev___closed__28; +static lean_object* l_Lean_bracketedExplicitBinders___closed__16; LEAN_EXPORT lean_object* l_term___xd7____1; static lean_object* l_calc___closed__4; static lean_object* l_unexpandUnit___rarg___closed__1; +static lean_object* l_Lean_doElemRepeat____Until_____closed__2; LEAN_EXPORT lean_object* l_Lean_expandExplicitBinders___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__4; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____1___closed__5; static lean_object* l_Lean_unifConstraint___closed__2; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); -static lean_object* l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__42; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__3; LEAN_EXPORT lean_object* l_Lean_expandBrackedBindersAux___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__7; 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___aux__Init__NotationExtra______macroRules__Lean__doElemWhile___x3a__Do____1___closed__3; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__14; static lean_object* l_unexpandGetElem_x21___closed__2; static lean_object* l_Lean_termMacro_x2etrace_x5b___x5d_____closed__13; +LEAN_EXPORT lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_classAbbrev___closed__12; -static lean_object* l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__16; static lean_object* l___aux__Init__NotationExtra______macroRules__Lean__Parser__Command__classAbbrev__1___lambda__3___closed__10; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___spec__7(lean_object*, size_t, size_t, lean_object*); lean_object* l_Lean_Syntax_node7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_doElemRepeat__Until_____closed__2; static lean_object* l_Lean_Parser_Command_classAbbrev___closed__8; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__27; static lean_object* l_Lean_termMacro_x2etrace_x5b___x5d_____closed__14; lean_object* l_Lean_Syntax_node1(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__11; lean_object* l_Lean_mkSepArray(lean_object*, lean_object*); +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__25; LEAN_EXPORT lean_object* l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_unexpandMkArray0___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_instForInLoopUnit___rarg___boxed(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___lambda__2___closed__43; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__30; -static lean_object* l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__22; extern lean_object* l_Lean_binderIdent; +static lean_object* l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__16; static lean_object* l_Lean_doElemWhile__Do_____closed__3; 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__command____Unif__hint________Where___x7c_x2d_u22a2____1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_expandExplicitBindersAux_loop___closed__2; static lean_object* l_Lean_explicitBinders___closed__2; static lean_object* l_calcSteps___closed__6; -static lean_object* l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__15; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__3; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__32; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__21; static lean_object* l_unexpandSorryAx___closed__2; uint8_t l_Lean_Syntax_isNone(lean_object*); static lean_object* l_solve___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___lambda__2___closed__41; LEAN_EXPORT lean_object* l_unexpandMkArray0___rarg(lean_object*, lean_object*); +static lean_object* l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__6; static lean_object* l_Lean_unifConstraintElem___closed__4; static lean_object* l_unexpandGetElem_x21___closed__3; static lean_object* l_Lean_doElemWhile__Do_____closed__6; static lean_object* l_term_u03a3_x27___x2c_____closed__3; static lean_object* l_Lean_expandExplicitBindersAux_loop___closed__5; +static lean_object* l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__15; static lean_object* l_term_u03a3___x2c_____closed__7; lean_object* lean_nat_sub(lean_object*, lean_object*); +static lean_object* l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__25; static lean_object* l_Lean_termMacro_x2etrace_x5b___x5d_____closed__3; static lean_object* l_cdotTk___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__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_solve; LEAN_EXPORT lean_object* l_Array_sequenceMap___at___aux__Init__NotationExtra______macroRules__solve__1___spec__1___boxed(lean_object*, lean_object*); lean_object* l_Array_mkArray1___rarg(lean_object*); -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat__Until____1___closed__2; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_expandExplicitBinders___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_doElemWhile___x3a__Do_____closed__12; -static lean_object* l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__2; LEAN_EXPORT lean_object* l_Lean_unbracketedExplicitBinders; static lean_object* l_Lean_doElemWhile__Do_____closed__5; -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*); +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__43; static lean_object* l_Lean_doElemRepeat_____closed__7; LEAN_EXPORT lean_object* l_Lean_Loop_toCtorIdx(lean_object*); -static lean_object* l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__14; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____1___closed__10; static lean_object* l_calcTactic___closed__1; LEAN_EXPORT lean_object* l_Lean_doElemRepeat__; -static lean_object* l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__41; +static lean_object* l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__11; LEAN_EXPORT lean_object* l_Lean_Loop_forIn___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_cdotTk___closed__11; LEAN_EXPORT lean_object* l_Array_sequenceMap___at___aux__Init__NotationExtra______macroRules__solve__1___spec__1(lean_object*, lean_object*); @@ -699,101 +722,98 @@ static lean_object* l_Lean_unbracketedExplicitBinders___closed__10; static lean_object* l_cdot___closed__5; static lean_object* l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__9; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__8; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__23; static lean_object* l_Lean_Parser_Command_classAbbrev___closed__26; static lean_object* l_Lean_unifConstraintElem___closed__10; lean_object* l_Array_ofSubarray___rarg(lean_object*); -static lean_object* l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__4; -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_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__24; static lean_object* l_term___xd7_x27_____closed__5; -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_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_tacticFunext_________closed__10; static lean_object* l_calcStep___closed__4; lean_object* lean_erase_macro_scopes(lean_object*); static lean_object* l_Lean_Parser_Command_classAbbrev___closed__29; -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_Lean_doElemRepeat____Until_____closed__1; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__24; static lean_object* l_calcSteps___closed__4; LEAN_EXPORT lean_object* l_unexpandListNil___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_unexpandGetElem(lean_object*, lean_object*, lean_object*); static lean_object* l_solve___closed__14; +static lean_object* l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__35; 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__36; lean_object* l_String_intercalate(lean_object*, lean_object*); size_t lean_usize_sub(size_t, size_t); +static lean_object* l_Lean_unbracketedExplicitBinders___closed__15; static lean_object* l_calcStep___closed__1; static lean_object* l_calcSteps___closed__8; static lean_object* l___aux__Init__NotationExtra______macroRules__tacticFunext________1___closed__9; static lean_object* l___aux__Init__NotationExtra______macroRules__term_u03a3___x2c____1___closed__1; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__21; LEAN_EXPORT lean_object* l_Lean_doElemWhile___x3a__Do__; +static lean_object* l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__28; 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__29; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__19; static lean_object* l_calcFirstStep___closed__5; static lean_object* l_tacticFunext_________closed__2; +static lean_object* l_Lean_unbracketedExplicitBinders___closed__16; size_t lean_usize_add(size_t, size_t); LEAN_EXPORT lean_object* l_unexpandSorryAx(lean_object*, lean_object*, lean_object*); -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*); LEAN_EXPORT lean_object* l_unexpandListNil___boxed(lean_object*); static lean_object* l_cdotTk___closed__1; static lean_object* l_Lean_bracketedExplicitBinders___closed__6; static lean_object* l___aux__Init__NotationExtra______macroRules__solve__1___closed__5; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_doElemWhile___x3a__Do_____closed__5; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__42; 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___lambda__2___closed__29; lean_object* lean_array_uget(lean_object*, size_t); static lean_object* l_Lean_Parser_Command_classAbbrev___closed__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__4(size_t, size_t, lean_object*); -static lean_object* l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__12; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__34; LEAN_EXPORT lean_object* l_unexpandListNil(lean_object*); static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__26; static lean_object* l_Lean_Parser_Command_classAbbrev___closed__4; 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__2; -static lean_object* l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__43; static lean_object* l_calcTactic___closed__4; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__9; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemWhile___x3a__Do____1___closed__1; static lean_object* l_calcTactic___closed__2; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__6; extern lean_object* l_Lean_instInhabitedSyntax; LEAN_EXPORT lean_object* l_unexpandArrayEmpty___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_expandExplicitBindersAux_loop(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___aux__Init__NotationExtra______macroRules__termExists___x2c____1(lean_object*, lean_object*, lean_object*); 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__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__45; +static lean_object* l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__14; lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_term__Matches___x7c___closed__1; -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__4; static lean_object* l_Lean_doElemWhile___x3a__Do_____closed__9; static lean_object* l___aux__Init__NotationExtra______macroRules__Lean__Parser__Command__classAbbrev__1___lambda__3___closed__8; static lean_object* l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__9; LEAN_EXPORT lean_object* l_Lean_doElemWhile__Do__; -static lean_object* l_tacticFunext_________closed__16; static lean_object* l_Lean_unifConstraint___closed__11; LEAN_EXPORT lean_object* l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_termMacro_x2etrace_x5b___x5d_____closed__6; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__12; lean_object* lean_string_append(lean_object*, lean_object*); +static lean_object* l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__40; static lean_object* l_term___xd7____1___closed__1; static lean_object* l___aux__Init__NotationExtra______macroRules__tacticFunext________1___closed__8; static lean_object* l_Lean_doElemRepeat_____closed__4; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__5; +static lean_object* l_solve___closed__15; LEAN_EXPORT lean_object* l_unexpandPSigma(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__38; static lean_object* l_Lean_unbracketedExplicitBinders___closed__9; static lean_object* l_Lean_doElemRepeat_____closed__9; 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___aux__Init__NotationExtra______macroRules__tacticFunext________1___closed__4; lean_object* lean_array_get_size(lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___spec__6(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l___aux__Init__NotationExtra______macroRules__Lean__Parser__Command__classAbbrev__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_unexpandMkStr6(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_doElemWhile___x3a__Do_____closed__7; 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_unexpandIte___closed__4; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__12; static lean_object* l_Lean_termMacro_x2etrace_x5b___x5d_____closed__2; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__7; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__15; uint8_t lean_nat_dec_le(lean_object*, lean_object*); static lean_object* l_Lean_termMacro_x2etrace_x5b___x5d_____closed__17; static lean_object* l_unexpandExists___closed__4; @@ -803,29 +823,27 @@ static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__ter LEAN_EXPORT lean_object* l_Lean_expandBrackedBinders(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_calcStep___closed__3; lean_object* lean_nat_add(lean_object*, lean_object*); +static lean_object* l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__30; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__17; static lean_object* l_Lean_Parser_Command_classAbbrev___closed__21; static lean_object* l_Lean_unifConstraint___closed__12; static lean_object* l___aux__Init__NotationExtra______macroRules__term_u03a3_x27___x2c____1___closed__2; static lean_object* l_Lean_doElemWhile___x3a__Do_____closed__15; +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__39; static lean_object* l_Lean_bracketedExplicitBinders___closed__10; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____1___closed__3; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__1; 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___lambda__2___closed__31; static lean_object* l_calcFirstStep___closed__7; LEAN_EXPORT lean_object* l_tacticFunext______; static lean_object* l___aux__Init__NotationExtra______macroRules__Lean__Parser__Command__classAbbrev__1___lambda__3___closed__13; -LEAN_EXPORT lean_object* l_Lean_doElemRepeat__Until__; static lean_object* l_Lean_term__Matches___x7c___closed__5; LEAN_EXPORT lean_object* l_unexpandMkArray4(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__26; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____1___closed__13; +static lean_object* l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__19; LEAN_EXPORT lean_object* l_unexpandExists(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__6; static lean_object* l_tacticFunext_________closed__11; lean_object* l_String_toSubstring_x27(lean_object*); -static lean_object* l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__1; static lean_object* l_Lean_unifConstraintElem___closed__8; LEAN_EXPORT lean_object* l_unexpandMkArray3(lean_object*, lean_object*, lean_object*); static lean_object* l_solve___closed__8; @@ -838,29 +856,25 @@ static lean_object* l_unexpandGetElem_x3f___closed__1; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemWhile___x3a__Do____1___closed__10; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__20; LEAN_EXPORT lean_object* l_unexpandMkArray8(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__38; static lean_object* l___aux__Init__NotationExtra______macroRules__term_u03a3_x27___x2c____1___closed__1; -static lean_object* l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__9; static lean_object* l_solve___closed__4; static lean_object* l___aux__Init__NotationExtra______macroRules__tacticFunext________1___closed__7; static lean_object* l_Lean_doElemWhile__Do_____closed__4; static lean_object* l_Lean_unifConstraint___closed__8; -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*); static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__24; +static lean_object* l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__31; static lean_object* l_term___xd7____1___closed__7; static lean_object* l_Lean_Parser_Command_classAbbrev___closed__11; -static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__2; static lean_object* l_unexpandPSigma___closed__1; -static lean_object* l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__37; static lean_object* l_Lean_Parser_Command_classAbbrev___closed__14; static lean_object* l_unexpandMkStr1___closed__2; static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____1___closed__4; LEAN_EXPORT lean_object* l___aux__Init__NotationExtra______macroRules__tacticFunext________1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___aux__Init__NotationExtra______macroRules__Lean__Parser__Command__classAbbrev__1___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_unexpandMkStr3___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__29; LEAN_EXPORT lean_object* l_unexpandMkStr5___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__10; 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(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_calcFirstStep; @@ -1517,16 +1531,58 @@ return x_3; static lean_object* _init_l_Lean_unbracketedExplicitBinders___closed__5() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("ppSpace", 7); +return x_1; +} +} +static lean_object* _init_l_Lean_unbracketedExplicitBinders___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_unbracketedExplicitBinders___closed__5; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_unbracketedExplicitBinders___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_unbracketedExplicitBinders___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_unbracketedExplicitBinders___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__5; +x_2 = l_Lean_unbracketedExplicitBinders___closed__7; +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); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_unbracketedExplicitBinders___closed__9() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_unbracketedExplicitBinders___closed__4; -x_2 = l_Lean_binderIdent; +x_2 = l_Lean_unbracketedExplicitBinders___closed__8; 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_unbracketedExplicitBinders___closed__6() { +static lean_object* _init_l_Lean_unbracketedExplicitBinders___closed__10() { _start: { lean_object* x_1; @@ -1534,17 +1590,17 @@ x_1 = lean_mk_string_from_bytes("optional", 8); return x_1; } } -static lean_object* _init_l_Lean_unbracketedExplicitBinders___closed__7() { +static lean_object* _init_l_Lean_unbracketedExplicitBinders___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_unbracketedExplicitBinders___closed__6; +x_2 = l_Lean_unbracketedExplicitBinders___closed__10; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_unbracketedExplicitBinders___closed__8() { +static lean_object* _init_l_Lean_unbracketedExplicitBinders___closed__12() { _start: { lean_object* x_1; @@ -1552,22 +1608,22 @@ x_1 = lean_mk_string_from_bytes(" : ", 3); return x_1; } } -static lean_object* _init_l_Lean_unbracketedExplicitBinders___closed__9() { +static lean_object* _init_l_Lean_unbracketedExplicitBinders___closed__13() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_unbracketedExplicitBinders___closed__8; +x_1 = l_Lean_unbracketedExplicitBinders___closed__12; 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_unbracketedExplicitBinders___closed__10() { +static lean_object* _init_l_Lean_unbracketedExplicitBinders___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__5; -x_2 = l_Lean_unbracketedExplicitBinders___closed__9; +x_2 = l_Lean_unbracketedExplicitBinders___closed__13; x_3 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__19; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -1576,25 +1632,25 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_unbracketedExplicitBinders___closed__11() { +static lean_object* _init_l_Lean_unbracketedExplicitBinders___closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_unbracketedExplicitBinders___closed__7; -x_2 = l_Lean_unbracketedExplicitBinders___closed__10; +x_1 = l_Lean_unbracketedExplicitBinders___closed__11; +x_2 = l_Lean_unbracketedExplicitBinders___closed__14; 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_unbracketedExplicitBinders___closed__12() { +static lean_object* _init_l_Lean_unbracketedExplicitBinders___closed__16() { _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__5; -x_2 = l_Lean_unbracketedExplicitBinders___closed__5; -x_3 = l_Lean_unbracketedExplicitBinders___closed__11; +x_2 = l_Lean_unbracketedExplicitBinders___closed__9; +x_3 = l_Lean_unbracketedExplicitBinders___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); @@ -1602,13 +1658,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_unbracketedExplicitBinders___closed__13() { +static lean_object* _init_l_Lean_unbracketedExplicitBinders___closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_unbracketedExplicitBinders___closed__1; x_2 = l_Lean_unbracketedExplicitBinders___closed__2; -x_3 = l_Lean_unbracketedExplicitBinders___closed__12; +x_3 = l_Lean_unbracketedExplicitBinders___closed__16; x_4 = lean_alloc_ctor(9, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -1620,7 +1676,7 @@ static lean_object* _init_l_Lean_unbracketedExplicitBinders() { _start: { lean_object* x_1; -x_1 = l_Lean_unbracketedExplicitBinders___closed__13; +x_1 = l_Lean_unbracketedExplicitBinders___closed__17; return x_1; } } @@ -1675,8 +1731,8 @@ static lean_object* _init_l_Lean_bracketedExplicitBinders___closed__6() { { 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__5; -x_2 = l_Lean_unbracketedExplicitBinders___closed__5; -x_3 = l_Lean_unbracketedExplicitBinders___closed__9; +x_2 = l_Lean_binderIdent; +x_3 = l_Lean_unbracketedExplicitBinders___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); @@ -1687,9 +1743,53 @@ return x_4; static lean_object* _init_l_Lean_bracketedExplicitBinders___closed__7() { _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___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_Lean_bracketedExplicitBinders___closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(": ", 2); +return x_1; +} +} +static lean_object* _init_l_Lean_bracketedExplicitBinders___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_bracketedExplicitBinders___closed__8; +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_bracketedExplicitBinders___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__5; -x_2 = l_Lean_bracketedExplicitBinders___closed__6; +x_2 = l_Lean_bracketedExplicitBinders___closed__7; +x_3 = l_Lean_bracketedExplicitBinders___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); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_bracketedExplicitBinders___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__5; +x_2 = l_Lean_bracketedExplicitBinders___closed__10; x_3 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__19; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -1698,25 +1798,25 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_bracketedExplicitBinders___closed__8() { +static lean_object* _init_l_Lean_bracketedExplicitBinders___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_bracketedExplicitBinders___closed__5; -x_2 = l_Lean_bracketedExplicitBinders___closed__7; +x_2 = l_Lean_bracketedExplicitBinders___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_bracketedExplicitBinders___closed__9() { +static lean_object* _init_l_Lean_bracketedExplicitBinders___closed__13() { _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__5; x_2 = l_Lean_bracketedExplicitBinders___closed__3; -x_3 = l_Lean_bracketedExplicitBinders___closed__8; +x_3 = l_Lean_bracketedExplicitBinders___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); @@ -1724,7 +1824,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_bracketedExplicitBinders___closed__10() { +static lean_object* _init_l_Lean_bracketedExplicitBinders___closed__14() { _start: { lean_object* x_1; lean_object* x_2; @@ -1734,13 +1834,13 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_bracketedExplicitBinders___closed__11() { +static lean_object* _init_l_Lean_bracketedExplicitBinders___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__5; -x_2 = l_Lean_bracketedExplicitBinders___closed__9; -x_3 = l_Lean_bracketedExplicitBinders___closed__10; +x_2 = l_Lean_bracketedExplicitBinders___closed__13; +x_3 = l_Lean_bracketedExplicitBinders___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); @@ -1748,13 +1848,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_bracketedExplicitBinders___closed__12() { +static lean_object* _init_l_Lean_bracketedExplicitBinders___closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_bracketedExplicitBinders___closed__1; x_2 = l_Lean_bracketedExplicitBinders___closed__2; -x_3 = l_Lean_bracketedExplicitBinders___closed__11; +x_3 = l_Lean_bracketedExplicitBinders___closed__15; x_4 = lean_alloc_ctor(9, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -1766,7 +1866,7 @@ static lean_object* _init_l_Lean_bracketedExplicitBinders() { _start: { lean_object* x_1; -x_1 = l_Lean_bracketedExplicitBinders___closed__12; +x_1 = l_Lean_bracketedExplicitBinders___closed__16; return x_1; } } @@ -1809,21 +1909,35 @@ 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; lean_object* x_4; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__5; +x_2 = l_Lean_unbracketedExplicitBinders___closed__7; +x_3 = l_Lean_bracketedExplicitBinders; +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_explicitBinders___closed__6() { +_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; +x_2 = l_Lean_explicitBinders___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_explicitBinders___closed__6() { +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__4; -x_2 = l_Lean_explicitBinders___closed__5; +x_2 = l_Lean_explicitBinders___closed__6; x_3 = l_Lean_unbracketedExplicitBinders; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -1832,13 +1946,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_explicitBinders___closed__7() { +static lean_object* _init_l_Lean_explicitBinders___closed__8() { _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__6; +x_3 = l_Lean_explicitBinders___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); @@ -1850,7 +1964,7 @@ static lean_object* _init_l_Lean_explicitBinders() { _start: { lean_object* x_1; -x_1 = l_Lean_explicitBinders___closed__7; +x_1 = l_Lean_explicitBinders___closed__8; return x_1; } } @@ -2825,7 +2939,7 @@ static lean_object* _init_l_Lean_unifConstraintElem___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_unbracketedExplicitBinders___closed__7; +x_1 = l_Lean_unbracketedExplicitBinders___closed__11; x_2 = l_Lean_unifConstraintElem___closed__8; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -2869,25 +2983,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_|-⊢_", 34); 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__1; -x_2 = l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__1; +x_2 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__1; x_3 = l_Lean_Name_mkStr2(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; @@ -2895,39 +3009,39 @@ x_1 = lean_mk_string_from_bytes("docComment", 10); 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 = l_Lean_Name_str___override(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; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_unbracketedExplicitBinders___closed__7; -x_2 = l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__5; +x_1 = l_Lean_unbracketedExplicitBinders___closed__11; +x_2 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____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_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; @@ -2935,33 +3049,33 @@ 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__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; x_1 = lean_box(0); -x_2 = l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__7; +x_2 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__7; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -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; -x_1 = l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__8; +x_1 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____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_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__5; -x_2 = l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__6; -x_3 = l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__9; +x_2 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__6; +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); @@ -2969,31 +3083,45 @@ 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; -x_1 = lean_mk_string_from_bytes("unif_hint ", 10); +x_1 = lean_mk_string_from_bytes("unif_hint", 9); 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; -x_1 = l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__11; +x_1 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____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_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; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__5; +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__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); +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__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__5; -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__12; +x_2 = l_Lean_unbracketedExplicitBinders___closed__7; +x_3 = l_Lean_termMacro_x2etrace_x5b___x5d_____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); @@ -3001,25 +3129,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__14() { +static lean_object* _init_l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__15() { _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__10; +x_1 = l_Lean_unbracketedExplicitBinders___closed__11; +x_2 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__14; 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__15() { +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; lean_object* x_4; x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__5; -x_2 = l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__13; -x_3 = l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__14; +x_2 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__13; +x_3 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____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); @@ -3027,7 +3155,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__16() { +static lean_object* _init_l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__17() { _start: { lean_object* x_1; @@ -3035,17 +3163,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__17() { +static lean_object* _init_l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__18() { _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__16; +x_2 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__17; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -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__19() { _start: { lean_object* x_1; @@ -3053,45 +3181,59 @@ 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__19() { +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; x_1 = lean_box(0); -x_2 = l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__18; +x_2 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__19; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -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__21() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__19; +x_1 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____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_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__21() { +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; lean_object* x_4; +x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__5; +x_2 = l_Lean_unbracketedExplicitBinders___closed__7; +x_3 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____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); +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__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__17; -x_2 = l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__20; +x_1 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__18; +x_2 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____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_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__22() { +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; lean_object* x_4; x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__5; -x_2 = l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__15; -x_3 = l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__21; +x_2 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__16; +x_3 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____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); @@ -3099,7 +3241,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__23() { +static lean_object* _init_l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__25() { _start: { lean_object* x_1; @@ -3107,23 +3249,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__24() { +static lean_object* _init_l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__26() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__23; +x_1 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__25; 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__25() { +static lean_object* _init_l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__27() { _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__5; -x_2 = l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__22; -x_3 = l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__24; +x_2 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__24; +x_3 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__26; 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,7 +3273,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__28() { _start: { lean_object* x_1; @@ -3139,21 +3281,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__27() { +static lean_object* _init_l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__29() { _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__26; +x_2 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__28; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -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__30() { _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__17; +x_1 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__18; x_2 = l_Lean_unifConstraintElem; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -3161,25 +3303,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__29() { +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; -x_1 = l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__27; -x_2 = l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__28; +x_1 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__29; +x_2 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__30; 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__30() { +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__5; -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__29; +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__31; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -3187,7 +3329,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__31() { +static lean_object* _init_l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__33() { _start: { lean_object* x_1; @@ -3195,33 +3337,33 @@ 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__32() { +static lean_object* _init_l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__34() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_unifConstraint___closed__5; -x_2 = l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__31; +x_2 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__33; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -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__35() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__31; +x_1 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__33; 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__34() { +static lean_object* _init_l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__36() { _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__31; -x_2 = l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__32; -x_3 = l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__33; +x_1 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__33; +x_2 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__34; +x_3 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__35; x_4 = lean_alloc_ctor(9, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -3229,7 +3371,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__35() { +static lean_object* _init_l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__37() { _start: { lean_object* x_1; @@ -3237,33 +3379,33 @@ 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__36() { +static lean_object* _init_l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__38() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_unifConstraint___closed__5; -x_2 = l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__35; +x_2 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__37; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__37() { +static lean_object* _init_l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__39() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__35; +x_1 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__37; 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__38() { +static lean_object* _init_l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__40() { _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__35; -x_2 = l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__36; -x_3 = l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__37; +x_1 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__37; +x_2 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__38; +x_3 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__39; x_4 = lean_alloc_ctor(9, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -3271,13 +3413,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__39() { +static lean_object* _init_l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__41() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_explicitBinders___closed__4; -x_2 = l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__34; -x_3 = l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__38; +x_2 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__36; +x_3 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__40; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -3285,25 +3427,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__40() { +static lean_object* _init_l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__42() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_unifConstraint___closed__4; -x_2 = l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__39; +x_2 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__41; 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__41() { +static lean_object* _init_l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__43() { _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__5; -x_2 = l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__30; -x_3 = l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__40; +x_2 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__32; +x_3 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__42; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -3311,12 +3453,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__42() { +static lean_object* _init_l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__44() { _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__5; -x_2 = l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__41; +x_2 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__43; x_3 = l_Lean_unifConstraint; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -3325,13 +3467,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__43() { +static lean_object* _init_l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__45() { _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__42; +x_3 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__44; x_4 = lean_alloc_ctor(3, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -3339,15 +3481,15 @@ 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__43; +x_1 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__45; return x_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) { +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_6; uint8_t x_7; @@ -3416,18 +3558,18 @@ return x_20; } } } -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) { +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: { 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); +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) { +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: { uint8_t x_4; @@ -3454,7 +3596,7 @@ 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) { +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; @@ -3481,7 +3623,7 @@ 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() { +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; @@ -3489,19 +3631,19 @@ 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() { +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; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__1; x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__1; x_3 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__2; -x_4 = l_Array_forInUnsafe_loop___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___spec__5___closed__1; +x_4 = l_Array_forInUnsafe_loop___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___spec__5___closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -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() { +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; @@ -3509,7 +3651,7 @@ 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__4() { +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__4() { _start: { lean_object* x_1; @@ -3517,7 +3659,7 @@ 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, 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___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, 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; @@ -3545,7 +3687,7 @@ x_14 = lean_ctor_get(x_7, 5); lean_inc(x_14); x_15 = 0; x_16 = l_Lean_SourceInfo_fromRef(x_14, x_15); -x_17 = l_Array_forInUnsafe_loop___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___spec__5___closed__3; +x_17 = l_Array_forInUnsafe_loop___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___spec__5___closed__3; lean_inc(x_16); x_18 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_18, 0, x_16); @@ -3553,12 +3695,12 @@ lean_ctor_set(x_18, 1, x_17); lean_inc(x_2); lean_inc(x_16); x_19 = l_Lean_Syntax_node3(x_16, x_2, x_12, x_18, x_13); -x_20 = l_Array_forInUnsafe_loop___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___spec__5___closed__4; +x_20 = l_Array_forInUnsafe_loop___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___spec__5___closed__4; 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_Array_forInUnsafe_loop___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___spec__5___closed__2; +x_22 = l_Array_forInUnsafe_loop___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___spec__5___closed__2; x_23 = l_Lean_Syntax_node3(x_16, x_22, x_19, x_21, x_6); x_24 = 1; x_25 = lean_usize_add(x_5, x_24); @@ -3568,7 +3710,7 @@ 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__6(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_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___spec__6(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { _start: { uint8_t x_5; @@ -3592,7 +3734,7 @@ 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__7(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_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___spec__7(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { _start: { uint8_t x_5; @@ -3616,7 +3758,7 @@ goto _start; } } } -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_EXPORT lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__1(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; @@ -3679,27 +3821,27 @@ return x_18; } } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__1() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___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_termMacro_x2etrace_x5b___x5d_____closed__1; x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__1; x_3 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__2; -x_4 = l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__7; +x_4 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__7; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__2() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__2() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__1), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__1), 1, 0); return x_1; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__3() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -3711,7 +3853,7 @@ 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___lambda__2___closed__4() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__4() { _start: { lean_object* x_1; @@ -3719,17 +3861,17 @@ x_1 = lean_mk_string_from_bytes("term_=_", 7); return x_1; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__5() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__5() { _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___lambda__2___closed__4; +x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__4; x_3 = l_Lean_Name_str___override(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___lambda__2___closed__6() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__6() { _start: { lean_object* x_1; @@ -3737,26 +3879,26 @@ 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___lambda__2___closed__7() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__7() { _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___lambda__2___closed__6; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__6; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__8() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___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___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__6; +x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__6; x_3 = l_Lean_Name_str___override(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___lambda__2___closed__9() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__9() { _start: { lean_object* x_1; @@ -3764,7 +3906,7 @@ 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___lambda__2___closed__10() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__10() { _start: { lean_object* x_1; @@ -3772,19 +3914,19 @@ 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___lambda__2___closed__11() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__11() { _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_termMacro_x2etrace_x5b___x5d_____closed__1; x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__1; -x_3 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__9; -x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__10; +x_3 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__9; +x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__10; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__12() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__12() { _start: { lean_object* x_1; @@ -3792,19 +3934,19 @@ 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___lambda__2___closed__13() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__13() { _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_termMacro_x2etrace_x5b___x5d_____closed__1; x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__1; -x_3 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__9; -x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__12; +x_3 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__9; +x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__12; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__14() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__14() { _start: { lean_object* x_1; @@ -3812,19 +3954,19 @@ 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___lambda__2___closed__15() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__15() { _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_termMacro_x2etrace_x5b___x5d_____closed__1; x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__1; x_3 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__2; -x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__14; +x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__14; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__16() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__16() { _start: { lean_object* x_1; @@ -3832,7 +3974,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___lambda__2___closed__17() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__17() { _start: { lean_object* x_1; @@ -3840,19 +3982,19 @@ 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___lambda__2___closed__18() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__18() { _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_termMacro_x2etrace_x5b___x5d_____closed__1; x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__1; x_3 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__2; -x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__17; +x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__17; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__19() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__19() { _start: { lean_object* x_1; @@ -3860,7 +4002,7 @@ 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___lambda__2___closed__20() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__20() { _start: { lean_object* x_1; @@ -3868,19 +4010,19 @@ 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___lambda__2___closed__21() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__21() { _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_termMacro_x2etrace_x5b___x5d_____closed__1; x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__1; -x_3 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__19; -x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__20; +x_3 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__19; +x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__20; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__22() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__22() { _start: { lean_object* x_1; @@ -3888,26 +4030,26 @@ x_1 = lean_mk_string_from_bytes("unification_hint", 16); return x_1; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__23() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__23() { _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___lambda__2___closed__22; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__22; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__24() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____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___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__22; +x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__22; x_3 = l_Lean_Name_str___override(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___lambda__2___closed__25() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__25() { _start: { lean_object* x_1; @@ -3915,19 +4057,19 @@ 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___lambda__2___closed__26() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__26() { _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_termMacro_x2etrace_x5b___x5d_____closed__1; x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__1; -x_3 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__9; -x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__25; +x_3 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__9; +x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__25; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__27() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__27() { _start: { lean_object* x_1; @@ -3935,19 +4077,19 @@ 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___lambda__2___closed__28() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__28() { _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_termMacro_x2etrace_x5b___x5d_____closed__1; x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__1; -x_3 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__9; -x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__27; +x_3 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__9; +x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__27; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__29() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__29() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; @@ -3961,7 +4103,7 @@ 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___lambda__2___closed__30() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__30() { _start: { lean_object* x_1; lean_object* x_2; @@ -3970,7 +4112,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__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__31() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__31() { _start: { lean_object* x_1; @@ -3978,67 +4120,67 @@ 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___lambda__2___closed__32() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__32() { _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_termMacro_x2etrace_x5b___x5d_____closed__1; x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__1; -x_3 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__9; -x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__31; +x_3 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__9; +x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__31; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__33() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__33() { _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_termMacro_x2etrace_x5b___x5d_____closed__1; x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__1; x_3 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__2; -x_4 = l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__18; +x_4 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__19; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__34() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__34() { _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___lambda__2___closed__33; +x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___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___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__35() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__35() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_expandExplicitBindersAux_loop___closed__6; -x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__34; +x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__34; 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___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__36() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__36() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__9; -x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__35; +x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__35; 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___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__37() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__37() { _start: { lean_object* x_1; @@ -4046,19 +4188,19 @@ 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___lambda__2___closed__38() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__38() { _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_termMacro_x2etrace_x5b___x5d_____closed__1; x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__1; x_3 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__2; -x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__37; +x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__37; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__39() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__39() { _start: { lean_object* x_1; @@ -4066,7 +4208,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___lambda__2___closed__40() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__40() { _start: { lean_object* x_1; @@ -4074,19 +4216,19 @@ 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___lambda__2___closed__41() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__41() { _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_termMacro_x2etrace_x5b___x5d_____closed__1; x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__1; -x_3 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__40; +x_3 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__40; x_4 = l_Lean_expandExplicitBindersAux_loop___closed__5; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__42() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__42() { _start: { lean_object* x_1; @@ -4094,19 +4236,19 @@ 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___lambda__2___closed__43() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__43() { _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_termMacro_x2etrace_x5b___x5d_____closed__1; x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__1; -x_3 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__9; -x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__42; +x_3 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__9; +x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__42; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__44() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__44() { _start: { lean_object* x_1; @@ -4114,7 +4256,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___lambda__2___closed__45() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__45() { _start: { lean_object* x_1; lean_object* x_2; @@ -4123,13 +4265,13 @@ x_2 = l_Array_append___rarg(x_1, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__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___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__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; uint8_t x_9; x_6 = lean_unsigned_to_nat(1u); x_7 = l_Lean_Syntax_getArg(x_1, x_6); -x_8 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__1; +x_8 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__1; lean_inc(x_7); x_9 = l_Lean_Syntax_isOfKind(x_7, x_8); if (x_9 == 0) @@ -4155,8 +4297,8 @@ x_16 = lean_unsigned_to_nat(6u); 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___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__2; -x_20 = l_Array_sequenceMap___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___spec__1(x_18, x_19); +x_19 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__2; +x_20 = l_Array_sequenceMap___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___spec__1(x_18, x_19); lean_dec(x_18); if (lean_obj_tag(x_20) == 0) { @@ -4183,8 +4325,8 @@ x_25 = lean_usize_of_nat(x_24); lean_dec(x_24); x_26 = 0; lean_inc(x_23); -x_27 = l_Array_mapMUnsafe_map___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___spec__3(x_25, x_26, x_23); -x_28 = l_Array_mapMUnsafe_map___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___spec__4(x_25, x_26, x_23); +x_27 = l_Array_mapMUnsafe_map___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___spec__3(x_25, x_26, x_23); +x_28 = l_Array_mapMUnsafe_map___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___spec__4(x_25, x_26, x_23); x_29 = lean_unsigned_to_nat(8u); x_30 = l_Lean_Syntax_getArg(x_1, x_29); x_31 = l_Lean_unifConstraint___closed__2; @@ -4255,12 +4397,12 @@ lean_inc(x_42); x_43 = 0; x_44 = l_Lean_SourceInfo_fromRef(x_42, x_43); x_45 = lean_box(0); -x_46 = l_Array_forInUnsafe_loop___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___spec__5___closed__3; +x_46 = l_Array_forInUnsafe_loop___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___spec__5___closed__3; lean_inc(x_44); x_47 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_47, 0, x_44); lean_ctor_set(x_47, 1, x_46); -x_48 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__5; +x_48 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__5; lean_inc(x_44); x_49 = l_Lean_Syntax_node3(x_44, x_48, x_36, x_47, x_38); x_50 = l_Array_zip___rarg(x_28, x_27); @@ -4270,9 +4412,9 @@ x_51 = l_Array_reverse___rarg(x_50); x_52 = lean_array_get_size(x_51); x_53 = lean_usize_of_nat(x_52); lean_dec(x_52); -x_54 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__3; +x_54 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__3; lean_inc(x_4); -x_55 = l_Array_forInUnsafe_loop___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___spec__5(x_54, x_48, x_51, x_53, x_26, x_49, x_4, x_5); +x_55 = l_Array_forInUnsafe_loop___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___spec__5(x_54, x_48, x_51, x_53, x_26, x_49, x_4, x_5); lean_dec(x_51); x_56 = !lean_is_exclusive(x_55); if (x_56 == 0) @@ -4284,25 +4426,25 @@ lean_inc(x_58); x_59 = lean_ctor_get(x_4, 1); lean_inc(x_59); lean_dec(x_4); -x_60 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__8; +x_60 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__8; lean_inc(x_58); lean_inc(x_59); x_61 = l_Lean_addMacroScope(x_59, x_60, x_58); -x_62 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__7; +x_62 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__7; lean_inc(x_44); x_63 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_63, 0, x_44); lean_ctor_set(x_63, 1, x_62); lean_ctor_set(x_63, 2, x_61); lean_ctor_set(x_63, 3, x_45); -x_64 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__16; +x_64 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__16; lean_inc(x_44); x_65 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_65, 0, x_44); 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___lambda__2___closed__24; +x_66 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__24; x_67 = l_Lean_addMacroScope(x_59, x_66, x_58); -x_68 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__23; +x_68 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__23; lean_inc(x_44); x_69 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_69, 0, x_44); @@ -4316,11 +4458,11 @@ x_72 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_72, 0, x_44); lean_ctor_set(x_72, 1, x_70); lean_ctor_set(x_72, 2, x_71); -x_73 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__21; +x_73 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__21; lean_inc(x_72); lean_inc(x_44); x_74 = l_Lean_Syntax_node2(x_44, x_73, x_69, x_72); -x_75 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__18; +x_75 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__18; lean_inc(x_44); x_76 = l_Lean_Syntax_node2(x_44, x_75, x_7, x_74); lean_inc(x_44); @@ -4330,12 +4472,12 @@ lean_inc(x_44); x_79 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_79, 0, x_44); lean_ctor_set(x_79, 1, x_78); -x_80 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__15; +x_80 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__15; lean_inc(x_44); x_81 = l_Lean_Syntax_node3(x_44, x_80, x_65, x_77, x_79); lean_inc(x_44); x_82 = l_Lean_Syntax_node1(x_44, x_70, x_81); -x_83 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__25; +x_83 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__25; lean_inc(x_44); x_84 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_84, 0, x_44); @@ -4343,13 +4485,13 @@ lean_ctor_set(x_84, 1, x_83); x_85 = lean_array_get_size(x_39); x_86 = lean_usize_of_nat(x_85); lean_dec(x_85); -x_87 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__35; -x_88 = l_Array_mapMUnsafe_map___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___spec__6(x_87, x_86, x_26, x_39); +x_87 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__35; +x_88 = l_Array_mapMUnsafe_map___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___spec__6(x_87, x_86, x_26, x_39); x_89 = lean_array_get_size(x_88); x_90 = lean_usize_of_nat(x_89); lean_dec(x_89); -x_91 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__36; -x_92 = l_Array_mapMUnsafe_map___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___spec__7(x_91, x_90, x_26, x_88); +x_91 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__36; +x_92 = l_Array_mapMUnsafe_map___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___spec__7(x_91, x_90, x_26, x_88); x_93 = l_Array_append___rarg(x_71, x_92); lean_inc(x_44); x_94 = lean_alloc_ctor(1, 3, 0); @@ -4361,7 +4503,7 @@ lean_inc(x_44); x_96 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_96, 0, x_44); lean_ctor_set(x_96, 1, x_95); -x_97 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__39; +x_97 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__39; lean_inc(x_44); x_98 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_98, 0, x_44); @@ -4371,12 +4513,12 @@ lean_inc(x_44); x_100 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_100, 0, x_44); lean_ctor_set(x_100, 1, x_99); -x_101 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__41; +x_101 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__41; lean_inc(x_44); x_102 = l_Lean_Syntax_node1(x_44, x_101, x_100); lean_inc(x_44); x_103 = l_Lean_Syntax_node1(x_44, x_70, x_102); -x_104 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__38; +x_104 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__38; lean_inc(x_44); x_105 = l_Lean_Syntax_node2(x_44, x_104, x_98, x_103); x_106 = l_Lean_expandExplicitBindersAux_loop___closed__11; @@ -4384,49 +4526,49 @@ lean_inc(x_44); x_107 = l_Lean_Syntax_node2(x_44, x_106, x_96, x_105); lean_inc(x_44); x_108 = l_Lean_Syntax_node1(x_44, x_70, x_107); -x_109 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__32; +x_109 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__32; lean_inc(x_44); x_110 = l_Lean_Syntax_node2(x_44, x_109, x_94, x_108); -x_111 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__44; +x_111 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__44; lean_inc(x_44); x_112 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_112, 0, x_44); lean_ctor_set(x_112, 1, x_111); -x_113 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__43; +x_113 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__43; lean_inc(x_72); lean_inc(x_44); x_114 = l_Lean_Syntax_node3(x_44, x_113, x_112, x_57, x_72); if (lean_obj_tag(x_3) == 0) { lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; -x_115 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__45; +x_115 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__45; lean_inc(x_44); x_116 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_116, 0, x_44); lean_ctor_set(x_116, 1, x_70); lean_ctor_set(x_116, 2, x_115); -x_117 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__13; +x_117 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__13; lean_inc_n(x_72, 4); lean_inc(x_44); x_118 = l_Lean_Syntax_node6(x_44, x_117, x_116, x_82, x_72, x_72, x_72, x_72); if (lean_obj_tag(x_41) == 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; -x_119 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__30; +x_119 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__30; x_120 = lean_array_push(x_119, x_63); -x_121 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__29; +x_121 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__29; x_122 = lean_array_push(x_120, x_121); x_123 = lean_box(2); -x_124 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__28; +x_124 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__28; 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_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__26; +x_126 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__26; lean_inc_n(x_72, 2); lean_inc(x_44); x_127 = l_Lean_Syntax_node7(x_44, x_126, x_84, x_125, x_110, x_114, x_72, x_72, x_72); -x_128 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__11; +x_128 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__11; x_129 = l_Lean_Syntax_node2(x_44, x_128, x_118, x_127); lean_ctor_set(x_55, 0, x_129); return x_55; @@ -4438,21 +4580,21 @@ lean_dec(x_63); x_130 = lean_ctor_get(x_41, 0); lean_inc(x_130); lean_dec(x_41); -x_131 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__30; +x_131 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__30; x_132 = lean_array_push(x_131, x_130); -x_133 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__29; +x_133 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__29; x_134 = lean_array_push(x_132, x_133); x_135 = lean_box(2); -x_136 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__28; +x_136 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__28; x_137 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_137, 0, x_135); lean_ctor_set(x_137, 1, x_136); lean_ctor_set(x_137, 2, x_134); -x_138 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__26; +x_138 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__26; lean_inc_n(x_72, 2); lean_inc(x_44); x_139 = l_Lean_Syntax_node7(x_44, x_138, x_84, x_137, x_110, x_114, x_72, x_72, x_72); -x_140 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__11; +x_140 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__11; x_141 = l_Lean_Syntax_node2(x_44, x_140, x_118, x_139); lean_ctor_set(x_55, 0, x_141); return x_55; @@ -4471,28 +4613,28 @@ x_145 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_145, 0, x_44); lean_ctor_set(x_145, 1, x_70); lean_ctor_set(x_145, 2, x_144); -x_146 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__13; +x_146 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__13; lean_inc_n(x_72, 4); lean_inc(x_44); x_147 = l_Lean_Syntax_node6(x_44, x_146, x_145, x_82, x_72, x_72, x_72, x_72); if (lean_obj_tag(x_41) == 0) { 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_148 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__30; +x_148 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__30; x_149 = lean_array_push(x_148, x_63); -x_150 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__29; +x_150 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__29; x_151 = lean_array_push(x_149, x_150); x_152 = lean_box(2); -x_153 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__28; +x_153 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__28; 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_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__26; +x_155 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__26; lean_inc_n(x_72, 2); lean_inc(x_44); x_156 = l_Lean_Syntax_node7(x_44, x_155, x_84, x_154, x_110, x_114, x_72, x_72, x_72); -x_157 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__11; +x_157 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__11; x_158 = l_Lean_Syntax_node2(x_44, x_157, x_147, x_156); lean_ctor_set(x_55, 0, x_158); return x_55; @@ -4504,21 +4646,21 @@ lean_dec(x_63); x_159 = lean_ctor_get(x_41, 0); lean_inc(x_159); lean_dec(x_41); -x_160 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__30; +x_160 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__30; x_161 = lean_array_push(x_160, x_159); -x_162 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__29; +x_162 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__29; x_163 = lean_array_push(x_161, x_162); x_164 = lean_box(2); -x_165 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__28; +x_165 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__28; 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); -x_167 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__26; +x_167 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__26; lean_inc_n(x_72, 2); lean_inc(x_44); x_168 = l_Lean_Syntax_node7(x_44, x_167, x_84, x_166, x_110, x_114, x_72, x_72, x_72); -x_169 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__11; +x_169 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__11; x_170 = l_Lean_Syntax_node2(x_44, x_169, x_147, x_168); lean_ctor_set(x_55, 0, x_170); return x_55; @@ -4538,25 +4680,25 @@ lean_inc(x_173); x_174 = lean_ctor_get(x_4, 1); lean_inc(x_174); lean_dec(x_4); -x_175 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__8; +x_175 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__8; lean_inc(x_173); lean_inc(x_174); x_176 = l_Lean_addMacroScope(x_174, x_175, x_173); -x_177 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__7; +x_177 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__7; lean_inc(x_44); x_178 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_178, 0, x_44); lean_ctor_set(x_178, 1, x_177); lean_ctor_set(x_178, 2, x_176); lean_ctor_set(x_178, 3, x_45); -x_179 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__16; +x_179 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__16; lean_inc(x_44); x_180 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_180, 0, x_44); lean_ctor_set(x_180, 1, x_179); -x_181 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__24; +x_181 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__24; x_182 = l_Lean_addMacroScope(x_174, x_181, x_173); -x_183 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__23; +x_183 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__23; lean_inc(x_44); x_184 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_184, 0, x_44); @@ -4570,11 +4712,11 @@ x_187 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_187, 0, x_44); lean_ctor_set(x_187, 1, x_185); lean_ctor_set(x_187, 2, x_186); -x_188 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__21; +x_188 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__21; lean_inc(x_187); lean_inc(x_44); x_189 = l_Lean_Syntax_node2(x_44, x_188, x_184, x_187); -x_190 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__18; +x_190 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__18; lean_inc(x_44); x_191 = l_Lean_Syntax_node2(x_44, x_190, x_7, x_189); lean_inc(x_44); @@ -4584,12 +4726,12 @@ lean_inc(x_44); x_194 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_194, 0, x_44); lean_ctor_set(x_194, 1, x_193); -x_195 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__15; +x_195 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__15; lean_inc(x_44); x_196 = l_Lean_Syntax_node3(x_44, x_195, x_180, x_192, x_194); lean_inc(x_44); x_197 = l_Lean_Syntax_node1(x_44, x_185, x_196); -x_198 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__25; +x_198 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__25; lean_inc(x_44); x_199 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_199, 0, x_44); @@ -4597,13 +4739,13 @@ lean_ctor_set(x_199, 1, x_198); x_200 = lean_array_get_size(x_39); x_201 = lean_usize_of_nat(x_200); lean_dec(x_200); -x_202 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__35; -x_203 = l_Array_mapMUnsafe_map___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___spec__6(x_202, x_201, x_26, x_39); +x_202 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__35; +x_203 = l_Array_mapMUnsafe_map___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___spec__6(x_202, x_201, x_26, x_39); x_204 = lean_array_get_size(x_203); x_205 = lean_usize_of_nat(x_204); lean_dec(x_204); -x_206 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__36; -x_207 = l_Array_mapMUnsafe_map___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___spec__7(x_206, x_205, x_26, x_203); +x_206 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__36; +x_207 = l_Array_mapMUnsafe_map___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___spec__7(x_206, x_205, x_26, x_203); x_208 = l_Array_append___rarg(x_186, x_207); lean_inc(x_44); x_209 = lean_alloc_ctor(1, 3, 0); @@ -4615,7 +4757,7 @@ lean_inc(x_44); x_211 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_211, 0, x_44); lean_ctor_set(x_211, 1, x_210); -x_212 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__39; +x_212 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__39; lean_inc(x_44); x_213 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_213, 0, x_44); @@ -4625,12 +4767,12 @@ lean_inc(x_44); x_215 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_215, 0, x_44); lean_ctor_set(x_215, 1, x_214); -x_216 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__41; +x_216 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__41; lean_inc(x_44); x_217 = l_Lean_Syntax_node1(x_44, x_216, x_215); lean_inc(x_44); x_218 = l_Lean_Syntax_node1(x_44, x_185, x_217); -x_219 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__38; +x_219 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__38; lean_inc(x_44); x_220 = l_Lean_Syntax_node2(x_44, x_219, x_213, x_218); x_221 = l_Lean_expandExplicitBindersAux_loop___closed__11; @@ -4638,49 +4780,49 @@ lean_inc(x_44); x_222 = l_Lean_Syntax_node2(x_44, x_221, x_211, x_220); lean_inc(x_44); x_223 = l_Lean_Syntax_node1(x_44, x_185, x_222); -x_224 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__32; +x_224 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__32; lean_inc(x_44); x_225 = l_Lean_Syntax_node2(x_44, x_224, x_209, x_223); -x_226 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__44; +x_226 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__44; lean_inc(x_44); x_227 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_227, 0, x_44); lean_ctor_set(x_227, 1, x_226); -x_228 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__43; +x_228 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__43; lean_inc(x_187); lean_inc(x_44); x_229 = l_Lean_Syntax_node3(x_44, x_228, x_227, x_171, x_187); if (lean_obj_tag(x_3) == 0) { lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; -x_230 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__45; +x_230 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__45; lean_inc(x_44); x_231 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_231, 0, x_44); lean_ctor_set(x_231, 1, x_185); lean_ctor_set(x_231, 2, x_230); -x_232 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__13; +x_232 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__13; lean_inc_n(x_187, 4); lean_inc(x_44); x_233 = l_Lean_Syntax_node6(x_44, x_232, x_231, x_197, x_187, x_187, x_187, x_187); if (lean_obj_tag(x_41) == 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; -x_234 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__30; +x_234 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__30; x_235 = lean_array_push(x_234, x_178); -x_236 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__29; +x_236 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__29; x_237 = lean_array_push(x_235, x_236); x_238 = lean_box(2); -x_239 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__28; +x_239 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__28; 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__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__26; +x_241 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__26; lean_inc_n(x_187, 2); lean_inc(x_44); x_242 = l_Lean_Syntax_node7(x_44, x_241, x_199, x_240, x_225, x_229, x_187, x_187, x_187); -x_243 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__11; +x_243 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__11; x_244 = l_Lean_Syntax_node2(x_44, x_243, x_233, x_242); x_245 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_245, 0, x_244); @@ -4694,21 +4836,21 @@ lean_dec(x_178); x_246 = lean_ctor_get(x_41, 0); lean_inc(x_246); lean_dec(x_41); -x_247 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__30; +x_247 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__30; x_248 = lean_array_push(x_247, x_246); -x_249 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__29; +x_249 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__29; x_250 = lean_array_push(x_248, x_249); x_251 = lean_box(2); -x_252 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__28; +x_252 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__28; x_253 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_253, 0, x_251); lean_ctor_set(x_253, 1, x_252); lean_ctor_set(x_253, 2, x_250); -x_254 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__26; +x_254 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__26; lean_inc_n(x_187, 2); lean_inc(x_44); x_255 = l_Lean_Syntax_node7(x_44, x_254, x_199, x_253, x_225, x_229, x_187, x_187, x_187); -x_256 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__11; +x_256 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__11; x_257 = l_Lean_Syntax_node2(x_44, x_256, x_233, x_255); x_258 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_258, 0, x_257); @@ -4729,28 +4871,28 @@ x_262 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_262, 0, x_44); lean_ctor_set(x_262, 1, x_185); lean_ctor_set(x_262, 2, x_261); -x_263 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__13; +x_263 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__13; lean_inc_n(x_187, 4); lean_inc(x_44); x_264 = l_Lean_Syntax_node6(x_44, x_263, x_262, x_197, x_187, x_187, x_187, x_187); if (lean_obj_tag(x_41) == 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; 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_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__30; +x_265 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__30; x_266 = lean_array_push(x_265, x_178); -x_267 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__29; +x_267 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__29; x_268 = lean_array_push(x_266, x_267); x_269 = lean_box(2); -x_270 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__28; +x_270 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__28; 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); -x_272 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__26; +x_272 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__26; lean_inc_n(x_187, 2); lean_inc(x_44); x_273 = l_Lean_Syntax_node7(x_44, x_272, x_199, x_271, x_225, x_229, x_187, x_187, x_187); -x_274 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__11; +x_274 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__11; x_275 = l_Lean_Syntax_node2(x_44, x_274, x_264, x_273); x_276 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_276, 0, x_275); @@ -4764,21 +4906,21 @@ lean_dec(x_178); x_277 = lean_ctor_get(x_41, 0); lean_inc(x_277); lean_dec(x_41); -x_278 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__30; +x_278 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__30; x_279 = lean_array_push(x_278, x_277); -x_280 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__29; +x_280 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__29; x_281 = lean_array_push(x_279, x_280); x_282 = lean_box(2); -x_283 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__28; +x_283 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__28; x_284 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_284, 0, x_282); lean_ctor_set(x_284, 1, x_283); lean_ctor_set(x_284, 2, x_281); -x_285 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__26; +x_285 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__26; lean_inc_n(x_187, 2); lean_inc(x_44); x_286 = l_Lean_Syntax_node7(x_44, x_285, x_199, x_284, x_225, x_229, x_187, x_187, x_187); -x_287 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__11; +x_287 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__11; x_288 = l_Lean_Syntax_node2(x_44, x_287, x_264, x_286); x_289 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_289, 0, x_288); @@ -4793,23 +4935,23 @@ return x_289; } } } -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__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_termMacro_x2etrace_x5b___x5d_____closed__1; x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__1; -x_3 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__9; -x_4 = l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__3; +x_3 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__9; +x_4 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__3; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -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) { +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; +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) @@ -4852,7 +4994,7 @@ else lean_object* x_15; lean_object* x_16; uint8_t x_17; x_15 = l_Lean_Syntax_getArg(x_9, x_8); lean_dec(x_9); -x_16 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___closed__1; +x_16 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___closed__1; lean_inc(x_15); x_17 = l_Lean_Syntax_isOfKind(x_15, x_16); if (x_17 == 0) @@ -4873,7 +5015,7 @@ lean_object* x_20; lean_object* x_21; lean_object* x_22; x_20 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_20, 0, x_15); x_21 = lean_box(0); -x_22 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2(x_1, x_21, x_20, x_2, x_3); +x_22 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2(x_1, x_21, x_20, x_2, x_3); lean_dec(x_1); return x_22; } @@ -4885,32 +5027,32 @@ lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_dec(x_9); x_23 = lean_box(0); x_24 = lean_box(0); -x_25 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2(x_1, x_24, x_23, x_2, x_3); +x_25 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2(x_1, x_24, x_23, x_2, x_3); lean_dec(x_1); return x_25; } } } } -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) { +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); +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) { +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); +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) { +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; @@ -4918,11 +5060,11 @@ 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__3(x_4, x_5, x_3); +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) { +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; @@ -4930,11 +5072,11 @@ 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); +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_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) { _start: { size_t x_9; size_t x_10; lean_object* x_11; @@ -4942,13 +5084,13 @@ 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___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___spec__5(x_1, x_2, x_3, x_9, x_10, x_6, x_7, x_8); +x_11 = 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_9, x_10, x_6, x_7, x_8); lean_dec(x_3); lean_dec(x_1); return x_11; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___spec__6___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_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___spec__6___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; @@ -4956,12 +5098,12 @@ 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_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___spec__6(x_1, x_5, x_6, x_4); +x_7 = l_Array_mapMUnsafe_map___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___spec__6(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___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___spec__7___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_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___spec__7___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; @@ -4969,16 +5111,16 @@ 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_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___spec__7(x_1, x_5, x_6, x_4); +x_7 = l_Array_mapMUnsafe_map___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___spec__7(x_1, x_5, x_6, x_4); lean_dec(x_1); return x_7; } } -LEAN_EXPORT lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____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_EXPORT lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__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_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2(x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2(x_1, x_2, x_3, x_4, x_5); lean_dec(x_2); lean_dec(x_1); return x_6; @@ -5006,7 +5148,7 @@ static lean_object* _init_l_term_u2203___x2c_____closed__3() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("∃ ", 4); +x_1 = lean_mk_string_from_bytes("∃", 3); return x_1; } } @@ -6110,7 +6252,7 @@ static lean_object* _init_l_calcFirstStep___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_unbracketedExplicitBinders___closed__7; +x_1 = l_Lean_unbracketedExplicitBinders___closed__11; x_2 = l_calcFirstStep___closed__8; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -6296,7 +6438,7 @@ static lean_object* _init_l_calcSteps___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__27; +x_1 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__29; x_2 = l_calcFirstStep; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -6323,8 +6465,8 @@ static lean_object* _init_l_calcSteps___closed__8() { { 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__5; -x_2 = l_calcSteps___closed__7; -x_3 = l_calcSteps___closed__5; +x_2 = l_calcSteps___closed__5; +x_3 = l_calcStep; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -6335,48 +6477,34 @@ return x_4; static lean_object* _init_l_calcSteps___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__5; -x_2 = l_calcStep; -x_3 = l_calcSteps___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); -lean_ctor_set(x_4, 2, x_3); -return x_4; -} -} -static lean_object* _init_l_calcSteps___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__17; -x_2 = l_calcSteps___closed__9; +x_1 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__18; +x_2 = l_calcSteps___closed__8; 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_calcSteps___closed__11() { +static lean_object* _init_l_calcSteps___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__27; -x_2 = l_calcSteps___closed__10; +x_1 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__29; +x_2 = l_calcSteps___closed__9; 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_calcSteps___closed__12() { +static lean_object* _init_l_calcSteps___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__5; -x_2 = l_calcSteps___closed__8; -x_3 = l_calcSteps___closed__11; +x_2 = l_calcSteps___closed__7; +x_3 = l_calcSteps___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); @@ -6384,13 +6512,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_calcSteps___closed__13() { +static lean_object* _init_l_calcSteps___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_calcSteps___closed__1; x_2 = l_calcSteps___closed__2; -x_3 = l_calcSteps___closed__12; +x_3 = l_calcSteps___closed__11; x_4 = lean_alloc_ctor(9, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -6402,7 +6530,7 @@ static lean_object* _init_l_calcSteps() { _start: { lean_object* x_1; -x_1 = l_calcSteps___closed__13; +x_1 = l_calcSteps___closed__12; return x_1; } } @@ -7560,24 +7688,16 @@ static lean_object* _init_l_unexpandExists___closed__3() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("∃", 3); -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() { +static lean_object* _init_l_unexpandExists___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__1; -x_2 = l_unexpandExists___closed__4; +x_2 = l_unexpandExists___closed__3; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -7769,7 +7889,7 @@ x_49 = l_Lean_Syntax_getArg(x_19, x_39); lean_dec(x_19); x_50 = 0; x_51 = l_Lean_SourceInfo_fromRef(x_2, x_50); -x_52 = l_unexpandExists___closed__3; +x_52 = l_term_u2203___x2c_____closed__3; lean_inc(x_51); x_53 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_53, 0, x_51); @@ -7779,7 +7899,7 @@ lean_inc(x_51); x_55 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_55, 0, x_51); lean_ctor_set(x_55, 1, x_54); -x_56 = l_unexpandExists___closed__5; +x_56 = l_unexpandExists___closed__4; lean_inc(x_51); x_57 = l_Lean_Syntax_node1(x_51, x_56, x_35); x_58 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__16; @@ -7850,12 +7970,12 @@ if (x_81 == 0) uint8_t 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; x_82 = 0; x_83 = l_Lean_SourceInfo_fromRef(x_2, x_82); -x_84 = l_unexpandExists___closed__3; +x_84 = l_term_u2203___x2c_____closed__3; lean_inc(x_83); x_85 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_85, 0, x_83); lean_ctor_set(x_85, 1, x_84); -x_86 = l_unexpandExists___closed__5; +x_86 = l_unexpandExists___closed__4; lean_inc(x_83); x_87 = l_Lean_Syntax_node1(x_83, x_86, x_28); x_88 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__16; @@ -7897,12 +8017,12 @@ uint8_t x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_ lean_dec(x_100); x_103 = 0; x_104 = l_Lean_SourceInfo_fromRef(x_2, x_103); -x_105 = l_unexpandExists___closed__3; +x_105 = l_term_u2203___x2c_____closed__3; lean_inc(x_104); x_106 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_106, 0, x_104); lean_ctor_set(x_106, 1, x_105); -x_107 = l_unexpandExists___closed__5; +x_107 = l_unexpandExists___closed__4; lean_inc(x_104); x_108 = l_Lean_Syntax_node1(x_104, x_107, x_28); x_109 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__16; @@ -7944,12 +8064,12 @@ uint8_t x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_ lean_dec(x_120); x_123 = 0; x_124 = l_Lean_SourceInfo_fromRef(x_2, x_123); -x_125 = l_unexpandExists___closed__3; +x_125 = l_term_u2203___x2c_____closed__3; 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_unexpandExists___closed__5; +x_127 = l_unexpandExists___closed__4; lean_inc(x_124); x_128 = l_Lean_Syntax_node1(x_124, x_127, x_28); x_129 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__16; @@ -7989,12 +8109,12 @@ uint8_t x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_ lean_dec(x_139); x_142 = 0; x_143 = l_Lean_SourceInfo_fromRef(x_2, x_142); -x_144 = l_unexpandExists___closed__3; +x_144 = l_term_u2203___x2c_____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_unexpandExists___closed__5; +x_146 = l_unexpandExists___closed__4; lean_inc(x_143); x_147 = l_Lean_Syntax_node1(x_143, x_146, x_28); x_148 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__16; @@ -8030,12 +8150,12 @@ x_159 = l_Lean_Syntax_getArgs(x_139); lean_dec(x_139); x_160 = 0; x_161 = l_Lean_SourceInfo_fromRef(x_2, x_160); -x_162 = l_unexpandExists___closed__3; +x_162 = l_term_u2203___x2c_____closed__3; lean_inc(x_161); x_163 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_163, 0, x_161); lean_ctor_set(x_163, 1, x_162); -x_164 = l_unexpandExists___closed__5; +x_164 = l_unexpandExists___closed__4; lean_inc(x_161); x_165 = l_Lean_Syntax_node1(x_161, x_164, x_28); x_166 = l_Array_mkArray1___rarg(x_165); @@ -8274,7 +8394,7 @@ lean_inc(x_50); x_52 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_52, 0, x_50); lean_ctor_set(x_52, 1, x_51); -x_53 = l_unexpandExists___closed__5; +x_53 = l_unexpandExists___closed__4; lean_inc(x_50); x_54 = l_Lean_Syntax_node1(x_50, x_53, x_33); x_55 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__16; @@ -8510,7 +8630,7 @@ lean_inc(x_50); x_52 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_52, 0, x_50); lean_ctor_set(x_52, 1, x_51); -x_53 = l_unexpandExists___closed__5; +x_53 = l_unexpandExists___closed__4; lean_inc(x_50); x_54 = l_Lean_Syntax_node1(x_50, x_53, x_33); x_55 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__16; @@ -11906,7 +12026,7 @@ static lean_object* _init_l_tacticFunext_________closed__5() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("ppSpace", 7); +x_1 = lean_mk_string_from_bytes("colGt", 5); return x_1; } } @@ -11933,38 +12053,10 @@ return x_2; static lean_object* _init_l_tacticFunext_________closed__8() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("colGt", 5); -return x_1; -} -} -static lean_object* _init_l_tacticFunext_________closed__9() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_tacticFunext_________closed__8; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_tacticFunext_________closed__10() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_tacticFunext_________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_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__5; -x_2 = l_tacticFunext_________closed__7; -x_3 = l_tacticFunext_________closed__10; +x_2 = l_Lean_unbracketedExplicitBinders___closed__7; +x_3 = l_tacticFunext_________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); @@ -11972,7 +12064,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_tacticFunext_________closed__12() { +static lean_object* _init_l_tacticFunext_________closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -11984,13 +12076,13 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_tacticFunext_________closed__13() { +static lean_object* _init_l_tacticFunext_________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__5; -x_2 = l_tacticFunext_________closed__11; -x_3 = l_tacticFunext_________closed__12; +x_2 = l_tacticFunext_________closed__8; +x_3 = l_tacticFunext_________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); @@ -11998,25 +12090,25 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_tacticFunext_________closed__14() { +static lean_object* _init_l_tacticFunext_________closed__11() { _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__17; -x_2 = l_tacticFunext_________closed__13; +x_1 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__18; +x_2 = l_tacticFunext_________closed__10; 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_tacticFunext_________closed__15() { +static lean_object* _init_l_tacticFunext_________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__5; x_2 = l_tacticFunext_________closed__4; -x_3 = l_tacticFunext_________closed__14; +x_3 = l_tacticFunext_________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); @@ -12024,13 +12116,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_tacticFunext_________closed__16() { +static lean_object* _init_l_tacticFunext_________closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_tacticFunext_________closed__2; x_2 = lean_unsigned_to_nat(1022u); -x_3 = l_tacticFunext_________closed__15; +x_3 = l_tacticFunext_________closed__12; x_4 = lean_alloc_ctor(3, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -12042,7 +12134,7 @@ static lean_object* _init_l_tacticFunext______() { _start: { lean_object* x_1; -x_1 = l_tacticFunext_________closed__16; +x_1 = l_tacticFunext_________closed__13; return x_1; } } @@ -13158,7 +13250,7 @@ x_44 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_44, 0, x_33); lean_ctor_set(x_44, 1, x_42); lean_ctor_set(x_44, 2, x_43); -x_45 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__44; +x_45 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__44; lean_inc(x_33); x_46 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_46, 0, x_33); @@ -13253,7 +13345,7 @@ size_t x_81; size_t x_82; lean_object* x_83; lean_object* x_84; uint8_t x_85; x_81 = lean_usize_of_nat(x_23); lean_dec(x_23); x_82 = 0; -x_83 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__3; +x_83 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__3; x_84 = l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__4(x_83, x_19, x_81, x_82, x_21, x_2, x_3); lean_dec(x_19); x_85 = !lean_is_exclusive(x_84); @@ -13297,7 +13389,7 @@ size_t x_91; size_t x_92; lean_object* x_93; lean_object* x_94; uint8_t x_95; x_91 = lean_usize_of_nat(x_23); lean_dec(x_23); x_92 = 0; -x_93 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__3; +x_93 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__3; x_94 = l_Array_foldrMUnsafe_fold___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__5(x_93, x_19, x_91, x_92, x_21, x_2, x_3); lean_dec(x_19); x_95 = !lean_is_exclusive(x_94); @@ -13398,7 +13490,7 @@ static lean_object* _init_l_Lean_Parser_Command_classAbbrev___closed__2() { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__1; x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__1; -x_3 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__9; +x_3 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__9; x_4 = l_Lean_Parser_Command_classAbbrev___closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; @@ -13409,7 +13501,7 @@ static lean_object* _init_l_Lean_Parser_Command_classAbbrev___closed__3() { { 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___lambda__2___closed__12; +x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__12; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -13493,7 +13585,7 @@ static lean_object* _init_l_Lean_Parser_Command_classAbbrev___closed__11() { { 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___lambda__2___closed__27; +x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__27; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -13525,10 +13617,22 @@ return x_4; static lean_object* _init_l_Lean_Parser_Command_classAbbrev___closed__14() { _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__18; +x_2 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____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); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Command_classAbbrev___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__5; x_2 = l_Lean_Parser_Command_classAbbrev___closed__13; -x_3 = l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__21; +x_3 = l_Lean_Parser_Command_classAbbrev___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); @@ -13536,7 +13640,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Command_classAbbrev___closed__15() { +static lean_object* _init_l_Lean_Parser_Command_classAbbrev___closed__16() { _start: { lean_object* x_1; lean_object* x_2; @@ -13546,12 +13650,12 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_classAbbrev___closed__16() { +static lean_object* _init_l_Lean_Parser_Command_classAbbrev___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__5; -x_2 = l_Lean_Parser_Command_classAbbrev___closed__15; +x_2 = l_Lean_Parser_Command_classAbbrev___closed__16; x_3 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__19; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -13560,25 +13664,25 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Command_classAbbrev___closed__17() { +static lean_object* _init_l_Lean_Parser_Command_classAbbrev___closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_unbracketedExplicitBinders___closed__7; -x_2 = l_Lean_Parser_Command_classAbbrev___closed__16; +x_1 = l_Lean_unbracketedExplicitBinders___closed__11; +x_2 = l_Lean_Parser_Command_classAbbrev___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_Command_classAbbrev___closed__18() { +static lean_object* _init_l_Lean_Parser_Command_classAbbrev___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__5; -x_2 = l_Lean_Parser_Command_classAbbrev___closed__14; -x_3 = l_Lean_Parser_Command_classAbbrev___closed__17; +x_2 = l_Lean_Parser_Command_classAbbrev___closed__15; +x_3 = l_Lean_Parser_Command_classAbbrev___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); @@ -13586,23 +13690,23 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Command_classAbbrev___closed__19() { +static lean_object* _init_l_Lean_Parser_Command_classAbbrev___closed__20() { _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___lambda__2___closed__44; +x_1 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__44; 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_Command_classAbbrev___closed__20() { +static lean_object* _init_l_Lean_Parser_Command_classAbbrev___closed__21() { _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__5; -x_2 = l_Lean_Parser_Command_classAbbrev___closed__18; -x_3 = l_Lean_Parser_Command_classAbbrev___closed__19; +x_2 = l_Lean_Parser_Command_classAbbrev___closed__19; +x_3 = l_Lean_Parser_Command_classAbbrev___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); @@ -13610,7 +13714,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Command_classAbbrev___closed__21() { +static lean_object* _init_l_Lean_Parser_Command_classAbbrev___closed__22() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; @@ -13624,7 +13728,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Command_classAbbrev___closed__22() { +static lean_object* _init_l_Lean_Parser_Command_classAbbrev___closed__23() { _start: { lean_object* x_1; lean_object* x_2; @@ -13634,25 +13738,25 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_classAbbrev___closed__23() { +static lean_object* _init_l_Lean_Parser_Command_classAbbrev___closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_unbracketedExplicitBinders___closed__7; -x_2 = l_Lean_Parser_Command_classAbbrev___closed__22; +x_1 = l_Lean_unbracketedExplicitBinders___closed__11; +x_2 = l_Lean_Parser_Command_classAbbrev___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_Command_classAbbrev___closed__24() { +static lean_object* _init_l_Lean_Parser_Command_classAbbrev___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__5; -x_2 = l_Lean_Parser_Command_classAbbrev___closed__21; -x_3 = l_Lean_Parser_Command_classAbbrev___closed__23; +x_2 = l_Lean_Parser_Command_classAbbrev___closed__22; +x_3 = l_Lean_Parser_Command_classAbbrev___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); @@ -13660,49 +13764,49 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Command_classAbbrev___closed__25() { +static lean_object* _init_l_Lean_Parser_Command_classAbbrev___closed__26() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_unexpandGetElem_x21___closed__4; -x_2 = l_Lean_Parser_Command_classAbbrev___closed__24; +x_2 = l_Lean_Parser_Command_classAbbrev___closed__25; 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_Command_classAbbrev___closed__26() { +static lean_object* _init_l_Lean_Parser_Command_classAbbrev___closed__27() { _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__17; -x_2 = l_Lean_Parser_Command_classAbbrev___closed__25; +x_1 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__18; +x_2 = l_Lean_Parser_Command_classAbbrev___closed__26; 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_Command_classAbbrev___closed__27() { +static lean_object* _init_l_Lean_Parser_Command_classAbbrev___closed__28() { _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__27; -x_2 = l_Lean_Parser_Command_classAbbrev___closed__26; +x_1 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__29; +x_2 = l_Lean_Parser_Command_classAbbrev___closed__27; 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_Command_classAbbrev___closed__28() { +static lean_object* _init_l_Lean_Parser_Command_classAbbrev___closed__29() { _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__5; -x_2 = l_Lean_Parser_Command_classAbbrev___closed__20; -x_3 = l_Lean_Parser_Command_classAbbrev___closed__27; +x_2 = l_Lean_Parser_Command_classAbbrev___closed__21; +x_3 = l_Lean_Parser_Command_classAbbrev___closed__28; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -13710,13 +13814,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Command_classAbbrev___closed__29() { +static lean_object* _init_l_Lean_Parser_Command_classAbbrev___closed__30() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_classAbbrev___closed__2; x_2 = lean_unsigned_to_nat(1022u); -x_3 = l_Lean_Parser_Command_classAbbrev___closed__28; +x_3 = l_Lean_Parser_Command_classAbbrev___closed__29; x_4 = lean_alloc_ctor(3, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -13728,7 +13832,7 @@ static lean_object* _init_l_Lean_Parser_Command_classAbbrev() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Command_classAbbrev___closed__29; +x_1 = l_Lean_Parser_Command_classAbbrev___closed__30; return x_1; } } @@ -13816,7 +13920,7 @@ static lean_object* _init_l___aux__Init__NotationExtra______macroRules__Lean__Pa lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__1; x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__1; -x_3 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__9; +x_3 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__9; x_4 = l___aux__Init__NotationExtra______macroRules__Lean__Parser__Command__classAbbrev__1___lambda__3___closed__2; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; @@ -13836,7 +13940,7 @@ static lean_object* _init_l___aux__Init__NotationExtra______macroRules__Lean__Pa lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__1; x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__1; -x_3 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__9; +x_3 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__9; x_4 = l___aux__Init__NotationExtra______macroRules__Lean__Parser__Command__classAbbrev__1___lambda__3___closed__4; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; @@ -13864,7 +13968,7 @@ static lean_object* _init_l___aux__Init__NotationExtra______macroRules__Lean__Pa lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__1; x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__1; -x_3 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__9; +x_3 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__9; x_4 = l___aux__Init__NotationExtra______macroRules__Lean__Parser__Command__classAbbrev__1___lambda__3___closed__7; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; @@ -13884,7 +13988,7 @@ static lean_object* _init_l___aux__Init__NotationExtra______macroRules__Lean__Pa lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__1; x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__1; -x_3 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__9; +x_3 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__9; x_4 = l___aux__Init__NotationExtra______macroRules__Lean__Parser__Command__classAbbrev__1___lambda__3___closed__9; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; @@ -13904,7 +14008,7 @@ static lean_object* _init_l___aux__Init__NotationExtra______macroRules__Lean__Pa lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__1; x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__1; -x_3 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__9; +x_3 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__9; x_4 = l___aux__Init__NotationExtra______macroRules__Lean__Parser__Command__classAbbrev__1___lambda__3___closed__11; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; @@ -13924,7 +14028,7 @@ static lean_object* _init_l___aux__Init__NotationExtra______macroRules__Lean__Pa lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__1; x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__1; -x_3 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__19; +x_3 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__19; x_4 = l___aux__Init__NotationExtra______macroRules__Lean__Parser__Command__classAbbrev__1___lambda__3___closed__13; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; @@ -14096,7 +14200,7 @@ lean_inc(x_22); x_46 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_46, 0, x_22); 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___lambda__2___closed__1; +x_47 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__1; lean_inc(x_40); lean_inc(x_22); x_48 = l_Lean_Syntax_node1(x_22, x_47, x_40); @@ -14109,7 +14213,7 @@ x_51 = l___aux__Init__NotationExtra______macroRules__Lean__Parser__Command__clas lean_inc(x_40); lean_inc(x_22); x_52 = l_Lean_Syntax_node2(x_22, x_51, x_50, x_40); -x_53 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__18; +x_53 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__18; lean_inc(x_22); x_54 = l_Lean_Syntax_node2(x_22, x_53, x_48, x_52); lean_inc(x_22); @@ -14127,7 +14231,7 @@ x_60 = l_Lean_Syntax_node5(x_22, x_59, x_44, x_46, x_55, x_57, x_58); if (lean_obj_tag(x_6) == 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___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__45; +x_61 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__45; lean_inc(x_22); x_62 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_62, 0, x_22); @@ -14136,7 +14240,7 @@ lean_ctor_set(x_62, 2, x_61); x_63 = l___aux__Init__NotationExtra______macroRules__Lean__Parser__Command__classAbbrev__1___lambda__3___closed__3; lean_inc(x_22); x_64 = l_Lean_Syntax_node7(x_22, x_63, x_26, x_3, x_30, x_39, x_62, x_40, x_42); -x_65 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__11; +x_65 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__11; lean_inc(x_22); x_66 = l_Lean_Syntax_node2(x_22, x_65, x_4, x_64); x_67 = l_Lean_Syntax_node2(x_22, x_29, x_66, x_60); @@ -14169,7 +14273,7 @@ lean_ctor_set(x_76, 2, x_75); x_77 = l___aux__Init__NotationExtra______macroRules__Lean__Parser__Command__classAbbrev__1___lambda__3___closed__3; lean_inc(x_22); x_78 = l_Lean_Syntax_node7(x_22, x_77, x_26, x_3, x_30, x_39, x_76, x_40, x_42); -x_79 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__11; +x_79 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__11; lean_inc(x_22); x_80 = l_Lean_Syntax_node2(x_22, x_79, x_4, x_78); x_81 = l_Lean_Syntax_node2(x_22, x_29, x_80, x_60); @@ -14205,7 +14309,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___lambda__2___closed__13; +x_10 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__13; lean_inc(x_9); x_11 = l_Lean_Syntax_isOfKind(x_9, x_10); if (x_11 == 0) @@ -14554,40 +14658,84 @@ return x_3; static lean_object* _init_l_solve___closed__3() { _start: { +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l_solve___closed__1; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 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_solve___closed__4() { +_start: +{ lean_object* x_1; -x_1 = lean_mk_string_from_bytes("solve ", 6); +x_1 = lean_mk_string_from_bytes("ppDedent", 8); return x_1; } } -static lean_object* _init_l_solve___closed__4() { +static lean_object* _init_l_solve___closed__5() { _start: { -lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l_solve___closed__3; -x_2 = 0; -x_3 = lean_alloc_ctor(6, 1, 1); +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_solve___closed__4; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_solve___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_solve___closed__5; +x_2 = l_calcSteps___closed__5; +x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); -lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_solve___closed__5() { +static lean_object* _init_l_solve___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__5; +x_2 = l_solve___closed__6; +x_3 = l_Lean_unifConstraintElem___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); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_solve___closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("| ", 2); +return x_1; +} +} +static lean_object* _init_l_solve___closed__9() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__16; +x_1 = l_solve___closed__8; x_2 = lean_alloc_ctor(5, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_solve___closed__6() { +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__5; -x_2 = l_Lean_unifConstraintElem___closed__5; -x_3 = l_solve___closed__5; +x_2 = l_solve___closed__7; +x_3 = l_solve___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); @@ -14595,7 +14743,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_solve___closed__7() { +static lean_object* _init_l_solve___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -14605,23 +14753,23 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_solve___closed__8() { +static lean_object* _init_l_solve___closed__12() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_solve___closed__7; +x_1 = l_solve___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_solve___closed__9() { +static lean_object* _init_l_solve___closed__13() { _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__5; -x_2 = l_solve___closed__6; -x_3 = l_solve___closed__8; +x_2 = l_solve___closed__10; +x_3 = l_solve___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); @@ -14629,49 +14777,49 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_solve___closed__10() { +static lean_object* _init_l_solve___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_unexpandGetElem_x21___closed__4; -x_2 = l_solve___closed__9; +x_2 = l_solve___closed__13; 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_solve___closed__11() { +static lean_object* _init_l_solve___closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_unbracketedExplicitBinders___closed__4; -x_2 = l_solve___closed__10; +x_2 = l_solve___closed__14; 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_solve___closed__12() { +static lean_object* _init_l_solve___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__27; -x_2 = l_solve___closed__11; +x_1 = l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__29; +x_2 = l_solve___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_solve___closed__13() { +static lean_object* _init_l_solve___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__5; -x_2 = l_solve___closed__4; -x_3 = l_solve___closed__12; +x_2 = l_solve___closed__3; +x_3 = l_solve___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); @@ -14679,13 +14827,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_solve___closed__14() { +static lean_object* _init_l_solve___closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_solve___closed__2; x_2 = lean_unsigned_to_nat(1022u); -x_3 = l_solve___closed__13; +x_3 = l_solve___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); @@ -14697,7 +14845,7 @@ static lean_object* _init_l_solve() { _start: { lean_object* x_1; -x_1 = l_solve___closed__14; +x_1 = l_solve___closed__18; return x_1; } } @@ -15637,7 +15785,7 @@ static lean_object* _init_l_Lean_doElemWhile___x3a__Do_____closed__6() { 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__5; x_2 = l_Lean_doElemWhile___x3a__Do_____closed__5; -x_3 = l_Lean_unbracketedExplicitBinders___closed__9; +x_3 = l_Lean_unbracketedExplicitBinders___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); @@ -16153,49 +16301,63 @@ return x_45; } } } -static lean_object* _init_l_Lean_doElemRepeat__Until_____closed__1() { +static lean_object* _init_l_Lean_doElemRepeat____Until_____closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("doElemRepeat_Until_", 19); +x_1 = lean_mk_string_from_bytes("doElemRepeat__Until_", 20); return x_1; } } -static lean_object* _init_l_Lean_doElemRepeat__Until_____closed__2() { +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__1; -x_2 = l_Lean_doElemRepeat__Until_____closed__1; +x_2 = l_Lean_doElemRepeat____Until_____closed__1; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_doElemRepeat__Until_____closed__3() { +static lean_object* _init_l_Lean_doElemRepeat____Until_____closed__3() { +_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__5; +x_2 = l_Lean_doElemRepeat_____closed__8; +x_3 = l_solve___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_doElemRepeat____Until_____closed__4() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes(" until ", 7); +x_1 = lean_mk_string_from_bytes("until ", 6); return x_1; } } -static lean_object* _init_l_Lean_doElemRepeat__Until_____closed__4() { +static lean_object* _init_l_Lean_doElemRepeat____Until_____closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_doElemRepeat__Until_____closed__3; +x_1 = l_Lean_doElemRepeat____Until_____closed__4; 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_doElemRepeat__Until_____closed__5() { +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__5; -x_2 = l_Lean_doElemRepeat_____closed__8; -x_3 = l_Lean_doElemRepeat__Until_____closed__4; +x_2 = l_Lean_doElemRepeat____Until_____closed__3; +x_3 = l_Lean_doElemRepeat____Until_____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); @@ -16203,12 +16365,12 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_doElemRepeat__Until_____closed__6() { +static lean_object* _init_l_Lean_doElemRepeat____Until_____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__5; -x_2 = l_Lean_doElemRepeat__Until_____closed__5; +x_2 = l_Lean_doElemRepeat____Until_____closed__6; x_3 = l_Lean_termMacro_x2etrace_x5b___x5d_____closed__19; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -16217,13 +16379,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_doElemRepeat__Until_____closed__7() { +static lean_object* _init_l_Lean_doElemRepeat____Until_____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_doElemRepeat__Until_____closed__2; +x_1 = l_Lean_doElemRepeat____Until_____closed__2; x_2 = lean_unsigned_to_nat(1022u); -x_3 = l_Lean_doElemRepeat__Until_____closed__6; +x_3 = l_Lean_doElemRepeat____Until_____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); @@ -16231,15 +16393,15 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_doElemRepeat__Until__() { +static lean_object* _init_l_Lean_doElemRepeat____Until__() { _start: { lean_object* x_1; -x_1 = l_Lean_doElemRepeat__Until_____closed__7; +x_1 = l_Lean_doElemRepeat____Until_____closed__8; return x_1; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat__Until____1___closed__1() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____Until____1___closed__1() { _start: { lean_object* x_1; @@ -16247,23 +16409,23 @@ x_1 = lean_mk_string_from_bytes("doNested", 8); return x_1; } } -static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat__Until____1___closed__2() { +static lean_object* _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____Until____1___closed__2() { _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_termMacro_x2etrace_x5b___x5d_____closed__1; x_2 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__1; x_3 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__termMacro_x2etrace_x5b___x5d____1___closed__2; -x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat__Until____1___closed__1; +x_4 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____Until____1___closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -LEAN_EXPORT lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat__Until____1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____Until____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_doElemRepeat__Until_____closed__2; +x_4 = l_Lean_doElemRepeat____Until_____closed__2; lean_inc(x_1); x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); if (x_5 == 0) @@ -16300,7 +16462,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__doElemRepeat__Until____1___closed__2; +x_19 = l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____Until____1___closed__2; lean_inc(x_14); x_20 = l_Lean_Syntax_node2(x_14, x_19, x_18, x_9); x_21 = l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__10; @@ -16420,10 +16582,28 @@ return x_3; static lean_object* _init_l_Lean_term__Matches___x7c___closed__6() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(" | ", 3); +return x_1; +} +} +static lean_object* _init_l_Lean_term__Matches___x7c___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_term__Matches___x7c___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_term__Matches___x7c___closed__8() { +_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_term__Matches___x7c___closed__5; -x_2 = l___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___closed__16; -x_3 = l_solve___closed__5; +x_2 = l_Lean_term__Matches___x7c___closed__6; +x_3 = l_Lean_term__Matches___x7c___closed__7; x_4 = 0; x_5 = lean_alloc_ctor(11, 3, 1); lean_ctor_set(x_5, 0, x_1); @@ -16433,13 +16613,13 @@ lean_ctor_set_uint8(x_5, sizeof(void*)*3, x_4); return x_5; } } -static lean_object* _init_l_Lean_term__Matches___x7c___closed__7() { +static lean_object* _init_l_Lean_term__Matches___x7c___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__5; x_2 = l_Lean_term__Matches___x7c___closed__4; -x_3 = l_Lean_term__Matches___x7c___closed__6; +x_3 = l_Lean_term__Matches___x7c___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); @@ -16447,14 +16627,14 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_term__Matches___x7c___closed__8() { +static lean_object* _init_l_Lean_term__Matches___x7c___closed__10() { _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_term__Matches___x7c___closed__2; x_2 = lean_unsigned_to_nat(50u); x_3 = lean_unsigned_to_nat(51u); -x_4 = l_Lean_term__Matches___x7c___closed__7; +x_4 = l_Lean_term__Matches___x7c___closed__9; x_5 = lean_alloc_ctor(4, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -16467,7 +16647,7 @@ static lean_object* _init_l_Lean_term__Matches___x7c() { _start: { lean_object* x_1; -x_1 = l_Lean_term__Matches___x7c___closed__8; +x_1 = l_Lean_term__Matches___x7c___closed__10; return x_1; } } @@ -17132,6 +17312,14 @@ l_Lean_unbracketedExplicitBinders___closed__12 = _init_l_Lean_unbracketedExplici lean_mark_persistent(l_Lean_unbracketedExplicitBinders___closed__12); l_Lean_unbracketedExplicitBinders___closed__13 = _init_l_Lean_unbracketedExplicitBinders___closed__13(); lean_mark_persistent(l_Lean_unbracketedExplicitBinders___closed__13); +l_Lean_unbracketedExplicitBinders___closed__14 = _init_l_Lean_unbracketedExplicitBinders___closed__14(); +lean_mark_persistent(l_Lean_unbracketedExplicitBinders___closed__14); +l_Lean_unbracketedExplicitBinders___closed__15 = _init_l_Lean_unbracketedExplicitBinders___closed__15(); +lean_mark_persistent(l_Lean_unbracketedExplicitBinders___closed__15); +l_Lean_unbracketedExplicitBinders___closed__16 = _init_l_Lean_unbracketedExplicitBinders___closed__16(); +lean_mark_persistent(l_Lean_unbracketedExplicitBinders___closed__16); +l_Lean_unbracketedExplicitBinders___closed__17 = _init_l_Lean_unbracketedExplicitBinders___closed__17(); +lean_mark_persistent(l_Lean_unbracketedExplicitBinders___closed__17); l_Lean_unbracketedExplicitBinders = _init_l_Lean_unbracketedExplicitBinders(); lean_mark_persistent(l_Lean_unbracketedExplicitBinders); l_Lean_bracketedExplicitBinders___closed__1 = _init_l_Lean_bracketedExplicitBinders___closed__1(); @@ -17158,6 +17346,14 @@ l_Lean_bracketedExplicitBinders___closed__11 = _init_l_Lean_bracketedExplicitBin lean_mark_persistent(l_Lean_bracketedExplicitBinders___closed__11); l_Lean_bracketedExplicitBinders___closed__12 = _init_l_Lean_bracketedExplicitBinders___closed__12(); lean_mark_persistent(l_Lean_bracketedExplicitBinders___closed__12); +l_Lean_bracketedExplicitBinders___closed__13 = _init_l_Lean_bracketedExplicitBinders___closed__13(); +lean_mark_persistent(l_Lean_bracketedExplicitBinders___closed__13); +l_Lean_bracketedExplicitBinders___closed__14 = _init_l_Lean_bracketedExplicitBinders___closed__14(); +lean_mark_persistent(l_Lean_bracketedExplicitBinders___closed__14); +l_Lean_bracketedExplicitBinders___closed__15 = _init_l_Lean_bracketedExplicitBinders___closed__15(); +lean_mark_persistent(l_Lean_bracketedExplicitBinders___closed__15); +l_Lean_bracketedExplicitBinders___closed__16 = _init_l_Lean_bracketedExplicitBinders___closed__16(); +lean_mark_persistent(l_Lean_bracketedExplicitBinders___closed__16); l_Lean_bracketedExplicitBinders = _init_l_Lean_bracketedExplicitBinders(); lean_mark_persistent(l_Lean_bracketedExplicitBinders); l_Lean_explicitBinders___closed__1 = _init_l_Lean_explicitBinders___closed__1(); @@ -17174,6 +17370,8 @@ 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___closed__8 = _init_l_Lean_explicitBinders___closed__8(); +lean_mark_persistent(l_Lean_explicitBinders___closed__8); 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(); @@ -17264,194 +17462,198 @@ 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_____closed__34 = _init_l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__34(); -lean_mark_persistent(l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__34); -l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__35 = _init_l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__35(); -lean_mark_persistent(l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__35); -l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__36 = _init_l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__36(); -lean_mark_persistent(l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__36); -l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__37 = _init_l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__37(); -lean_mark_persistent(l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__37); -l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__38 = _init_l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__38(); -lean_mark_persistent(l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__38); -l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__39 = _init_l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__39(); -lean_mark_persistent(l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__39); -l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__40 = _init_l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__40(); -lean_mark_persistent(l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__40); -l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__41 = _init_l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__41(); -lean_mark_persistent(l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__41); -l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__42 = _init_l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__42(); -lean_mark_persistent(l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__42); -l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__43 = _init_l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__43(); -lean_mark_persistent(l_Lean_command____Unif__hint____Where___x7c_x2d_u22a2_____closed__43); -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_Array_forInUnsafe_loop___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___spec__5___closed__4 = _init_l_Array_forInUnsafe_loop___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___spec__5___closed__4(); -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__4); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__1 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__1(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__1); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__2 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__2(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__2); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__3 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__3(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__3); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__4 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__4(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__4); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__5 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__5(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__5); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__6 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__6(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__6); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__7 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__7(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__7); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__8 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__8(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__8); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__9 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__9(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__9); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__10 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__10(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__10); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__11 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__11(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__11); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__12 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__12(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__12); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__13 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__13(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__13); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__14 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__14(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__14); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__15 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__15(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__15); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__16 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__16(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__16); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__17 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__17(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__17); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__18 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__18(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__18); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__19 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__19(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__19); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__20 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__20(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__20); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__21 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__21(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__21); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__22 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__22(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__22); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__23 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__23(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__23); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__24 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__24(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__24); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__25 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__25(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__25); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__26 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__26(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__26); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__27 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__27(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__27); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__28 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__28(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__28); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__29 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__29(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__29); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__30 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__30(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__30); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__31 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__31(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__31); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__32 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__32(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__32); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__33 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__33(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__33); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__34 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__34(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__34); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__35 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__35(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__35); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__36 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__36(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__36); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__37 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__37(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__37); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__38 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__38(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__38); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__39 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__39(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__39); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__40 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__40(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__40); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__41 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__41(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__41); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__42 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__42(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__42); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__43 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__43(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__43); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__44 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__44(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__44); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__45 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__45(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint____Where___x7c_x2d_u22a2____1___lambda__2___closed__45); -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_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_____closed__34 = _init_l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__34(); +lean_mark_persistent(l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__34); +l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__35 = _init_l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__35(); +lean_mark_persistent(l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__35); +l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__36 = _init_l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__36(); +lean_mark_persistent(l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__36); +l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__37 = _init_l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__37(); +lean_mark_persistent(l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__37); +l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__38 = _init_l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__38(); +lean_mark_persistent(l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__38); +l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__39 = _init_l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__39(); +lean_mark_persistent(l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__39); +l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__40 = _init_l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__40(); +lean_mark_persistent(l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__40); +l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__41 = _init_l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__41(); +lean_mark_persistent(l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__41); +l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__42 = _init_l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__42(); +lean_mark_persistent(l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__42); +l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__43 = _init_l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__43(); +lean_mark_persistent(l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__43); +l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__44 = _init_l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__44(); +lean_mark_persistent(l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__44); +l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__45 = _init_l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__45(); +lean_mark_persistent(l_Lean_command____Unif__hint________Where___x7c_x2d_u22a2_____closed__45); +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_Array_forInUnsafe_loop___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___spec__5___closed__4 = _init_l_Array_forInUnsafe_loop___at_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___spec__5___closed__4(); +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__4); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__1 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__1(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__1); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__2 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__2(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__2); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__3 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__3(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__3); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__4 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__4(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__4); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__5 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__5(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__5); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__6 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__6(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__6); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__7 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__7(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__7); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__8 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__8(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__8); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__9 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__9(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__9); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__10 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__10(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__10); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__11 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__11(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__11); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__12 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__12(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__12); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__13 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__13(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__13); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__14 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__14(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__14); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__15 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__15(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__15); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__16 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__16(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__16); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__17 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__17(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__17); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__18 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__18(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__18); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__19 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__19(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__19); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__20 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__20(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__20); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__21 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__21(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__21); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__22 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__22(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__22); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__23 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__23(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__23); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__24 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__24(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__24); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__25 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__25(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__25); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__26 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__26(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__26); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__27 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__27(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__27); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__28 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__28(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__28); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__29 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__29(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__29); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__30 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__30(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__30); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__31 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__31(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__31); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__32 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__32(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__32); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__33 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__33(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__33); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__34 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__34(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__34); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__35 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__35(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__35); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__36 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__36(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__36); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__37 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__37(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__37); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__38 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__38(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__38); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__39 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__39(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__39); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__40 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__40(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__40); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__41 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__41(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__41); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__42 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__42(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__42); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__43 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__43(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__43); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__44 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__44(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__44); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__45 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__45(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__command____Unif__hint________Where___x7c_x2d_u22a2____1___lambda__2___closed__45); +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_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(); @@ -17634,8 +17836,6 @@ l_calcSteps___closed__11 = _init_l_calcSteps___closed__11(); lean_mark_persistent(l_calcSteps___closed__11); l_calcSteps___closed__12 = _init_l_calcSteps___closed__12(); lean_mark_persistent(l_calcSteps___closed__12); -l_calcSteps___closed__13 = _init_l_calcSteps___closed__13(); -lean_mark_persistent(l_calcSteps___closed__13); l_calcSteps = _init_l_calcSteps(); lean_mark_persistent(l_calcSteps); l_calc___closed__1 = _init_l_calc___closed__1(); @@ -17708,8 +17908,6 @@ l_unexpandExists___closed__3 = _init_l_unexpandExists___closed__3(); lean_mark_persistent(l_unexpandExists___closed__3); 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_unexpandSigma___closed__1 = _init_l_unexpandSigma___closed__1(); lean_mark_persistent(l_unexpandSigma___closed__1); l_unexpandPSigma___closed__1 = _init_l_unexpandPSigma___closed__1(); @@ -17784,12 +17982,6 @@ l_tacticFunext_________closed__12 = _init_l_tacticFunext_________closed__12(); lean_mark_persistent(l_tacticFunext_________closed__12); l_tacticFunext_________closed__13 = _init_l_tacticFunext_________closed__13(); lean_mark_persistent(l_tacticFunext_________closed__13); -l_tacticFunext_________closed__14 = _init_l_tacticFunext_________closed__14(); -lean_mark_persistent(l_tacticFunext_________closed__14); -l_tacticFunext_________closed__15 = _init_l_tacticFunext_________closed__15(); -lean_mark_persistent(l_tacticFunext_________closed__15); -l_tacticFunext_________closed__16 = _init_l_tacticFunext_________closed__16(); -lean_mark_persistent(l_tacticFunext_________closed__16); l_tacticFunext______ = _init_l_tacticFunext______(); lean_mark_persistent(l_tacticFunext______); l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__1 = _init_l___aux__Init__NotationExtra______macroRules__tacticFunext________1___lambda__2___closed__1(); @@ -17942,6 +18134,8 @@ l_Lean_Parser_Command_classAbbrev___closed__28 = _init_l_Lean_Parser_Command_cla lean_mark_persistent(l_Lean_Parser_Command_classAbbrev___closed__28); l_Lean_Parser_Command_classAbbrev___closed__29 = _init_l_Lean_Parser_Command_classAbbrev___closed__29(); lean_mark_persistent(l_Lean_Parser_Command_classAbbrev___closed__29); +l_Lean_Parser_Command_classAbbrev___closed__30 = _init_l_Lean_Parser_Command_classAbbrev___closed__30(); +lean_mark_persistent(l_Lean_Parser_Command_classAbbrev___closed__30); l_Lean_Parser_Command_classAbbrev = _init_l_Lean_Parser_Command_classAbbrev(); lean_mark_persistent(l_Lean_Parser_Command_classAbbrev); l___aux__Init__NotationExtra______macroRules__Lean__Parser__Command__classAbbrev__1___lambda__3___closed__1 = _init_l___aux__Init__NotationExtra______macroRules__Lean__Parser__Command__classAbbrev__1___lambda__3___closed__1(); @@ -18048,6 +18242,14 @@ l_solve___closed__13 = _init_l_solve___closed__13(); lean_mark_persistent(l_solve___closed__13); l_solve___closed__14 = _init_l_solve___closed__14(); lean_mark_persistent(l_solve___closed__14); +l_solve___closed__15 = _init_l_solve___closed__15(); +lean_mark_persistent(l_solve___closed__15); +l_solve___closed__16 = _init_l_solve___closed__16(); +lean_mark_persistent(l_solve___closed__16); +l_solve___closed__17 = _init_l_solve___closed__17(); +lean_mark_persistent(l_solve___closed__17); +l_solve___closed__18 = _init_l_solve___closed__18(); +lean_mark_persistent(l_solve___closed__18); l_solve = _init_l_solve(); lean_mark_persistent(l_solve); 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(); @@ -18186,26 +18388,28 @@ l_Lean_doElemWhile__Do_____closed__6 = _init_l_Lean_doElemWhile__Do_____closed__ lean_mark_persistent(l_Lean_doElemWhile__Do_____closed__6); l_Lean_doElemWhile__Do__ = _init_l_Lean_doElemWhile__Do__(); lean_mark_persistent(l_Lean_doElemWhile__Do__); -l_Lean_doElemRepeat__Until_____closed__1 = _init_l_Lean_doElemRepeat__Until_____closed__1(); -lean_mark_persistent(l_Lean_doElemRepeat__Until_____closed__1); -l_Lean_doElemRepeat__Until_____closed__2 = _init_l_Lean_doElemRepeat__Until_____closed__2(); -lean_mark_persistent(l_Lean_doElemRepeat__Until_____closed__2); -l_Lean_doElemRepeat__Until_____closed__3 = _init_l_Lean_doElemRepeat__Until_____closed__3(); -lean_mark_persistent(l_Lean_doElemRepeat__Until_____closed__3); -l_Lean_doElemRepeat__Until_____closed__4 = _init_l_Lean_doElemRepeat__Until_____closed__4(); -lean_mark_persistent(l_Lean_doElemRepeat__Until_____closed__4); -l_Lean_doElemRepeat__Until_____closed__5 = _init_l_Lean_doElemRepeat__Until_____closed__5(); -lean_mark_persistent(l_Lean_doElemRepeat__Until_____closed__5); -l_Lean_doElemRepeat__Until_____closed__6 = _init_l_Lean_doElemRepeat__Until_____closed__6(); -lean_mark_persistent(l_Lean_doElemRepeat__Until_____closed__6); -l_Lean_doElemRepeat__Until_____closed__7 = _init_l_Lean_doElemRepeat__Until_____closed__7(); -lean_mark_persistent(l_Lean_doElemRepeat__Until_____closed__7); -l_Lean_doElemRepeat__Until__ = _init_l_Lean_doElemRepeat__Until__(); -lean_mark_persistent(l_Lean_doElemRepeat__Until__); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat__Until____1___closed__1 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat__Until____1___closed__1(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat__Until____1___closed__1); -l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat__Until____1___closed__2 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat__Until____1___closed__2(); -lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat__Until____1___closed__2); +l_Lean_doElemRepeat____Until_____closed__1 = _init_l_Lean_doElemRepeat____Until_____closed__1(); +lean_mark_persistent(l_Lean_doElemRepeat____Until_____closed__1); +l_Lean_doElemRepeat____Until_____closed__2 = _init_l_Lean_doElemRepeat____Until_____closed__2(); +lean_mark_persistent(l_Lean_doElemRepeat____Until_____closed__2); +l_Lean_doElemRepeat____Until_____closed__3 = _init_l_Lean_doElemRepeat____Until_____closed__3(); +lean_mark_persistent(l_Lean_doElemRepeat____Until_____closed__3); +l_Lean_doElemRepeat____Until_____closed__4 = _init_l_Lean_doElemRepeat____Until_____closed__4(); +lean_mark_persistent(l_Lean_doElemRepeat____Until_____closed__4); +l_Lean_doElemRepeat____Until_____closed__5 = _init_l_Lean_doElemRepeat____Until_____closed__5(); +lean_mark_persistent(l_Lean_doElemRepeat____Until_____closed__5); +l_Lean_doElemRepeat____Until_____closed__6 = _init_l_Lean_doElemRepeat____Until_____closed__6(); +lean_mark_persistent(l_Lean_doElemRepeat____Until_____closed__6); +l_Lean_doElemRepeat____Until_____closed__7 = _init_l_Lean_doElemRepeat____Until_____closed__7(); +lean_mark_persistent(l_Lean_doElemRepeat____Until_____closed__7); +l_Lean_doElemRepeat____Until_____closed__8 = _init_l_Lean_doElemRepeat____Until_____closed__8(); +lean_mark_persistent(l_Lean_doElemRepeat____Until_____closed__8); +l_Lean_doElemRepeat____Until__ = _init_l_Lean_doElemRepeat____Until__(); +lean_mark_persistent(l_Lean_doElemRepeat____Until__); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____Until____1___closed__1 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____Until____1___closed__1(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____Until____1___closed__1); +l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____Until____1___closed__2 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____Until____1___closed__2(); +lean_mark_persistent(l_Lean___aux__Init__NotationExtra______macroRules__Lean__doElemRepeat____Until____1___closed__2); l_Lean_term__Matches___x7c___closed__1 = _init_l_Lean_term__Matches___x7c___closed__1(); lean_mark_persistent(l_Lean_term__Matches___x7c___closed__1); l_Lean_term__Matches___x7c___closed__2 = _init_l_Lean_term__Matches___x7c___closed__2(); @@ -18222,6 +18426,10 @@ l_Lean_term__Matches___x7c___closed__7 = _init_l_Lean_term__Matches___x7c___clos lean_mark_persistent(l_Lean_term__Matches___x7c___closed__7); l_Lean_term__Matches___x7c___closed__8 = _init_l_Lean_term__Matches___x7c___closed__8(); lean_mark_persistent(l_Lean_term__Matches___x7c___closed__8); +l_Lean_term__Matches___x7c___closed__9 = _init_l_Lean_term__Matches___x7c___closed__9(); +lean_mark_persistent(l_Lean_term__Matches___x7c___closed__9); +l_Lean_term__Matches___x7c___closed__10 = _init_l_Lean_term__Matches___x7c___closed__10(); +lean_mark_persistent(l_Lean_term__Matches___x7c___closed__10); l_Lean_term__Matches___x7c = _init_l_Lean_term__Matches___x7c(); lean_mark_persistent(l_Lean_term__Matches___x7c); l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__1 = _init_l_Lean___aux__Init__NotationExtra______macroRules__Lean__term__Matches___x7c__1___closed__1(); diff --git a/stage0/stdlib/Init/Prelude.c b/stage0/stdlib/Init/Prelude.c index 833a425ec446..90c3fc8aae1c 100644 --- a/stage0/stdlib/Init/Prelude.c +++ b/stage0/stdlib/Init/Prelude.c @@ -124,6 +124,7 @@ LEAN_EXPORT lean_object* l_not___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_getId(lean_object*); LEAN_EXPORT lean_object* l_ReaderT_instApplicativeReaderT___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_instDecidableNot___rarg___boxed(lean_object*); +static lean_object* l_Lean_hygieneInfoKind___closed__1; LEAN_EXPORT lean_object* l_Lean_TSyntaxArray_rawImpl___rarg___boxed(lean_object*); LEAN_EXPORT lean_object* l_instHAppend(lean_object*); LEAN_EXPORT lean_object* l_instMonadWithReaderOfReaderT___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -679,6 +680,7 @@ LEAN_EXPORT lean_object* l_instHMod___rarg(lean_object*, lean_object*, lean_obje LEAN_EXPORT lean_object* l_Fin_decLt___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Macro_instMonadQuotationMacroM___lambda__2___boxed(lean_object*, lean_object*); static lean_object* l_instInhabitedSubstring___closed__1; +static lean_object* l_Lean_hygieneInfoKind___closed__2; LEAN_EXPORT lean_object* l_Lean_Syntax_node7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Macro_instMonadRefMacroM___closed__1; LEAN_EXPORT lean_object* l_Lean_Syntax_getTailPos_x3f_loop___boxed(lean_object*, lean_object*, lean_object*); @@ -855,6 +857,7 @@ LEAN_EXPORT lean_object* l_Array_data___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mkArray2___rarg(lean_object*, lean_object*); LEAN_EXPORT uint32_t l_Char_utf8Size(uint32_t); static lean_object* l_Lean_Macro_instInhabitedMethods___closed__3; +LEAN_EXPORT lean_object* l_Lean_hygieneInfoKind; LEAN_EXPORT lean_object* l_ReaderT_instMonadExceptOfReaderT(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_USize_ofNat32___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_PrettyPrinter_instMonadQuotationUnexpandM___spec__2___rarg(lean_object*, lean_object*, lean_object*); @@ -8405,6 +8408,32 @@ x_1 = l_Lean_fieldIdxKind___closed__2; return x_1; } } +static lean_object* _init_l_Lean_hygieneInfoKind___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("hygieneInfo", 11); +return x_1; +} +} +static lean_object* _init_l_Lean_hygieneInfoKind___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_hygieneInfoKind___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_hygieneInfoKind() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_hygieneInfoKind___closed__2; +return x_1; +} +} static lean_object* _init_l_Lean_interpolatedStrLitKind___closed__1() { _start: { @@ -12339,6 +12368,12 @@ l_Lean_fieldIdxKind___closed__2 = _init_l_Lean_fieldIdxKind___closed__2(); lean_mark_persistent(l_Lean_fieldIdxKind___closed__2); l_Lean_fieldIdxKind = _init_l_Lean_fieldIdxKind(); lean_mark_persistent(l_Lean_fieldIdxKind); +l_Lean_hygieneInfoKind___closed__1 = _init_l_Lean_hygieneInfoKind___closed__1(); +lean_mark_persistent(l_Lean_hygieneInfoKind___closed__1); +l_Lean_hygieneInfoKind___closed__2 = _init_l_Lean_hygieneInfoKind___closed__2(); +lean_mark_persistent(l_Lean_hygieneInfoKind___closed__2); +l_Lean_hygieneInfoKind = _init_l_Lean_hygieneInfoKind(); +lean_mark_persistent(l_Lean_hygieneInfoKind); l_Lean_interpolatedStrLitKind___closed__1 = _init_l_Lean_interpolatedStrLitKind___closed__1(); lean_mark_persistent(l_Lean_interpolatedStrLitKind___closed__1); l_Lean_interpolatedStrLitKind___closed__2 = _init_l_Lean_interpolatedStrLitKind___closed__2(); diff --git a/stage0/stdlib/Init/Tactics.c b/stage0/stdlib/Init/Tactics.c index 0aabc13db2f5..d9ba63458c56 100644 --- a/stage0/stdlib/Init/Tactics.c +++ b/stage0/stdlib/Init/Tactics.c @@ -45,6 +45,7 @@ static lean_object* l_Lean_Parser_Tactic_injection___closed__7; static lean_object* l_term_u2039___u203a___closed__5; static lean_object* l_term_u2039___u203a___closed__9; 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__27; static lean_object* l_Lean_Parser_Tactic_inductionAlt___closed__10; lean_object* lean_mk_empty_array_with_capacity(lean_object*); static lean_object* l_Lean_Parser_Tactic_inductionAltLHS___closed__10; @@ -140,6 +141,7 @@ static lean_object* l_Lean_Parser_Tactic_traceState___closed__2; static lean_object* l_Lean_Parser_Tactic_letrec___closed__10; static lean_object* l_Lean_Parser_Tactic_specialize___closed__2; static lean_object* l_Lean_Parser_Tactic_tacticSuffices_____closed__8; +static lean_object* l_Lean_Parser_Tactic_locationHyp___closed__19; static lean_object* l_Lean_Parser_Tactic_rewriteSeq___closed__5; static lean_object* l_Lean_Parser_Tactic_case___closed__1; static lean_object* l_Lean_Parser_Tactic_withAnnotateState___closed__7; @@ -197,7 +199,6 @@ static lean_object* l_Lean_Parser_Tactic_simpAll___closed__3; static lean_object* l_Lean_Parser_Tactic_dsimp___closed__9; static lean_object* l_Lean_Parser_Tactic_tacticUnhygienic_____closed__2; static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticRfl_x27__1___closed__3; -static lean_object* l_Lean_Parser_Tactic_inductionAlts___closed__11; static lean_object* l_Lean_Parser_Tactic_inductionAlt___closed__9; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tactic___x3c_x3b_x3e____1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_first___closed__17; @@ -245,7 +246,6 @@ static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__L LEAN_EXPORT lean_object* l_tacticGet__elem__tactic; static lean_object* l_Lean_Parser_Tactic_paren___closed__1; static lean_object* l_Lean_Parser_Tactic_discharger___closed__7; -static lean_object* l_Lean_Parser_Tactic_unfold___closed__7; static lean_object* l_Lean_Parser_Tactic_withAnnotateState___closed__4; static lean_object* l_Lean_Parser_Tactic_intro___closed__10; static lean_object* l_Lean_Parser_Tactic_inductionAltLHS___closed__1; @@ -332,9 +332,7 @@ static lean_object* l_Lean_Parser_Tactic_rewriteSeq___closed__9; static lean_object* l_Lean_Parser_Tactic_tacticLet_____closed__6; static lean_object* l_Lean_Parser_Tactic_delta___closed__7; LEAN_EXPORT lean_object* l___aux__Init__Tactics______macroRules__tacticGet__elem__tactic__trivial__1(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic_inductionAltLHS___closed__12; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_tacticInfer__instance; -static lean_object* l_Lean_Parser_Tactic_inductionAltLHS___closed__13; static lean_object* l_term_____x5b___x5d___closed__7; static lean_object* l_Lean_Parser_Tactic_congr___closed__2; static lean_object* l_Lean_Parser_Tactic_anyGoals___closed__6; @@ -348,7 +346,6 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_cases; static lean_object* l_Lean_Parser_Tactic_case___closed__12; static lean_object* l_Lean_Parser_Tactic_cases___closed__8; static lean_object* l_Lean_Parser_Tactic_simpPre___closed__4; -static lean_object* l_Lean_Parser_Tactic_subst___closed__6; static lean_object* l_Lean_Parser_Tactic_rwRule___closed__10; static lean_object* l_Lean_Parser_Tactic_cases___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTrivial__2(lean_object*, lean_object*, lean_object*); @@ -387,7 +384,6 @@ static lean_object* l___aux__Init__Tactics______macroRules__tacticGet__elem__tac static lean_object* l_Lean_Parser_Tactic_renameI___closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_generalizeArg; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_withReducible; -static lean_object* l_Lean_Parser_Tactic_revert___closed__9; static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticSorry__1___closed__5; static lean_object* l_Lean_Parser_Tactic_traceMessage___closed__8; static lean_object* l_term_____x5b___x5d___closed__1; @@ -430,10 +426,8 @@ lean_object* l_Lean_SourceInfo_fromRef(lean_object*, uint8_t); static lean_object* l_Lean_Parser_Tactic_cases___closed__9; static lean_object* l_Lean_Parser_Tactic_induction___closed__17; static lean_object* l_Lean_Parser_Tactic_done___closed__2; -static lean_object* l_Lean_Parser_Tactic_inductionAlt___closed__17; static lean_object* l_Lean_Parser_Tactic_failIfSuccess___closed__2; static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__10; -static lean_object* l_Lean_Parser_Tactic_delta___closed__9; static lean_object* l_Lean_Parser_Tactic_paren___closed__8; static lean_object* l_Lean_Parser_Tactic_rewriteSeq___closed__6; static lean_object* l_Lean_Parser_Tactic_tacticTry_____closed__2; @@ -456,7 +450,6 @@ static lean_object* l_Lean_Parser_Tactic_induction___closed__16; static lean_object* l_Lean_Parser_Tactic_sleep___closed__3; static lean_object* l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__5; static lean_object* l_Lean_Parser_Tactic_simpPost___closed__5; -static lean_object* l_Lean_Parser_Tactic_intros___closed__17; static lean_object* l_Lean_Parser_Tactic_inductionAltLHS___closed__6; static lean_object* l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__8; static lean_object* l_Lean_Parser_Tactic_tacticTrivial___closed__1; @@ -515,7 +508,6 @@ static lean_object* l_Lean_Parser_Tactic_tacticUnhygienic_____closed__5; static lean_object* l_Lean_Parser_Tactic_tacticHave_x27_____closed__6; 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___aux__Init__Tactics______macroRules__term_____x5b___x5d_x27____1(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic_congr___closed__6; static lean_object* l_Lean_Parser_Tactic_rewriteSeq___closed__7; static lean_object* l_Lean_Parser_Tactic_rwSeq___closed__8; static lean_object* l_Lean_Parser_Tactic_tacticUnhygienic_____closed__1; @@ -538,6 +530,7 @@ static lean_object* l_Lean_Parser_Tactic_clear___closed__3; static lean_object* l_Lean_Parser_Tactic_refine___closed__6; static lean_object* l_Lean_Parser_Tactic_intro___closed__17; static lean_object* l_Lean_Parser_Tactic_rwRule___closed__3; +static lean_object* l_Lean_Parser_Tactic_first___closed__23; static lean_object* l_Lean_Parser_Tactic_tacticHave_____closed__9; static lean_object* l_Lean_Parser_Tactic_split___closed__7; static lean_object* l_Lean_Parser_Tactic_rotateLeft___closed__11; @@ -582,12 +575,14 @@ static lean_object* l_Lean_Parser_Tactic_skip___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTrivial__4(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_tactic___x3c_x3b_x3e_____closed__1; static lean_object* l_Lean_Parser_Tactic_inductionAltLHS___closed__7; +static lean_object* l_Lean_Parser_Tactic_injection___closed__11; static lean_object* l_Lean_Parser_Tactic_clear___closed__4; static lean_object* l_Lean_Parser_Tactic_apply___closed__1; 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_tacticNext___x3d_x3e_____closed__9; static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticShow____1___closed__4; static lean_object* l_Lean_Parser_Tactic_focus___closed__5; +static lean_object* l_Lean_Parser_Tactic_simpStar___closed__4; static lean_object* l_Lean_Parser_Tactic_renameI___closed__5; static lean_object* l_term_u2039___u203a___closed__2; static lean_object* l_Lean_Parser_Tactic_rotateLeft___closed__4; @@ -619,12 +614,14 @@ static lean_object* l_Lean_Parser_Tactic_simp___closed__14; static lean_object* l_Lean_Parser_Tactic_refine___closed__1; static lean_object* l_Lean_Parser_Tactic_simp___closed__3; static lean_object* l_Lean_Parser_Tactic_case___closed__9; +static lean_object* l_Lean_Parser_Tactic_first___closed__26; static lean_object* l_Lean_Parser_Tactic_induction___closed__9; lean_object* l_Lean_Syntax_node2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_tacticRefine__lift_x27_____closed__5; static lean_object* l_Lean_Parser_Tactic_locationHyp___closed__14; static lean_object* l_Lean_Parser_Tactic_apply___closed__2; static lean_object* l_Lean_Parser_Tactic_location___closed__2; +static lean_object* l_Lean_Parser_Tactic_first___closed__24; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_withUnfoldingAll; static lean_object* l_Lean_Parser_Tactic_tacticRfl_x27___closed__1; static lean_object* l_Lean_Parser_Tactic_rotateLeft___closed__9; @@ -706,8 +703,8 @@ static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__L LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_tacticTry__; static lean_object* l_Lean_Parser_Tactic_acRfl___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_simpPre; +static lean_object* l_Lean_Parser_Tactic_first___closed__21; static lean_object* l_Lean_Parser_Tactic_unfold___closed__6; -static lean_object* l_Lean_Parser_Tactic_inductionAlt___closed__13; static lean_object* l_Lean_Parser_Tactic_rotateRight___closed__4; static lean_object* l_Lean_Parser_Tactic_traceState___closed__5; static lean_object* l_Lean_Parser_Tactic_simp___closed__1; @@ -759,6 +756,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_tacticRfl_x27; static lean_object* l_term_____x5b___x5d___closed__9; static lean_object* l_Lean_Parser_Tactic_simpAll___closed__5; static lean_object* l_Lean_Parser_Tactic_rwSeq___closed__2; +static lean_object* l_Lean_Parser_Tactic_first___closed__22; static lean_object* l_Lean_Parser_Tactic_specialize___closed__5; static lean_object* l_Lean_Parser_Tactic_generalizeArg___closed__12; static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__letrec__1___closed__2; @@ -791,14 +789,15 @@ static lean_object* l_term_____x5b___x5d_x27_____closed__5; static lean_object* l_Lean_Parser_Tactic_induction___closed__3; static lean_object* l_Lean_Parser_Tactic_tacticTrivial___closed__2; static lean_object* l_term_u2039___u203a___closed__6; +static lean_object* l_Lean_Parser_Tactic_simpStar___closed__5; static lean_object* l_Lean_Parser_Tactic_simpPost___closed__3; static lean_object* l_Lean_Parser_Tactic_induction___closed__5; static lean_object* l_Lean_Parser_Tactic_intros___closed__5; static lean_object* l_Lean_Parser_Tactic_refine_x27___closed__4; static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticSorry__1___closed__8; -static lean_object* l_Lean_Parser_Tactic_inductionAlt___closed__16; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_tacticExists___x2c_x2c; static lean_object* l_tacticGet__elem__tactic___closed__4; +static lean_object* l_Lean_Parser_Tactic_first___closed__25; static lean_object* l_Lean_Parser_Tactic_intro___closed__13; lean_object* l_Array_appendCore___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_subst___closed__4; @@ -836,7 +835,6 @@ static lean_object* l_Lean_Parser_Tactic_tacticLet_____closed__3; static lean_object* l_Lean_Parser_Tactic_cases___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_injections; static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticHave_x27___x3a_x3d____1___closed__7; -static lean_object* l_Lean_Parser_Tactic_inductionAlt___closed__15; static lean_object* l_Lean_Parser_Tactic_tacticRefine__lift_____closed__6; static lean_object* l_Lean_Parser_Tactic_case_x27___closed__5; static lean_object* l_term_u2039___u203a___closed__1; @@ -874,14 +872,12 @@ static lean_object* l_Lean_Parser_Tactic_generalize___closed__8; static lean_object* l_Lean_Parser_Tactic_unfold___closed__5; static lean_object* l_Lean_Parser_Tactic_withAnnotateState___closed__9; static lean_object* l_Lean_Parser_Tactic_intro___closed__6; -static lean_object* l_Lean_Parser_Tactic_inductionAlt___closed__12; static lean_object* l_Lean_Parser_Tactic_simp___closed__20; static lean_object* l_Lean_Parser_Tactic_locationWildcard___closed__2; 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_simpAll___closed__12; static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__rwSeq__1___closed__4; lean_object* l_Lean_Syntax_node1(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic_split___closed__9; static lean_object* l_Lean_Parser_Tactic_discharger___closed__17; static lean_object* l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__5; static lean_object* l_Lean_Parser_Tactic_generalize___closed__4; @@ -926,6 +922,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_allGoals; static lean_object* l_Lean_Parser_Tactic_withAnnotateState___closed__11; static lean_object* l_Lean_Parser_Tactic_tacticLet_x27_____closed__1; static lean_object* l_Lean_Parser_Tactic_changeWith___closed__2; +static lean_object* l_Lean_Parser_Tactic_sleep___closed__6; static lean_object* l_Lean_Parser_Tactic_dsimp___closed__1; static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticSorry__1___closed__7; static lean_object* l_Lean_Parser_Tactic_intro___closed__18; @@ -940,7 +937,6 @@ static lean_object* l_Lean_Parser_Tactic_paren___closed__6; static lean_object* l_Lean_Parser_Tactic_generalizeArg___closed__5; static lean_object* l_Lean_Parser_Tactic_rwRule___closed__13; static lean_object* l_Lean_Parser_Tactic_intros___closed__10; -static lean_object* l_Lean_Parser_Tactic_inductionAlt___closed__14; extern lean_object* l_Lean_Parser_Tactic_caseArg; LEAN_EXPORT lean_object* l_term_____x5b___x5d_x27__; static lean_object* l_Lean_Parser_Tactic_casesTarget___closed__4; @@ -978,6 +974,7 @@ static lean_object* l_Lean_Parser_Tactic_rwRule___closed__2; static lean_object* l_Lean_Parser_Tactic_subst___closed__1; static lean_object* l_Lean_Parser_Tactic_first___closed__12; static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticSorry__1___closed__13; +static lean_object* l_Lean_Parser_Tactic_first___closed__20; static lean_object* l_Lean_Parser_Tactic_tacticStop_____closed__7; static lean_object* l_Lean_Parser_Tactic_tacticRfl___closed__3; static lean_object* l_Lean_Parser_Tactic_induction___closed__6; @@ -1028,6 +1025,7 @@ static lean_object* l_Lean_Parser_Tactic_simpPre___closed__5; static lean_object* l_Lean_Parser_Tactic_case___closed__10; lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___aux__Init__Tactics______macroRules__tacticGet__elem__tactic__1___closed__1; +static lean_object* l_Lean_Parser_Tactic_rotateLeft___closed__13; static lean_object* l_Lean_Parser_Tactic_intro___closed__8; static lean_object* l_Lean_Parser_Tactic_letrec___closed__4; static lean_object* l___aux__Init__Tactics______macroRules__term_____x5b___x5d__1___closed__3; @@ -1066,9 +1064,11 @@ static lean_object* l_Lean_Parser_Tactic_locationWildcard___closed__4; static lean_object* l___aux__Init__Tactics______macroRules__term_u2039___u203a__1___closed__1; static lean_object* l___aux__Init__Tactics______macroRules__term_u2039___u203a__1___closed__3; static lean_object* l_Lean_Parser_Tactic_locationWildcard___closed__1; +static lean_object* l_Lean_Parser_Tactic_injection___closed__12; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticRepeat____1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_case___closed__4; static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticSorry__1___closed__15; +static lean_object* l_Lean_Parser_Attr_simp___closed__10; static lean_object* l_Lean_Parser_Tactic_tacticLet_____closed__1; static lean_object* l_Lean_Parser_Tactic_inductionAlts___closed__4; static lean_object* l_Lean_Parser_Tactic_subst___closed__5; @@ -1086,7 +1086,6 @@ static lean_object* l_Lean_Parser_Tactic_injection___closed__6; static lean_object* l_Lean_Parser_Tactic_tacticSuffices_____closed__1; static lean_object* l_Lean_Parser_Tactic_locationHyp___closed__6; static lean_object* l_Lean_Parser_Tactic_generalizeArg___closed__1; -static lean_object* l_Lean_Parser_Tactic_clear___closed__6; static lean_object* l_Lean_Parser_Tactic_change___closed__1; static lean_object* l_Lean_Parser_Tactic_revert___closed__6; static lean_object* l_Lean_Parser_Tactic_allGoals___closed__6; @@ -1131,6 +1130,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_failIfSuccess; static lean_object* l_Lean_Parser_Tactic_simpAll___closed__6; static lean_object* l_Lean_Parser_Tactic_assumption___closed__1; static lean_object* l_Lean_Parser_Tactic_inductionAlt___closed__5; +static lean_object* l_Lean_Parser_Tactic_first___closed__19; static lean_object* l_Lean_Parser_Tactic_change___closed__6; static lean_object* l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticSuffices____1___closed__2; static lean_object* l_Lean_Parser_Tactic_induction___closed__2; @@ -1425,16 +1425,8 @@ return x_5; static lean_object* _init_l_Lean_Parser_Tactic_intro___closed__3() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("intro ", 6); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Tactic_intro___closed__4() { -_start: -{ lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_intro___closed__3; +x_1 = l_Lean_Parser_Tactic_intro___closed__1; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -1442,7 +1434,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_intro___closed__5() { +static lean_object* _init_l_Lean_Parser_Tactic_intro___closed__4() { _start: { lean_object* x_1; @@ -1450,17 +1442,17 @@ x_1 = lean_mk_string_from_bytes("notFollowedBy", 13); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_intro___closed__6() { +static lean_object* _init_l_Lean_Parser_Tactic_intro___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic_intro___closed__5; +x_2 = l_Lean_Parser_Tactic_intro___closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_intro___closed__7() { +static lean_object* _init_l_Lean_Parser_Tactic_intro___closed__6() { _start: { lean_object* x_1; @@ -1468,35 +1460,35 @@ x_1 = lean_mk_string_from_bytes("|", 1); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_intro___closed__8() { +static lean_object* _init_l_Lean_Parser_Tactic_intro___closed__7() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic_intro___closed__7; +x_1 = l_Lean_Parser_Tactic_intro___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_Parser_Tactic_intro___closed__9() { +static lean_object* _init_l_Lean_Parser_Tactic_intro___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_intro___closed__6; -x_2 = l_Lean_Parser_Tactic_intro___closed__8; +x_1 = l_Lean_Parser_Tactic_intro___closed__5; +x_2 = l_Lean_Parser_Tactic_intro___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_intro___closed__10() { +static lean_object* _init_l_Lean_Parser_Tactic_intro___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__7; -x_2 = l_Lean_Parser_Tactic_intro___closed__4; -x_3 = l_Lean_Parser_Tactic_intro___closed__9; +x_2 = l_Lean_Parser_Tactic_intro___closed__3; +x_3 = l_Lean_Parser_Tactic_intro___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); @@ -1504,7 +1496,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_intro___closed__11() { +static lean_object* _init_l_Lean_Parser_Tactic_intro___closed__10() { _start: { lean_object* x_1; @@ -1512,17 +1504,17 @@ x_1 = lean_mk_string_from_bytes("many", 4); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_intro___closed__12() { +static lean_object* _init_l_Lean_Parser_Tactic_intro___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_intro___closed__11; +x_2 = l_Lean_Parser_Tactic_intro___closed__10; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_intro___closed__13() { +static lean_object* _init_l_Lean_Parser_Tactic_intro___closed__12() { _start: { lean_object* x_1; @@ -1530,26 +1522,40 @@ x_1 = lean_mk_string_from_bytes("colGt", 5); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_intro___closed__14() { +static lean_object* _init_l_Lean_Parser_Tactic_intro___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_intro___closed__13; +x_2 = l_Lean_Parser_Tactic_intro___closed__12; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_intro___closed__15() { +static lean_object* _init_l_Lean_Parser_Tactic_intro___closed__14() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic_intro___closed__14; +x_1 = l_Lean_Parser_Tactic_intro___closed__13; 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_intro___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__7; +x_2 = l_Lean_Parser_Tactic_withAnnotateState___closed__16; +x_3 = l_Lean_Parser_Tactic_intro___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_Tactic_intro___closed__16() { _start: { @@ -1598,7 +1604,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_intro___closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_intro___closed__12; +x_1 = l_Lean_Parser_Tactic_intro___closed__11; x_2 = l_Lean_Parser_Tactic_intro___closed__19; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -1611,7 +1617,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_intro___closed__21() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__7; -x_2 = l_Lean_Parser_Tactic_intro___closed__10; +x_2 = l_Lean_Parser_Tactic_intro___closed__9; x_3 = l_Lean_Parser_Tactic_intro___closed__20; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -1665,16 +1671,8 @@ return x_5; static lean_object* _init_l_Lean_Parser_Tactic_intros___closed__3() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("intros ", 7); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Tactic_intros___closed__4() { -_start: -{ lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_intros___closed__3; +x_1 = l_Lean_Parser_Tactic_intros___closed__1; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -1682,7 +1680,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_intros___closed__5() { +static lean_object* _init_l_Lean_Parser_Tactic_intros___closed__4() { _start: { lean_object* x_1; @@ -1690,17 +1688,17 @@ x_1 = lean_mk_string_from_bytes("orelse", 6); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_intros___closed__6() { +static lean_object* _init_l_Lean_Parser_Tactic_intros___closed__5() { _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_2 = l_Lean_Parser_Tactic_intros___closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_intros___closed__7() { +static lean_object* _init_l_Lean_Parser_Tactic_intros___closed__6() { _start: { lean_object* x_1; @@ -1708,27 +1706,27 @@ x_1 = lean_mk_string_from_bytes("ident", 5); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_intros___closed__8() { +static lean_object* _init_l_Lean_Parser_Tactic_intros___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_intros___closed__7; +x_2 = l_Lean_Parser_Tactic_intros___closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_intros___closed__9() { +static lean_object* _init_l_Lean_Parser_Tactic_intros___closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic_intros___closed__8; +x_1 = l_Lean_Parser_Tactic_intros___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_Lean_Parser_Tactic_intros___closed__10() { +static lean_object* _init_l_Lean_Parser_Tactic_intros___closed__9() { _start: { lean_object* x_1; @@ -1736,33 +1734,33 @@ x_1 = lean_mk_string_from_bytes("hole", 4); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_intros___closed__11() { +static lean_object* _init_l_Lean_Parser_Tactic_intros___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_intros___closed__10; +x_2 = l_Lean_Parser_Tactic_intros___closed__9; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_intros___closed__12() { +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__11; +x_1 = l_Lean_Parser_Tactic_intros___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_Parser_Tactic_intros___closed__13() { +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__12; +x_1 = l_Lean_Parser_Tactic_intros___closed__5; +x_2 = l_Lean_Parser_Tactic_intros___closed__8; +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); @@ -1770,13 +1768,13 @@ 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__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__7; x_2 = l_Lean_Parser_Tactic_intro___closed__15; -x_3 = l_Lean_Parser_Tactic_intros___closed__13; +x_3 = l_Lean_Parser_Tactic_intros___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); @@ -1784,25 +1782,25 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_intros___closed__15() { +static lean_object* _init_l_Lean_Parser_Tactic_intros___closed__14() { _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__14; +x_1 = l_Lean_Parser_Tactic_intro___closed__11; +x_2 = l_Lean_Parser_Tactic_intros___closed__13; 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__16() { +static lean_object* _init_l_Lean_Parser_Tactic_intros___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__7; -x_2 = l_Lean_Parser_Tactic_intros___closed__4; -x_3 = l_Lean_Parser_Tactic_intros___closed__15; +x_2 = l_Lean_Parser_Tactic_intros___closed__3; +x_3 = l_Lean_Parser_Tactic_intros___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); @@ -1810,13 +1808,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_intros___closed__17() { +static lean_object* _init_l_Lean_Parser_Tactic_intros___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__2; x_2 = lean_unsigned_to_nat(1022u); -x_3 = l_Lean_Parser_Tactic_intros___closed__16; +x_3 = l_Lean_Parser_Tactic_intros___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); @@ -1828,7 +1826,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_intros() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Tactic_intros___closed__17; +x_1 = l_Lean_Parser_Tactic_intros___closed__16; return x_1; } } @@ -1936,7 +1934,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_rename___closed__10() { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__7; x_2 = l_Lean_Parser_Tactic_rename___closed__9; -x_3 = l_Lean_Parser_Tactic_intros___closed__9; +x_3 = l_Lean_Parser_Tactic_intros___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); @@ -1989,16 +1987,8 @@ return x_5; static lean_object* _init_l_Lean_Parser_Tactic_revert___closed__3() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("revert ", 7); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Tactic_revert___closed__4() { -_start: -{ lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_revert___closed__3; +x_1 = l_Lean_Parser_Tactic_revert___closed__1; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -2006,7 +1996,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_revert___closed__5() { +static lean_object* _init_l_Lean_Parser_Tactic_revert___closed__4() { _start: { lean_object* x_1; @@ -2014,21 +2004,21 @@ x_1 = lean_mk_string_from_bytes("many1", 5); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_revert___closed__6() { +static lean_object* _init_l_Lean_Parser_Tactic_revert___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic_revert___closed__5; +x_2 = l_Lean_Parser_Tactic_revert___closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_revert___closed__7() { +static lean_object* _init_l_Lean_Parser_Tactic_revert___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_1 = l_Lean_Parser_Tactic_revert___closed__5; x_2 = l_Lean_Parser_Tactic_intro___closed__19; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -2036,13 +2026,13 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_revert___closed__8() { +static lean_object* _init_l_Lean_Parser_Tactic_revert___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__7; -x_2 = l_Lean_Parser_Tactic_revert___closed__4; -x_3 = l_Lean_Parser_Tactic_revert___closed__7; +x_2 = l_Lean_Parser_Tactic_revert___closed__3; +x_3 = l_Lean_Parser_Tactic_revert___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); @@ -2050,13 +2040,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_revert___closed__9() { +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_revert___closed__2; x_2 = lean_unsigned_to_nat(1022u); -x_3 = l_Lean_Parser_Tactic_revert___closed__8; +x_3 = l_Lean_Parser_Tactic_revert___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); @@ -2068,7 +2058,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_revert() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Tactic_revert___closed__9; +x_1 = l_Lean_Parser_Tactic_revert___closed__8; return x_1; } } @@ -2095,16 +2085,8 @@ return x_5; static lean_object* _init_l_Lean_Parser_Tactic_clear___closed__3() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("clear ", 6); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Tactic_clear___closed__4() { -_start: -{ lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_clear___closed__3; +x_1 = l_Lean_Parser_Tactic_clear___closed__1; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -2112,13 +2094,13 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_clear___closed__5() { +static lean_object* _init_l_Lean_Parser_Tactic_clear___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__7; -x_2 = l_Lean_Parser_Tactic_clear___closed__4; -x_3 = l_Lean_Parser_Tactic_revert___closed__7; +x_2 = l_Lean_Parser_Tactic_clear___closed__3; +x_3 = l_Lean_Parser_Tactic_revert___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); @@ -2126,13 +2108,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_clear___closed__6() { +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_clear___closed__2; x_2 = lean_unsigned_to_nat(1022u); -x_3 = l_Lean_Parser_Tactic_clear___closed__5; +x_3 = l_Lean_Parser_Tactic_clear___closed__4; x_4 = lean_alloc_ctor(3, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -2144,7 +2126,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_clear() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Tactic_clear___closed__6; +x_1 = l_Lean_Parser_Tactic_clear___closed__5; return x_1; } } @@ -2171,16 +2153,8 @@ return x_5; static lean_object* _init_l_Lean_Parser_Tactic_subst___closed__3() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("subst ", 6); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Tactic_subst___closed__4() { -_start: -{ lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_subst___closed__3; +x_1 = l_Lean_Parser_Tactic_subst___closed__1; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -2188,13 +2162,13 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_subst___closed__5() { +static lean_object* _init_l_Lean_Parser_Tactic_subst___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__7; -x_2 = l_Lean_Parser_Tactic_subst___closed__4; -x_3 = l_Lean_Parser_Tactic_revert___closed__7; +x_2 = l_Lean_Parser_Tactic_subst___closed__3; +x_3 = l_Lean_Parser_Tactic_revert___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); @@ -2202,13 +2176,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_subst___closed__6() { +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_subst___closed__2; x_2 = lean_unsigned_to_nat(1022u); -x_3 = l_Lean_Parser_Tactic_subst___closed__5; +x_3 = l_Lean_Parser_Tactic_subst___closed__4; x_4 = lean_alloc_ctor(3, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -2220,7 +2194,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_subst() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Tactic_subst___closed__6; +x_1 = l_Lean_Parser_Tactic_subst___closed__5; return x_1; } } @@ -3066,7 +3040,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_tacticNext___x3d_x3e_____closed__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_intro___closed__12; +x_1 = l_Lean_Parser_Tactic_intro___closed__11; x_2 = l_Lean_binderIdent; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -3209,7 +3183,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__1; x_2 = l_Lean_Parser_Tactic_withAnnotateState___closed__2; x_3 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__7; -x_4 = l_Lean_Parser_Tactic_intros___closed__10; +x_4 = l_Lean_Parser_Tactic_intros___closed__9; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } @@ -4340,7 +4314,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_first___closed__9() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("colGe", 5); +x_1 = lean_mk_string_from_bytes("ppDedent", 8); return x_1; } } @@ -4357,20 +4331,78 @@ return x_3; static lean_object* _init_l_Lean_Parser_Tactic_first___closed__11() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("ppLine", 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_first___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_first___closed__11; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_first___closed__13() { +_start: +{ lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Tactic_first___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_Parser_Tactic_first___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Tactic_first___closed__10; +x_2 = l_Lean_Parser_Tactic_first___closed__13; +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_first___closed__15() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("colGe", 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_first___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_first___closed__15; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_first___closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Tactic_first___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_first___closed__12() { +static lean_object* _init_l_Lean_Parser_Tactic_first___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__7; -x_2 = l_Lean_Parser_Tactic_first___closed__11; -x_3 = l_Lean_Parser_Tactic_intro___closed__8; +x_2 = l_Lean_Parser_Tactic_first___closed__14; +x_3 = l_Lean_Parser_Tactic_first___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); @@ -4378,12 +4410,44 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_first___closed__13() { +static lean_object* _init_l_Lean_Parser_Tactic_first___closed__19() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("| ", 2); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_first___closed__20() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Tactic_first___closed__19; +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_first___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__7; +x_2 = l_Lean_Parser_Tactic_first___closed__18; +x_3 = l_Lean_Parser_Tactic_first___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); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_first___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__7; -x_2 = l_Lean_Parser_Tactic_first___closed__12; +x_2 = l_Lean_Parser_Tactic_first___closed__21; x_3 = l_Lean_Parser_Tactic_case___closed__12; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -4392,49 +4456,49 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_first___closed__14() { +static lean_object* _init_l_Lean_Parser_Tactic_first___closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Tactic_first___closed__8; -x_2 = l_Lean_Parser_Tactic_first___closed__13; +x_2 = l_Lean_Parser_Tactic_first___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_Parser_Tactic_first___closed__15() { +static lean_object* _init_l_Lean_Parser_Tactic_first___closed__24() { _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_first___closed__14; +x_1 = l_Lean_Parser_Tactic_revert___closed__5; +x_2 = l_Lean_Parser_Tactic_first___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_Tactic_first___closed__16() { +static lean_object* _init_l_Lean_Parser_Tactic_first___closed__25() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Tactic_first___closed__6; -x_2 = l_Lean_Parser_Tactic_first___closed__15; +x_2 = l_Lean_Parser_Tactic_first___closed__24; 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_first___closed__17() { +static lean_object* _init_l_Lean_Parser_Tactic_first___closed__26() { _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__7; x_2 = l_Lean_Parser_Tactic_first___closed__4; -x_3 = l_Lean_Parser_Tactic_first___closed__16; +x_3 = l_Lean_Parser_Tactic_first___closed__25; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -4442,13 +4506,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_first___closed__18() { +static lean_object* _init_l_Lean_Parser_Tactic_first___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_first___closed__2; x_2 = lean_unsigned_to_nat(1022u); -x_3 = l_Lean_Parser_Tactic_first___closed__17; +x_3 = l_Lean_Parser_Tactic_first___closed__26; x_4 = lean_alloc_ctor(3, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -4460,7 +4524,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_first() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Tactic_first___closed__18; +x_1 = l_Lean_Parser_Tactic_first___closed__27; return x_1; } } @@ -4553,22 +4617,36 @@ return x_2; static lean_object* _init_l_Lean_Parser_Tactic_rotateLeft___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__7; +x_2 = l_Lean_Parser_Tactic_withAnnotateState___closed__16; +x_3 = l_Lean_Parser_Tactic_rotateLeft___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); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_rotateLeft___closed__11() { +_start: +{ 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_rotateLeft___closed__9; +x_2 = l_Lean_Parser_Tactic_rotateLeft___closed__10; 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_rotateLeft___closed__11() { +static lean_object* _init_l_Lean_Parser_Tactic_rotateLeft___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__7; x_2 = l_Lean_Parser_Tactic_rotateLeft___closed__4; -x_3 = l_Lean_Parser_Tactic_rotateLeft___closed__10; +x_3 = l_Lean_Parser_Tactic_rotateLeft___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); @@ -4576,13 +4654,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_rotateLeft___closed__12() { +static lean_object* _init_l_Lean_Parser_Tactic_rotateLeft___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_rotateLeft___closed__2; x_2 = lean_unsigned_to_nat(1022u); -x_3 = l_Lean_Parser_Tactic_rotateLeft___closed__11; +x_3 = l_Lean_Parser_Tactic_rotateLeft___closed__12; x_4 = lean_alloc_ctor(3, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -4594,7 +4672,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_rotateLeft() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Tactic_rotateLeft___closed__12; +x_1 = l_Lean_Parser_Tactic_rotateLeft___closed__13; return x_1; } } @@ -4644,7 +4722,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_rotateRight___closed__5() { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__7; x_2 = l_Lean_Parser_Tactic_rotateRight___closed__4; -x_3 = l_Lean_Parser_Tactic_rotateLeft___closed__10; +x_3 = l_Lean_Parser_Tactic_rotateLeft___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); @@ -4816,7 +4894,7 @@ 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; +x_15 = l_Lean_Parser_Tactic_intro___closed__6; lean_inc(x_12); x_16 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_16, 0, x_12); @@ -6357,7 +6435,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_locationWildcard___closed__3() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("*", 1); +x_1 = lean_mk_string_from_bytes(" *", 2); return x_1; } } @@ -6416,22 +6494,36 @@ return x_5; static lean_object* _init_l_Lean_Parser_Tactic_locationHyp___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__7; +x_2 = l_Lean_Parser_Tactic_revert___closed__6; +x_3 = l_Lean_Parser_Tactic_withAnnotateState___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); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_locationHyp___closed__4() { +_start: +{ lean_object* x_1; x_1 = lean_mk_string_from_bytes("patternIgnore", 13); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_locationHyp___closed__4() { +static lean_object* _init_l_Lean_Parser_Tactic_locationHyp___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic_locationHyp___closed__3; +x_2 = l_Lean_Parser_Tactic_locationHyp___closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_locationHyp___closed__5() { +static lean_object* _init_l_Lean_Parser_Tactic_locationHyp___closed__6() { _start: { lean_object* x_1; @@ -6439,7 +6531,7 @@ x_1 = lean_mk_string_from_bytes("token", 5); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_locationHyp___closed__6() { +static lean_object* _init_l_Lean_Parser_Tactic_locationHyp___closed__7() { _start: { lean_object* x_1; @@ -6447,33 +6539,33 @@ x_1 = lean_mk_string_from_bytes("⊢", 3); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_locationHyp___closed__7() { +static lean_object* _init_l_Lean_Parser_Tactic_locationHyp___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_locationHyp___closed__5; -x_2 = l_Lean_Parser_Tactic_locationHyp___closed__6; +x_1 = l_Lean_Parser_Tactic_locationHyp___closed__6; +x_2 = l_Lean_Parser_Tactic_locationHyp___closed__7; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_locationHyp___closed__8() { +static lean_object* _init_l_Lean_Parser_Tactic_locationHyp___closed__9() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic_locationHyp___closed__6; +x_1 = l_Lean_Parser_Tactic_locationHyp___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_Parser_Tactic_locationHyp___closed__9() { +static lean_object* _init_l_Lean_Parser_Tactic_locationHyp___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_locationHyp___closed__6; -x_2 = l_Lean_Parser_Tactic_locationHyp___closed__7; -x_3 = l_Lean_Parser_Tactic_locationHyp___closed__8; +x_1 = l_Lean_Parser_Tactic_locationHyp___closed__7; +x_2 = l_Lean_Parser_Tactic_locationHyp___closed__8; +x_3 = l_Lean_Parser_Tactic_locationHyp___closed__9; x_4 = lean_alloc_ctor(9, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -6481,7 +6573,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_locationHyp___closed__10() { +static lean_object* _init_l_Lean_Parser_Tactic_locationHyp___closed__11() { _start: { lean_object* x_1; @@ -6489,33 +6581,33 @@ x_1 = lean_mk_string_from_bytes("|-", 2); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_locationHyp___closed__11() { +static lean_object* _init_l_Lean_Parser_Tactic_locationHyp___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_locationHyp___closed__5; -x_2 = l_Lean_Parser_Tactic_locationHyp___closed__10; +x_1 = l_Lean_Parser_Tactic_locationHyp___closed__6; +x_2 = l_Lean_Parser_Tactic_locationHyp___closed__11; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_locationHyp___closed__12() { +static lean_object* _init_l_Lean_Parser_Tactic_locationHyp___closed__13() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic_locationHyp___closed__10; +x_1 = l_Lean_Parser_Tactic_locationHyp___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_locationHyp___closed__13() { +static lean_object* _init_l_Lean_Parser_Tactic_locationHyp___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_locationHyp___closed__10; -x_2 = l_Lean_Parser_Tactic_locationHyp___closed__11; -x_3 = l_Lean_Parser_Tactic_locationHyp___closed__12; +x_1 = l_Lean_Parser_Tactic_locationHyp___closed__11; +x_2 = l_Lean_Parser_Tactic_locationHyp___closed__12; +x_3 = l_Lean_Parser_Tactic_locationHyp___closed__13; x_4 = lean_alloc_ctor(9, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -6523,13 +6615,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_locationHyp___closed__14() { +static lean_object* _init_l_Lean_Parser_Tactic_locationHyp___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_intros___closed__6; -x_2 = l_Lean_Parser_Tactic_locationHyp___closed__9; -x_3 = l_Lean_Parser_Tactic_locationHyp___closed__13; +x_1 = l_Lean_Parser_Tactic_intros___closed__5; +x_2 = l_Lean_Parser_Tactic_locationHyp___closed__10; +x_3 = l_Lean_Parser_Tactic_locationHyp___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); @@ -6537,37 +6629,37 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_locationHyp___closed__15() { +static lean_object* _init_l_Lean_Parser_Tactic_locationHyp___closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_locationHyp___closed__4; -x_2 = l_Lean_Parser_Tactic_locationHyp___closed__14; +x_1 = l_Lean_Parser_Tactic_locationHyp___closed__5; +x_2 = l_Lean_Parser_Tactic_locationHyp___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_Parser_Tactic_locationHyp___closed__16() { +static lean_object* _init_l_Lean_Parser_Tactic_locationHyp___closed__17() { _start: { 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_locationHyp___closed__15; +x_2 = l_Lean_Parser_Tactic_locationHyp___closed__16; 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_locationHyp___closed__17() { +static lean_object* _init_l_Lean_Parser_Tactic_locationHyp___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__7; -x_2 = l_Lean_Parser_Tactic_revert___closed__7; -x_3 = l_Lean_Parser_Tactic_locationHyp___closed__16; +x_2 = l_Lean_Parser_Tactic_locationHyp___closed__3; +x_3 = l_Lean_Parser_Tactic_locationHyp___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); @@ -6575,13 +6667,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_locationHyp___closed__18() { +static lean_object* _init_l_Lean_Parser_Tactic_locationHyp___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_locationHyp___closed__1; x_2 = l_Lean_Parser_Tactic_locationHyp___closed__2; -x_3 = l_Lean_Parser_Tactic_locationHyp___closed__17; +x_3 = l_Lean_Parser_Tactic_locationHyp___closed__18; x_4 = lean_alloc_ctor(9, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -6593,7 +6685,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_locationHyp() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Tactic_locationHyp___closed__18; +x_1 = l_Lean_Parser_Tactic_locationHyp___closed__19; return x_1; } } @@ -6621,7 +6713,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_location___closed__3() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes(" at ", 4); +x_1 = lean_mk_string_from_bytes(" at", 3); return x_1; } } @@ -6639,7 +6731,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_Parser_Tactic_intros___closed__5; x_2 = l_Lean_Parser_Tactic_locationWildcard; x_3 = l_Lean_Parser_Tactic_locationHyp; x_4 = lean_alloc_ctor(2, 3, 0); @@ -6933,7 +7025,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_rwRule___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_locationHyp___closed__5; +x_1 = l_Lean_Parser_Tactic_locationHyp___closed__6; x_2 = l_Lean_Parser_Tactic_rwRule___closed__3; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; @@ -6975,7 +7067,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_rwRule___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_locationHyp___closed__5; +x_1 = l_Lean_Parser_Tactic_locationHyp___closed__6; x_2 = l_Lean_Parser_Tactic_rwRule___closed__7; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; @@ -7009,7 +7101,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_rwRule___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_intros___closed__6; +x_1 = l_Lean_Parser_Tactic_intros___closed__5; x_2 = l_Lean_Parser_Tactic_rwRule___closed__6; x_3 = l_Lean_Parser_Tactic_rwRule___closed__10; x_4 = lean_alloc_ctor(2, 3, 0); @@ -7023,7 +7115,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_rwRule___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_locationHyp___closed__4; +x_1 = l_Lean_Parser_Tactic_locationHyp___closed__5; x_2 = l_Lean_Parser_Tactic_rwRule___closed__11; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7890,22 +7982,40 @@ return x_4; static lean_object* _init_l_Lean_Parser_Tactic_injection___closed__6() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(" with", 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_injection___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Tactic_injection___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_Parser_Tactic_injection___closed__8() { +_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_intros___closed__14; +x_1 = l_Lean_Parser_Tactic_revert___closed__5; +x_2 = l_Lean_Parser_Tactic_intros___closed__13; 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_injection___closed__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__7; -x_2 = l_Lean_Parser_Tactic_changeWith___closed__4; -x_3 = l_Lean_Parser_Tactic_injection___closed__6; +x_2 = l_Lean_Parser_Tactic_injection___closed__7; +x_3 = l_Lean_Parser_Tactic_injection___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); @@ -7913,25 +8023,25 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_injection___closed__8() { +static lean_object* _init_l_Lean_Parser_Tactic_injection___closed__10() { _start: { 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_injection___closed__7; +x_2 = l_Lean_Parser_Tactic_injection___closed__9; 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_injection___closed__9() { +static lean_object* _init_l_Lean_Parser_Tactic_injection___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__7; x_2 = l_Lean_Parser_Tactic_injection___closed__5; -x_3 = l_Lean_Parser_Tactic_injection___closed__8; +x_3 = l_Lean_Parser_Tactic_injection___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); @@ -7939,13 +8049,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_injection___closed__10() { +static lean_object* _init_l_Lean_Parser_Tactic_injection___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_injection___closed__2; x_2 = lean_unsigned_to_nat(1022u); -x_3 = l_Lean_Parser_Tactic_injection___closed__9; +x_3 = l_Lean_Parser_Tactic_injection___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); @@ -7957,7 +8067,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_injection() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Tactic_injection___closed__10; +x_1 = l_Lean_Parser_Tactic_injection___closed__12; return x_1; } } @@ -7999,7 +8109,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_injections___closed__4() { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__7; x_2 = l_Lean_Parser_Tactic_injections___closed__3; -x_3 = l_Lean_Parser_Tactic_intros___closed__15; +x_3 = l_Lean_Parser_Tactic_intros___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); @@ -8053,7 +8163,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_discharger___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_locationHyp___closed__5; +x_1 = l_Lean_Parser_Tactic_locationHyp___closed__6; x_2 = l_Lean_Parser_Tactic_discharger___closed__1; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; @@ -8097,7 +8207,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; -x_1 = l_Lean_Parser_Tactic_locationHyp___closed__5; +x_1 = l_Lean_Parser_Tactic_locationHyp___closed__6; x_2 = l_Lean_Parser_Tactic_discharger___closed__6; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; @@ -8133,7 +8243,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_intros___closed__6; +x_1 = l_Lean_Parser_Tactic_intros___closed__5; x_2 = l_Lean_Parser_Tactic_discharger___closed__5; x_3 = l_Lean_Parser_Tactic_discharger___closed__9; x_4 = lean_alloc_ctor(2, 3, 0); @@ -8147,7 +8257,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; -x_1 = l_Lean_Parser_Tactic_locationHyp___closed__4; +x_1 = l_Lean_Parser_Tactic_locationHyp___closed__5; x_2 = l_Lean_Parser_Tactic_discharger___closed__10; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -8389,7 +8499,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_Parser_Tactic_intros___closed__5; x_2 = l_Lean_Parser_Tactic_simpPre; x_3 = l_Lean_Parser_Tactic_simpPost; x_4 = lean_alloc_ctor(2, 3, 0); @@ -8558,10 +8668,28 @@ return x_5; static lean_object* _init_l_Lean_Parser_Tactic_simpStar___closed__3() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("*", 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_simpStar___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Tactic_simpStar___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_simpStar___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_simpStar___closed__1; x_2 = l_Lean_Parser_Tactic_simpStar___closed__2; -x_3 = l_Lean_Parser_Tactic_locationWildcard___closed__4; +x_3 = l_Lean_Parser_Tactic_simpStar___closed__4; x_4 = lean_alloc_ctor(9, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -8573,7 +8701,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_simpStar() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Tactic_simpStar___closed__3; +x_1 = l_Lean_Parser_Tactic_simpStar___closed__5; return x_1; } } @@ -8699,7 +8827,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_intros___closed__6; +x_1 = l_Lean_Parser_Tactic_intros___closed__5; x_2 = l_Lean_Parser_Tactic_simpErase; x_3 = l_Lean_Parser_Tactic_simpLemma; x_4 = lean_alloc_ctor(2, 3, 0); @@ -8713,7 +8841,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_Parser_Tactic_intros___closed__5; x_2 = l_Lean_Parser_Tactic_simpStar; x_3 = l_Lean_Parser_Tactic_simp___closed__11; x_4 = lean_alloc_ctor(2, 3, 0); @@ -9174,16 +9302,8 @@ return x_5; static lean_object* _init_l_Lean_Parser_Tactic_delta___closed__3() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("delta ", 6); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Tactic_delta___closed__4() { -_start: -{ lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_delta___closed__3; +x_1 = l_Lean_Parser_Tactic_delta___closed__1; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -9191,13 +9311,13 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_delta___closed__5() { +static lean_object* _init_l_Lean_Parser_Tactic_delta___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__7; x_2 = l_Lean_Parser_Tactic_intro___closed__15; -x_3 = l_Lean_Parser_Tactic_intros___closed__9; +x_3 = l_Lean_Parser_Tactic_intros___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); @@ -9205,25 +9325,25 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_delta___closed__6() { +static lean_object* _init_l_Lean_Parser_Tactic_delta___closed__5() { _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_delta___closed__5; +x_1 = l_Lean_Parser_Tactic_revert___closed__5; +x_2 = l_Lean_Parser_Tactic_delta___closed__4; 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_delta___closed__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__7; -x_2 = l_Lean_Parser_Tactic_delta___closed__4; -x_3 = l_Lean_Parser_Tactic_delta___closed__6; +x_2 = l_Lean_Parser_Tactic_delta___closed__3; +x_3 = l_Lean_Parser_Tactic_delta___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); @@ -9231,12 +9351,12 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_delta___closed__8() { +static lean_object* _init_l_Lean_Parser_Tactic_delta___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__7; -x_2 = l_Lean_Parser_Tactic_delta___closed__7; +x_2 = l_Lean_Parser_Tactic_delta___closed__6; x_3 = l_Lean_Parser_Tactic_change___closed__6; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -9245,13 +9365,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_delta___closed__9() { +static lean_object* _init_l_Lean_Parser_Tactic_delta___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_delta___closed__2; x_2 = lean_unsigned_to_nat(1022u); -x_3 = l_Lean_Parser_Tactic_delta___closed__8; +x_3 = l_Lean_Parser_Tactic_delta___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); @@ -9263,7 +9383,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_delta() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Tactic_delta___closed__9; +x_1 = l_Lean_Parser_Tactic_delta___closed__8; return x_1; } } @@ -9290,16 +9410,8 @@ return x_5; static lean_object* _init_l_Lean_Parser_Tactic_unfold___closed__3() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("unfold ", 7); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Tactic_unfold___closed__4() { -_start: -{ lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_unfold___closed__3; +x_1 = l_Lean_Parser_Tactic_unfold___closed__1; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -9307,13 +9419,13 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_unfold___closed__5() { +static lean_object* _init_l_Lean_Parser_Tactic_unfold___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__7; -x_2 = l_Lean_Parser_Tactic_unfold___closed__4; -x_3 = l_Lean_Parser_Tactic_delta___closed__6; +x_2 = l_Lean_Parser_Tactic_unfold___closed__3; +x_3 = l_Lean_Parser_Tactic_delta___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); @@ -9321,12 +9433,12 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_unfold___closed__6() { +static lean_object* _init_l_Lean_Parser_Tactic_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_withAnnotateState___closed__7; -x_2 = l_Lean_Parser_Tactic_unfold___closed__5; +x_2 = l_Lean_Parser_Tactic_unfold___closed__4; x_3 = l_Lean_Parser_Tactic_change___closed__6; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -9335,13 +9447,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_unfold___closed__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_unfold___closed__2; x_2 = lean_unsigned_to_nat(1022u); -x_3 = l_Lean_Parser_Tactic_unfold___closed__6; +x_3 = l_Lean_Parser_Tactic_unfold___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); @@ -9353,7 +9465,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_unfold() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Tactic_unfold___closed__7; +x_1 = l_Lean_Parser_Tactic_unfold___closed__6; return x_1; } } @@ -11026,7 +11138,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____clos lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__7; x_2 = l_Lean_Parser_Tactic_tacticHave_x27___x3a_x3d_____closed__4; -x_3 = l_Lean_Parser_Tactic_intros___closed__9; +x_3 = l_Lean_Parser_Tactic_intros___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); @@ -11406,24 +11518,6 @@ return x_5; static lean_object* _init_l_Lean_Parser_Tactic_inductionAltLHS___closed__3() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("| ", 2); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Tactic_inductionAltLHS___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic_inductionAltLHS___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_inductionAltLHS___closed__5() { -_start: -{ lean_object* x_1; lean_object* x_2; x_1 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticSorry__1___closed__5; x_2 = lean_alloc_ctor(5, 1, 0); @@ -11431,25 +11525,25 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic_inductionAltLHS___closed__6() { +static lean_object* _init_l_Lean_Parser_Tactic_inductionAltLHS___closed__4() { _start: { 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_inductionAltLHS___closed__5; +x_2 = l_Lean_Parser_Tactic_inductionAltLHS___closed__3; 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__7() { +static lean_object* _init_l_Lean_Parser_Tactic_inductionAltLHS___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__7; -x_2 = l_Lean_Parser_Tactic_inductionAltLHS___closed__6; -x_3 = l_Lean_Parser_Tactic_intros___closed__9; +x_2 = l_Lean_Parser_Tactic_inductionAltLHS___closed__4; +x_3 = l_Lean_Parser_Tactic_intros___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); @@ -11457,25 +11551,25 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_inductionAltLHS___closed__8() { +static lean_object* _init_l_Lean_Parser_Tactic_inductionAltLHS___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Tactic_first___closed__8; -x_2 = l_Lean_Parser_Tactic_inductionAltLHS___closed__7; +x_2 = l_Lean_Parser_Tactic_inductionAltLHS___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_inductionAltLHS___closed__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_intros___closed__6; -x_2 = l_Lean_Parser_Tactic_inductionAltLHS___closed__8; -x_3 = l_Lean_Parser_Tactic_intros___closed__12; +x_1 = l_Lean_Parser_Tactic_intros___closed__5; +x_2 = l_Lean_Parser_Tactic_inductionAltLHS___closed__6; +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); @@ -11483,13 +11577,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_inductionAltLHS___closed__10() { +static lean_object* _init_l_Lean_Parser_Tactic_inductionAltLHS___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__7; -x_2 = l_Lean_Parser_Tactic_inductionAltLHS___closed__4; -x_3 = l_Lean_Parser_Tactic_inductionAltLHS___closed__9; +x_2 = l_Lean_Parser_Tactic_first___closed__20; +x_3 = l_Lean_Parser_Tactic_inductionAltLHS___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); @@ -11497,25 +11591,25 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_inductionAltLHS___closed__11() { +static lean_object* _init_l_Lean_Parser_Tactic_inductionAltLHS___closed__9() { _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_1 = l_Lean_Parser_Tactic_intro___closed__11; +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_inductionAltLHS___closed__12() { +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__7; -x_2 = l_Lean_Parser_Tactic_inductionAltLHS___closed__10; -x_3 = l_Lean_Parser_Tactic_inductionAltLHS___closed__11; +x_2 = l_Lean_Parser_Tactic_inductionAltLHS___closed__8; +x_3 = l_Lean_Parser_Tactic_inductionAltLHS___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); @@ -11523,13 +11617,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_inductionAltLHS___closed__13() { +static lean_object* _init_l_Lean_Parser_Tactic_inductionAltLHS___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_inductionAltLHS___closed__1; x_2 = l_Lean_Parser_Tactic_inductionAltLHS___closed__2; -x_3 = l_Lean_Parser_Tactic_inductionAltLHS___closed__12; +x_3 = l_Lean_Parser_Tactic_inductionAltLHS___closed__10; x_4 = lean_alloc_ctor(9, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -11541,7 +11635,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_inductionAltLHS() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Tactic_inductionAltLHS___closed__13; +x_1 = l_Lean_Parser_Tactic_inductionAltLHS___closed__11; return x_1; } } @@ -11568,66 +11662,8 @@ return x_5; static lean_object* _init_l_Lean_Parser_Tactic_inductionAlt___closed__3() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("ppDedent", 8); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Tactic_inductionAlt___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_Tactic_inductionAlt___closed__3; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Parser_Tactic_inductionAlt___closed__5() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("ppLine", 6); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Tactic_inductionAlt___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_inductionAlt___closed__5; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Parser_Tactic_inductionAlt___closed__7() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic_inductionAlt___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_Tactic_inductionAlt___closed__8() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_inductionAlt___closed__4; -x_2 = l_Lean_Parser_Tactic_inductionAlt___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_inductionAlt___closed__9() { -_start: -{ lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_revert___closed__6; +x_1 = l_Lean_Parser_Tactic_revert___closed__5; x_2 = l_Lean_Parser_Tactic_inductionAltLHS; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11635,13 +11671,13 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_inductionAlt___closed__10() { +static lean_object* _init_l_Lean_Parser_Tactic_inductionAlt___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__7; -x_2 = l_Lean_Parser_Tactic_inductionAlt___closed__8; -x_3 = l_Lean_Parser_Tactic_inductionAlt___closed__9; +x_2 = l_Lean_Parser_Tactic_first___closed__14; +x_3 = l_Lean_Parser_Tactic_inductionAlt___closed__3; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -11649,12 +11685,12 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_inductionAlt___closed__11() { +static lean_object* _init_l_Lean_Parser_Tactic_inductionAlt___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__7; -x_2 = l_Lean_Parser_Tactic_inductionAlt___closed__10; +x_2 = l_Lean_Parser_Tactic_inductionAlt___closed__4; 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); @@ -11663,7 +11699,7 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_inductionAlt___closed__12() { +static lean_object* _init_l_Lean_Parser_Tactic_inductionAlt___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -11673,22 +11709,22 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_inductionAlt___closed__13() { +static lean_object* _init_l_Lean_Parser_Tactic_inductionAlt___closed__7() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic_inductionAlt___closed__12; +x_1 = l_Lean_Parser_Tactic_inductionAlt___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_Tactic_inductionAlt___closed__14() { +static lean_object* _init_l_Lean_Parser_Tactic_inductionAlt___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__6; -x_2 = l_Lean_Parser_Tactic_inductionAlt___closed__13; +x_1 = l_Lean_Parser_Tactic_intros___closed__5; +x_2 = l_Lean_Parser_Tactic_inductionAlt___closed__7; x_3 = l_Lean_Parser_Tactic_case___closed__12; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -11697,13 +11733,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_inductionAlt___closed__15() { +static lean_object* _init_l_Lean_Parser_Tactic_inductionAlt___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_2 = l_Lean_Parser_Tactic_intros___closed__12; -x_3 = l_Lean_Parser_Tactic_inductionAlt___closed__14; +x_1 = l_Lean_Parser_Tactic_intros___closed__5; +x_2 = l_Lean_Parser_Tactic_intros___closed__11; +x_3 = l_Lean_Parser_Tactic_inductionAlt___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); @@ -11711,13 +11747,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_inductionAlt___closed__16() { +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__7; -x_2 = l_Lean_Parser_Tactic_inductionAlt___closed__11; -x_3 = l_Lean_Parser_Tactic_inductionAlt___closed__15; +x_2 = l_Lean_Parser_Tactic_inductionAlt___closed__5; +x_3 = l_Lean_Parser_Tactic_inductionAlt___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); @@ -11725,13 +11761,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_inductionAlt___closed__17() { +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_inductionAlt___closed__1; x_2 = l_Lean_Parser_Tactic_inductionAlt___closed__2; -x_3 = l_Lean_Parser_Tactic_inductionAlt___closed__16; +x_3 = l_Lean_Parser_Tactic_inductionAlt___closed__10; x_4 = lean_alloc_ctor(9, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -11743,7 +11779,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_inductionAlt() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Tactic_inductionAlt___closed__17; +x_1 = l_Lean_Parser_Tactic_inductionAlt___closed__11; return x_1; } } @@ -11770,40 +11806,36 @@ return x_5; static lean_object* _init_l_Lean_Parser_Tactic_inductionAlts___closed__3() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("with ", 5); -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__7; +x_2 = l_Lean_Parser_Tactic_withAnnotateState___closed__16; +x_3 = l_Lean_Parser_Tactic_withAnnotateState___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); +lean_ctor_set(x_4, 2, x_3); +return x_4; } } static lean_object* _init_l_Lean_Parser_Tactic_inductionAlts___closed__4() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic_inductionAlts___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_inductionAlts___closed__5() { -_start: -{ 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__20; +x_2 = l_Lean_Parser_Tactic_inductionAlts___closed__3; 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_inductionAlts___closed__6() { +static lean_object* _init_l_Lean_Parser_Tactic_inductionAlts___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__7; -x_2 = l_Lean_Parser_Tactic_inductionAlts___closed__4; -x_3 = l_Lean_Parser_Tactic_inductionAlts___closed__5; +x_2 = l_Lean_Parser_Tactic_injection___closed__7; +x_3 = l_Lean_Parser_Tactic_inductionAlts___closed__4; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -11811,12 +11843,12 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_inductionAlts___closed__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__7; -x_2 = l_Lean_Parser_Tactic_first___closed__11; +x_2 = l_Lean_Parser_Tactic_first___closed__17; x_3 = l_Lean_Parser_Tactic_inductionAlt; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -11825,37 +11857,37 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_inductionAlts___closed__8() { +static lean_object* _init_l_Lean_Parser_Tactic_inductionAlts___closed__7() { _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_inductionAlts___closed__7; +x_1 = l_Lean_Parser_Tactic_revert___closed__5; +x_2 = l_Lean_Parser_Tactic_inductionAlts___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_Lean_Parser_Tactic_inductionAlts___closed__9() { +static lean_object* _init_l_Lean_Parser_Tactic_inductionAlts___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Tactic_first___closed__6; -x_2 = l_Lean_Parser_Tactic_inductionAlts___closed__8; +x_2 = l_Lean_Parser_Tactic_inductionAlts___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_inductionAlts___closed__10() { +static lean_object* _init_l_Lean_Parser_Tactic_inductionAlts___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__7; -x_2 = l_Lean_Parser_Tactic_inductionAlts___closed__6; -x_3 = l_Lean_Parser_Tactic_inductionAlts___closed__9; +x_2 = l_Lean_Parser_Tactic_inductionAlts___closed__5; +x_3 = l_Lean_Parser_Tactic_inductionAlts___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); @@ -11863,13 +11895,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_inductionAlts___closed__11() { +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_inductionAlts___closed__1; x_2 = l_Lean_Parser_Tactic_inductionAlts___closed__2; -x_3 = l_Lean_Parser_Tactic_inductionAlts___closed__10; +x_3 = l_Lean_Parser_Tactic_inductionAlts___closed__9; x_4 = lean_alloc_ctor(9, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -11881,7 +11913,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_inductionAlts() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Tactic_inductionAlts___closed__11; +x_1 = l_Lean_Parser_Tactic_inductionAlts___closed__10; return x_1; } } @@ -11979,7 +12011,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_induction___closed__9() { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__7; x_2 = l_Lean_Parser_Tactic_induction___closed__8; -x_3 = l_Lean_Parser_Tactic_intros___closed__9; +x_3 = l_Lean_Parser_Tactic_intros___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); @@ -12017,7 +12049,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_induction___closed__12() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("generalizing ", 13); +x_1 = lean_mk_string_from_bytes(" generalizing", 13); return x_1; } } @@ -12037,7 +12069,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_induction___closed__14() { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__7; x_2 = l_Lean_Parser_Tactic_induction___closed__13; -x_3 = l_Lean_Parser_Tactic_revert___closed__7; +x_3 = l_Lean_Parser_Tactic_revert___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); @@ -12162,7 +12194,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_generalizeArg___closed__5() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__7; -x_2 = l_Lean_Parser_Tactic_intros___closed__9; +x_2 = l_Lean_Parser_Tactic_intros___closed__8; 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); @@ -12259,7 +12291,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_generalizeArg___closed__13() { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__7; x_2 = l_Lean_Parser_Tactic_generalizeArg___closed__12; -x_3 = l_Lean_Parser_Tactic_intros___closed__9; +x_3 = l_Lean_Parser_Tactic_intros___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); @@ -12595,7 +12627,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_renameI___closed__3() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("rename_i ", 9); +x_1 = lean_mk_string_from_bytes("rename_i", 8); return x_1; } } @@ -12629,7 +12661,7 @@ 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_1 = l_Lean_Parser_Tactic_revert___closed__5; 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); @@ -12791,7 +12823,7 @@ 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; +x_15 = l_Lean_Parser_Tactic_intro___closed__6; lean_inc(x_12); x_16 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_16, 0, x_12); @@ -12947,16 +12979,8 @@ return x_5; static lean_object* _init_l_Lean_Parser_Tactic_split___closed__3() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("split ", 6); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Tactic_split___closed__4() { -_start: -{ lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_split___closed__3; +x_1 = l_Lean_Parser_Tactic_split___closed__1; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -12964,7 +12988,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_split___closed__5() { +static lean_object* _init_l_Lean_Parser_Tactic_split___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; @@ -12978,25 +13002,25 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_split___closed__6() { +static lean_object* _init_l_Lean_Parser_Tactic_split___closed__5() { _start: { 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_split___closed__5; +x_2 = l_Lean_Parser_Tactic_split___closed__4; 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_split___closed__7() { +static lean_object* _init_l_Lean_Parser_Tactic_split___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__7; -x_2 = l_Lean_Parser_Tactic_split___closed__4; -x_3 = l_Lean_Parser_Tactic_split___closed__6; +x_2 = l_Lean_Parser_Tactic_split___closed__3; +x_3 = l_Lean_Parser_Tactic_split___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); @@ -13004,12 +13028,12 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_split___closed__8() { +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__7; -x_2 = l_Lean_Parser_Tactic_split___closed__7; +x_2 = l_Lean_Parser_Tactic_split___closed__6; x_3 = l_Lean_Parser_Tactic_change___closed__6; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -13018,13 +13042,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_split___closed__9() { +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_split___closed__2; x_2 = lean_unsigned_to_nat(1022u); -x_3 = l_Lean_Parser_Tactic_split___closed__8; +x_3 = l_Lean_Parser_Tactic_split___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); @@ -13036,7 +13060,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_split() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Tactic_split___closed__9; +x_1 = l_Lean_Parser_Tactic_split___closed__8; return x_1; } } @@ -13953,16 +13977,8 @@ return x_5; static lean_object* _init_l_Lean_Parser_Tactic_fail___closed__3() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("fail ", 5); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Tactic_fail___closed__4() { -_start: -{ lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_fail___closed__3; +x_1 = l_Lean_Parser_Tactic_fail___closed__1; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -13970,12 +13986,26 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } +static lean_object* _init_l_Lean_Parser_Tactic_fail___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__7; +x_2 = l_Lean_Parser_Tactic_withAnnotateState___closed__16; +x_3 = l_Lean_Parser_Tactic_traceMessage___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); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} static lean_object* _init_l_Lean_Parser_Tactic_fail___closed__5() { _start: { 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_traceMessage___closed__7; +x_2 = l_Lean_Parser_Tactic_fail___closed__4; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -13987,7 +14017,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_fail___closed__6() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Tactic_withAnnotateState___closed__7; -x_2 = l_Lean_Parser_Tactic_fail___closed__4; +x_2 = l_Lean_Parser_Tactic_fail___closed__3; x_3 = l_Lean_Parser_Tactic_fail___closed__5; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -14209,8 +14239,16 @@ return x_5; static lean_object* _init_l_Lean_Parser_Tactic_sleep___closed__3() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("sleep ", 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_sleep___closed__4() { +_start: +{ lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_sleep___closed__1; +x_1 = l_Lean_Parser_Tactic_sleep___closed__3; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -14218,12 +14256,12 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_sleep___closed__4() { +static lean_object* _init_l_Lean_Parser_Tactic_sleep___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__7; -x_2 = l_Lean_Parser_Tactic_sleep___closed__3; +x_2 = l_Lean_Parser_Tactic_sleep___closed__4; x_3 = l_Lean_Parser_Tactic_rotateLeft___closed__9; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -14232,13 +14270,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_sleep___closed__5() { +static lean_object* _init_l_Lean_Parser_Tactic_sleep___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_sleep___closed__2; x_2 = lean_unsigned_to_nat(1022u); -x_3 = l_Lean_Parser_Tactic_sleep___closed__4; +x_3 = l_Lean_Parser_Tactic_sleep___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); @@ -14250,7 +14288,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_sleep() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Tactic_sleep___closed__5; +x_1 = l_Lean_Parser_Tactic_sleep___closed__6; return x_1; } } @@ -14527,16 +14565,8 @@ return x_5; static lean_object* _init_l_Lean_Parser_Tactic_congr___closed__3() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("congr ", 6); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Tactic_congr___closed__4() { -_start: -{ lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_congr___closed__3; +x_1 = l_Lean_Parser_Tactic_congr___closed__1; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -14544,13 +14574,13 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_congr___closed__5() { +static lean_object* _init_l_Lean_Parser_Tactic_congr___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__7; -x_2 = l_Lean_Parser_Tactic_congr___closed__4; -x_3 = l_Lean_Parser_Tactic_rotateLeft___closed__10; +x_2 = l_Lean_Parser_Tactic_congr___closed__3; +x_3 = l_Lean_Parser_Tactic_rotateLeft___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); @@ -14558,13 +14588,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Tactic_congr___closed__6() { +static lean_object* _init_l_Lean_Parser_Tactic_congr___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_congr___closed__2; x_2 = lean_unsigned_to_nat(1022u); -x_3 = l_Lean_Parser_Tactic_congr___closed__5; +x_3 = l_Lean_Parser_Tactic_congr___closed__4; x_4 = lean_alloc_ctor(3, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -14576,7 +14606,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_congr() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Tactic_congr___closed__6; +x_1 = l_Lean_Parser_Tactic_congr___closed__5; return x_1; } } @@ -14647,22 +14677,36 @@ return x_3; static lean_object* _init_l_Lean_Parser_Attr_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__7; +x_2 = l_Lean_Parser_Tactic_withAnnotateState___closed__16; +x_3 = l_Lean_Parser_Attr_simp___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_Attr_simp___closed__8() { +_start: +{ 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_Attr_simp___closed__6; +x_2 = l_Lean_Parser_Attr_simp___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_Attr_simp___closed__8() { +static lean_object* _init_l_Lean_Parser_Attr_simp___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__7; x_2 = l_Lean_Parser_Attr_simp___closed__3; -x_3 = l_Lean_Parser_Attr_simp___closed__7; +x_3 = l_Lean_Parser_Attr_simp___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); @@ -14670,13 +14714,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Attr_simp___closed__9() { +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_Attr_simp___closed__2; x_2 = lean_unsigned_to_nat(1022u); -x_3 = l_Lean_Parser_Attr_simp___closed__8; +x_3 = l_Lean_Parser_Attr_simp___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); @@ -14688,7 +14732,7 @@ static lean_object* _init_l_Lean_Parser_Attr_simp() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Attr_simp___closed__9; +x_1 = l_Lean_Parser_Attr_simp___closed__10; return x_1; } } @@ -15466,7 +15510,7 @@ 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_Lean_Parser_Tactic_intro___closed__7; +x_13 = l_Lean_Parser_Tactic_intro___closed__6; lean_inc(x_10); x_14 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_14, 0, x_10); @@ -16106,8 +16150,6 @@ l_Lean_Parser_Tactic_intros___closed__15 = _init_l_Lean_Parser_Tactic_intros___c 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___closed__17 = _init_l_Lean_Parser_Tactic_intros___closed__17(); -lean_mark_persistent(l_Lean_Parser_Tactic_intros___closed__17); 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(); @@ -16150,8 +16192,6 @@ l_Lean_Parser_Tactic_revert___closed__7 = _init_l_Lean_Parser_Tactic_revert___cl lean_mark_persistent(l_Lean_Parser_Tactic_revert___closed__7); l_Lean_Parser_Tactic_revert___closed__8 = _init_l_Lean_Parser_Tactic_revert___closed__8(); lean_mark_persistent(l_Lean_Parser_Tactic_revert___closed__8); -l_Lean_Parser_Tactic_revert___closed__9 = _init_l_Lean_Parser_Tactic_revert___closed__9(); -lean_mark_persistent(l_Lean_Parser_Tactic_revert___closed__9); l_Lean_Parser_Tactic_revert = _init_l_Lean_Parser_Tactic_revert(); lean_mark_persistent(l_Lean_Parser_Tactic_revert); l_Lean_Parser_Tactic_clear___closed__1 = _init_l_Lean_Parser_Tactic_clear___closed__1(); @@ -16164,8 +16204,6 @@ l_Lean_Parser_Tactic_clear___closed__4 = _init_l_Lean_Parser_Tactic_clear___clos lean_mark_persistent(l_Lean_Parser_Tactic_clear___closed__4); l_Lean_Parser_Tactic_clear___closed__5 = _init_l_Lean_Parser_Tactic_clear___closed__5(); lean_mark_persistent(l_Lean_Parser_Tactic_clear___closed__5); -l_Lean_Parser_Tactic_clear___closed__6 = _init_l_Lean_Parser_Tactic_clear___closed__6(); -lean_mark_persistent(l_Lean_Parser_Tactic_clear___closed__6); l_Lean_Parser_Tactic_clear = _init_l_Lean_Parser_Tactic_clear(); lean_mark_persistent(l_Lean_Parser_Tactic_clear); l_Lean_Parser_Tactic_subst___closed__1 = _init_l_Lean_Parser_Tactic_subst___closed__1(); @@ -16178,8 +16216,6 @@ l_Lean_Parser_Tactic_subst___closed__4 = _init_l_Lean_Parser_Tactic_subst___clos lean_mark_persistent(l_Lean_Parser_Tactic_subst___closed__4); l_Lean_Parser_Tactic_subst___closed__5 = _init_l_Lean_Parser_Tactic_subst___closed__5(); lean_mark_persistent(l_Lean_Parser_Tactic_subst___closed__5); -l_Lean_Parser_Tactic_subst___closed__6 = _init_l_Lean_Parser_Tactic_subst___closed__6(); -lean_mark_persistent(l_Lean_Parser_Tactic_subst___closed__6); l_Lean_Parser_Tactic_subst = _init_l_Lean_Parser_Tactic_subst(); lean_mark_persistent(l_Lean_Parser_Tactic_subst); l_Lean_Parser_Tactic_substVars___closed__1 = _init_l_Lean_Parser_Tactic_substVars___closed__1(); @@ -16582,6 +16618,24 @@ l_Lean_Parser_Tactic_first___closed__17 = _init_l_Lean_Parser_Tactic_first___clo lean_mark_persistent(l_Lean_Parser_Tactic_first___closed__17); l_Lean_Parser_Tactic_first___closed__18 = _init_l_Lean_Parser_Tactic_first___closed__18(); lean_mark_persistent(l_Lean_Parser_Tactic_first___closed__18); +l_Lean_Parser_Tactic_first___closed__19 = _init_l_Lean_Parser_Tactic_first___closed__19(); +lean_mark_persistent(l_Lean_Parser_Tactic_first___closed__19); +l_Lean_Parser_Tactic_first___closed__20 = _init_l_Lean_Parser_Tactic_first___closed__20(); +lean_mark_persistent(l_Lean_Parser_Tactic_first___closed__20); +l_Lean_Parser_Tactic_first___closed__21 = _init_l_Lean_Parser_Tactic_first___closed__21(); +lean_mark_persistent(l_Lean_Parser_Tactic_first___closed__21); +l_Lean_Parser_Tactic_first___closed__22 = _init_l_Lean_Parser_Tactic_first___closed__22(); +lean_mark_persistent(l_Lean_Parser_Tactic_first___closed__22); +l_Lean_Parser_Tactic_first___closed__23 = _init_l_Lean_Parser_Tactic_first___closed__23(); +lean_mark_persistent(l_Lean_Parser_Tactic_first___closed__23); +l_Lean_Parser_Tactic_first___closed__24 = _init_l_Lean_Parser_Tactic_first___closed__24(); +lean_mark_persistent(l_Lean_Parser_Tactic_first___closed__24); +l_Lean_Parser_Tactic_first___closed__25 = _init_l_Lean_Parser_Tactic_first___closed__25(); +lean_mark_persistent(l_Lean_Parser_Tactic_first___closed__25); +l_Lean_Parser_Tactic_first___closed__26 = _init_l_Lean_Parser_Tactic_first___closed__26(); +lean_mark_persistent(l_Lean_Parser_Tactic_first___closed__26); +l_Lean_Parser_Tactic_first___closed__27 = _init_l_Lean_Parser_Tactic_first___closed__27(); +lean_mark_persistent(l_Lean_Parser_Tactic_first___closed__27); l_Lean_Parser_Tactic_first = _init_l_Lean_Parser_Tactic_first(); lean_mark_persistent(l_Lean_Parser_Tactic_first); l_Lean_Parser_Tactic_rotateLeft___closed__1 = _init_l_Lean_Parser_Tactic_rotateLeft___closed__1(); @@ -16608,6 +16662,8 @@ l_Lean_Parser_Tactic_rotateLeft___closed__11 = _init_l_Lean_Parser_Tactic_rotate lean_mark_persistent(l_Lean_Parser_Tactic_rotateLeft___closed__11); l_Lean_Parser_Tactic_rotateLeft___closed__12 = _init_l_Lean_Parser_Tactic_rotateLeft___closed__12(); lean_mark_persistent(l_Lean_Parser_Tactic_rotateLeft___closed__12); +l_Lean_Parser_Tactic_rotateLeft___closed__13 = _init_l_Lean_Parser_Tactic_rotateLeft___closed__13(); +lean_mark_persistent(l_Lean_Parser_Tactic_rotateLeft___closed__13); l_Lean_Parser_Tactic_rotateLeft = _init_l_Lean_Parser_Tactic_rotateLeft(); lean_mark_persistent(l_Lean_Parser_Tactic_rotateLeft); l_Lean_Parser_Tactic_rotateRight___closed__1 = _init_l_Lean_Parser_Tactic_rotateRight___closed__1(); @@ -16892,6 +16948,8 @@ l_Lean_Parser_Tactic_locationHyp___closed__17 = _init_l_Lean_Parser_Tactic_locat lean_mark_persistent(l_Lean_Parser_Tactic_locationHyp___closed__17); l_Lean_Parser_Tactic_locationHyp___closed__18 = _init_l_Lean_Parser_Tactic_locationHyp___closed__18(); lean_mark_persistent(l_Lean_Parser_Tactic_locationHyp___closed__18); +l_Lean_Parser_Tactic_locationHyp___closed__19 = _init_l_Lean_Parser_Tactic_locationHyp___closed__19(); +lean_mark_persistent(l_Lean_Parser_Tactic_locationHyp___closed__19); l_Lean_Parser_Tactic_locationHyp = _init_l_Lean_Parser_Tactic_locationHyp(); lean_mark_persistent(l_Lean_Parser_Tactic_locationHyp); l_Lean_Parser_Tactic_location___closed__1 = _init_l_Lean_Parser_Tactic_location___closed__1(); @@ -17078,6 +17136,10 @@ l_Lean_Parser_Tactic_injection___closed__9 = _init_l_Lean_Parser_Tactic_injectio lean_mark_persistent(l_Lean_Parser_Tactic_injection___closed__9); l_Lean_Parser_Tactic_injection___closed__10 = _init_l_Lean_Parser_Tactic_injection___closed__10(); lean_mark_persistent(l_Lean_Parser_Tactic_injection___closed__10); +l_Lean_Parser_Tactic_injection___closed__11 = _init_l_Lean_Parser_Tactic_injection___closed__11(); +lean_mark_persistent(l_Lean_Parser_Tactic_injection___closed__11); +l_Lean_Parser_Tactic_injection___closed__12 = _init_l_Lean_Parser_Tactic_injection___closed__12(); +lean_mark_persistent(l_Lean_Parser_Tactic_injection___closed__12); l_Lean_Parser_Tactic_injection = _init_l_Lean_Parser_Tactic_injection(); lean_mark_persistent(l_Lean_Parser_Tactic_injection); l_Lean_Parser_Tactic_injections___closed__1 = _init_l_Lean_Parser_Tactic_injections___closed__1(); @@ -17188,6 +17250,10 @@ l_Lean_Parser_Tactic_simpStar___closed__2 = _init_l_Lean_Parser_Tactic_simpStar_ lean_mark_persistent(l_Lean_Parser_Tactic_simpStar___closed__2); l_Lean_Parser_Tactic_simpStar___closed__3 = _init_l_Lean_Parser_Tactic_simpStar___closed__3(); lean_mark_persistent(l_Lean_Parser_Tactic_simpStar___closed__3); +l_Lean_Parser_Tactic_simpStar___closed__4 = _init_l_Lean_Parser_Tactic_simpStar___closed__4(); +lean_mark_persistent(l_Lean_Parser_Tactic_simpStar___closed__4); +l_Lean_Parser_Tactic_simpStar___closed__5 = _init_l_Lean_Parser_Tactic_simpStar___closed__5(); +lean_mark_persistent(l_Lean_Parser_Tactic_simpStar___closed__5); l_Lean_Parser_Tactic_simpStar = _init_l_Lean_Parser_Tactic_simpStar(); lean_mark_persistent(l_Lean_Parser_Tactic_simpStar); l_Lean_Parser_Tactic_simp___closed__1 = _init_l_Lean_Parser_Tactic_simp___closed__1(); @@ -17298,8 +17364,6 @@ l_Lean_Parser_Tactic_delta___closed__7 = _init_l_Lean_Parser_Tactic_delta___clos lean_mark_persistent(l_Lean_Parser_Tactic_delta___closed__7); l_Lean_Parser_Tactic_delta___closed__8 = _init_l_Lean_Parser_Tactic_delta___closed__8(); lean_mark_persistent(l_Lean_Parser_Tactic_delta___closed__8); -l_Lean_Parser_Tactic_delta___closed__9 = _init_l_Lean_Parser_Tactic_delta___closed__9(); -lean_mark_persistent(l_Lean_Parser_Tactic_delta___closed__9); l_Lean_Parser_Tactic_delta = _init_l_Lean_Parser_Tactic_delta(); lean_mark_persistent(l_Lean_Parser_Tactic_delta); l_Lean_Parser_Tactic_unfold___closed__1 = _init_l_Lean_Parser_Tactic_unfold___closed__1(); @@ -17314,8 +17378,6 @@ l_Lean_Parser_Tactic_unfold___closed__5 = _init_l_Lean_Parser_Tactic_unfold___cl lean_mark_persistent(l_Lean_Parser_Tactic_unfold___closed__5); l_Lean_Parser_Tactic_unfold___closed__6 = _init_l_Lean_Parser_Tactic_unfold___closed__6(); lean_mark_persistent(l_Lean_Parser_Tactic_unfold___closed__6); -l_Lean_Parser_Tactic_unfold___closed__7 = _init_l_Lean_Parser_Tactic_unfold___closed__7(); -lean_mark_persistent(l_Lean_Parser_Tactic_unfold___closed__7); l_Lean_Parser_Tactic_unfold = _init_l_Lean_Parser_Tactic_unfold(); lean_mark_persistent(l_Lean_Parser_Tactic_unfold); l_Lean_Parser_Tactic_tacticRefine__lift_____closed__1 = _init_l_Lean_Parser_Tactic_tacticRefine__lift_____closed__1(); @@ -17572,10 +17634,6 @@ l_Lean_Parser_Tactic_inductionAltLHS___closed__10 = _init_l_Lean_Parser_Tactic_i lean_mark_persistent(l_Lean_Parser_Tactic_inductionAltLHS___closed__10); l_Lean_Parser_Tactic_inductionAltLHS___closed__11 = _init_l_Lean_Parser_Tactic_inductionAltLHS___closed__11(); 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(); @@ -17600,18 +17658,6 @@ l_Lean_Parser_Tactic_inductionAlt___closed__10 = _init_l_Lean_Parser_Tactic_indu lean_mark_persistent(l_Lean_Parser_Tactic_inductionAlt___closed__10); l_Lean_Parser_Tactic_inductionAlt___closed__11 = _init_l_Lean_Parser_Tactic_inductionAlt___closed__11(); lean_mark_persistent(l_Lean_Parser_Tactic_inductionAlt___closed__11); -l_Lean_Parser_Tactic_inductionAlt___closed__12 = _init_l_Lean_Parser_Tactic_inductionAlt___closed__12(); -lean_mark_persistent(l_Lean_Parser_Tactic_inductionAlt___closed__12); -l_Lean_Parser_Tactic_inductionAlt___closed__13 = _init_l_Lean_Parser_Tactic_inductionAlt___closed__13(); -lean_mark_persistent(l_Lean_Parser_Tactic_inductionAlt___closed__13); -l_Lean_Parser_Tactic_inductionAlt___closed__14 = _init_l_Lean_Parser_Tactic_inductionAlt___closed__14(); -lean_mark_persistent(l_Lean_Parser_Tactic_inductionAlt___closed__14); -l_Lean_Parser_Tactic_inductionAlt___closed__15 = _init_l_Lean_Parser_Tactic_inductionAlt___closed__15(); -lean_mark_persistent(l_Lean_Parser_Tactic_inductionAlt___closed__15); -l_Lean_Parser_Tactic_inductionAlt___closed__16 = _init_l_Lean_Parser_Tactic_inductionAlt___closed__16(); -lean_mark_persistent(l_Lean_Parser_Tactic_inductionAlt___closed__16); -l_Lean_Parser_Tactic_inductionAlt___closed__17 = _init_l_Lean_Parser_Tactic_inductionAlt___closed__17(); -lean_mark_persistent(l_Lean_Parser_Tactic_inductionAlt___closed__17); l_Lean_Parser_Tactic_inductionAlt = _init_l_Lean_Parser_Tactic_inductionAlt(); lean_mark_persistent(l_Lean_Parser_Tactic_inductionAlt); l_Lean_Parser_Tactic_inductionAlts___closed__1 = _init_l_Lean_Parser_Tactic_inductionAlts___closed__1(); @@ -17634,8 +17680,6 @@ l_Lean_Parser_Tactic_inductionAlts___closed__9 = _init_l_Lean_Parser_Tactic_indu lean_mark_persistent(l_Lean_Parser_Tactic_inductionAlts___closed__9); l_Lean_Parser_Tactic_inductionAlts___closed__10 = _init_l_Lean_Parser_Tactic_inductionAlts___closed__10(); lean_mark_persistent(l_Lean_Parser_Tactic_inductionAlts___closed__10); -l_Lean_Parser_Tactic_inductionAlts___closed__11 = _init_l_Lean_Parser_Tactic_inductionAlts___closed__11(); -lean_mark_persistent(l_Lean_Parser_Tactic_inductionAlts___closed__11); l_Lean_Parser_Tactic_inductionAlts = _init_l_Lean_Parser_Tactic_inductionAlts(); lean_mark_persistent(l_Lean_Parser_Tactic_inductionAlts); l_Lean_Parser_Tactic_induction___closed__1 = _init_l_Lean_Parser_Tactic_induction___closed__1(); @@ -17818,8 +17862,6 @@ l_Lean_Parser_Tactic_split___closed__7 = _init_l_Lean_Parser_Tactic_split___clos lean_mark_persistent(l_Lean_Parser_Tactic_split___closed__7); l_Lean_Parser_Tactic_split___closed__8 = _init_l_Lean_Parser_Tactic_split___closed__8(); lean_mark_persistent(l_Lean_Parser_Tactic_split___closed__8); -l_Lean_Parser_Tactic_split___closed__9 = _init_l_Lean_Parser_Tactic_split___closed__9(); -lean_mark_persistent(l_Lean_Parser_Tactic_split___closed__9); l_Lean_Parser_Tactic_split = _init_l_Lean_Parser_Tactic_split(); lean_mark_persistent(l_Lean_Parser_Tactic_split); l_Lean_Parser_Tactic_dbgTrace___closed__1 = _init_l_Lean_Parser_Tactic_dbgTrace___closed__1(); @@ -17968,6 +18010,8 @@ l_Lean_Parser_Tactic_sleep___closed__4 = _init_l_Lean_Parser_Tactic_sleep___clos lean_mark_persistent(l_Lean_Parser_Tactic_sleep___closed__4); l_Lean_Parser_Tactic_sleep___closed__5 = _init_l_Lean_Parser_Tactic_sleep___closed__5(); lean_mark_persistent(l_Lean_Parser_Tactic_sleep___closed__5); +l_Lean_Parser_Tactic_sleep___closed__6 = _init_l_Lean_Parser_Tactic_sleep___closed__6(); +lean_mark_persistent(l_Lean_Parser_Tactic_sleep___closed__6); l_Lean_Parser_Tactic_sleep = _init_l_Lean_Parser_Tactic_sleep(); lean_mark_persistent(l_Lean_Parser_Tactic_sleep); l_Lean_Parser_Tactic_tacticExists___x2c_x2c___closed__1 = _init_l_Lean_Parser_Tactic_tacticExists___x2c_x2c___closed__1(); @@ -18002,8 +18046,6 @@ l_Lean_Parser_Tactic_congr___closed__4 = _init_l_Lean_Parser_Tactic_congr___clos lean_mark_persistent(l_Lean_Parser_Tactic_congr___closed__4); l_Lean_Parser_Tactic_congr___closed__5 = _init_l_Lean_Parser_Tactic_congr___closed__5(); lean_mark_persistent(l_Lean_Parser_Tactic_congr___closed__5); -l_Lean_Parser_Tactic_congr___closed__6 = _init_l_Lean_Parser_Tactic_congr___closed__6(); -lean_mark_persistent(l_Lean_Parser_Tactic_congr___closed__6); l_Lean_Parser_Tactic_congr = _init_l_Lean_Parser_Tactic_congr(); lean_mark_persistent(l_Lean_Parser_Tactic_congr); l_Lean_Parser_Attr_simp___closed__1 = _init_l_Lean_Parser_Attr_simp___closed__1(); @@ -18024,6 +18066,8 @@ l_Lean_Parser_Attr_simp___closed__8 = _init_l_Lean_Parser_Attr_simp___closed__8( lean_mark_persistent(l_Lean_Parser_Attr_simp___closed__8); l_Lean_Parser_Attr_simp___closed__9 = _init_l_Lean_Parser_Attr_simp___closed__9(); lean_mark_persistent(l_Lean_Parser_Attr_simp___closed__9); +l_Lean_Parser_Attr_simp___closed__10 = _init_l_Lean_Parser_Attr_simp___closed__10(); +lean_mark_persistent(l_Lean_Parser_Attr_simp___closed__10); l_Lean_Parser_Attr_simp = _init_l_Lean_Parser_Attr_simp(); lean_mark_persistent(l_Lean_Parser_Attr_simp); l_term_u2039___u203a___closed__1 = _init_l_term_u2039___u203a___closed__1(); diff --git a/stage0/stdlib/Lean/Class.c b/stage0/stdlib/Lean/Class.c index 604daa7fad32..af91e434aaee 100644 --- a/stage0/stdlib/Lean/Class.c +++ b/stage0/stdlib/Lean/Class.c @@ -2724,7 +2724,7 @@ static lean_object* _init_l_Lean_mkOutParamArgsImplicit_go___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_mkOutParamArgsImplicit_go___closed__1; x_2 = l_Lean_mkOutParamArgsImplicit_go___closed__2; -x_3 = lean_unsigned_to_nat(1540u); +x_3 = lean_unsigned_to_nat(1544u); x_4 = lean_unsigned_to_nat(24u); x_5 = l_Lean_mkOutParamArgsImplicit_go___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -2745,7 +2745,7 @@ static lean_object* _init_l_Lean_mkOutParamArgsImplicit_go___closed__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; x_1 = l_Lean_mkOutParamArgsImplicit_go___closed__1; x_2 = l_Lean_mkOutParamArgsImplicit_go___closed__5; -x_3 = lean_unsigned_to_nat(1529u); +x_3 = lean_unsigned_to_nat(1533u); x_4 = lean_unsigned_to_nat(23u); x_5 = l_Lean_mkOutParamArgsImplicit_go___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); diff --git a/stage0/stdlib/Lean/Compiler/IR.c b/stage0/stdlib/Lean/Compiler/IR.c index 93fcac5f3a5e..917ca327bcc5 100644 --- a/stage0/stdlib/Lean/Compiler/IR.c +++ b/stage0/stdlib/Lean/Compiler/IR.c @@ -13,7 +13,6 @@ #ifdef __cplusplus extern "C" { #endif -uint8_t l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_IR_initFn____x40_Lean_Compiler_IR___hyg_6____closed__2; static lean_object* l___private_Lean_Compiler_IR_0__Lean_IR_compileAux___closed__8; static lean_object* l___private_Lean_Compiler_IR_0__Lean_IR_compileAux___lambda__2___closed__8; @@ -68,11 +67,11 @@ static lean_object* l___private_Lean_Compiler_IR_0__Lean_IR_compileAux___lambda_ static lean_object* l___private_Lean_Compiler_IR_0__Lean_IR_compileAux___lambda__2___closed__2; static lean_object* l___private_Lean_Compiler_IR_0__Lean_IR_compileAux___lambda__2___closed__5; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); uint8_t l_Lean_IR_ExplicitBoxing_requiresBoxedVersion(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Compiler_IR_0__Lean_IR_compileAux___spec__3(size_t, size_t, lean_object*); lean_object* l_Lean_IR_checkDecls(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Compiler_IR_0__Lean_IR_compileAux___spec__1(size_t, size_t, lean_object*); -lean_object* l_Lean_Option_register___at_Std_Format_initFn____x40_Lean_Data_Format___hyg_95____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Compiler_IR_0__Lean_IR_compileAux___spec__4___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_IR_0__Lean_IR_compileAux___lambda__2___closed__4; LEAN_EXPORT lean_object* l_Lean_IR_addBoxedVersionAux___boxed(lean_object*, lean_object*, lean_object*); @@ -89,6 +88,7 @@ LEAN_EXPORT lean_object* lean_ir_add_boxed_version(lean_object*, lean_object*); lean_object* l_Lean_IR_explicitRC(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_initFn____x40_Lean_Compiler_IR___hyg_6_(lean_object*); lean_object* lean_ir_add_decl(lean_object*, lean_object*); +lean_object* l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_IR_0__Lean_IR_compileAux___lambda__2___closed__9; extern lean_object* l_Lean_IR_tracePrefixOptionName; static lean_object* l_Lean_IR_initFn____x40_Lean_Compiler_IR___hyg_6____closed__6; @@ -212,7 +212,7 @@ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_2 = l_Lean_IR_initFn____x40_Lean_Compiler_IR___hyg_6____closed__3; x_3 = l_Lean_IR_initFn____x40_Lean_Compiler_IR___hyg_6____closed__6; x_4 = l_Lean_IR_initFn____x40_Lean_Compiler_IR___hyg_6____closed__9; -x_5 = l_Lean_Option_register___at_Std_Format_initFn____x40_Lean_Data_Format___hyg_95____spec__1(x_2, x_3, x_4, x_1); +x_5 = l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__1(x_2, x_3, x_4, x_1); return x_5; } } @@ -707,7 +707,7 @@ x_45 = lean_ctor_get(x_44, 1); lean_inc(x_45); lean_dec(x_44); x_46 = l___private_Lean_Compiler_IR_0__Lean_IR_compileAux___lambda__2___closed__16; -x_47 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_6, x_46); +x_47 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_6, x_46); if (x_47 == 0) { lean_object* x_48; lean_object* x_49; @@ -895,7 +895,7 @@ x_24 = lean_ctor_get(x_23, 1); lean_inc(x_24); lean_dec(x_23); x_25 = l___private_Lean_Compiler_IR_0__Lean_IR_compileAux___lambda__2___closed__16; -x_26 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_2, x_25); +x_26 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_2, x_25); if (x_26 == 0) { lean_object* x_27; lean_object* x_28; diff --git a/stage0/stdlib/Lean/Compiler/LCNF/FVarUtil.c b/stage0/stdlib/Lean/Compiler/LCNF/FVarUtil.c index 7bfbadc83e81..ec4c32e629c8 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/FVarUtil.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/FVarUtil.c @@ -438,7 +438,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_Expr_mapFVarM___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; x_1 = l_Lean_Compiler_LCNF_Expr_mapFVarM___rarg___lambda__1___closed__1; x_2 = l_Lean_Compiler_LCNF_Expr_mapFVarM___rarg___lambda__1___closed__2; -x_3 = lean_unsigned_to_nat(1476u); +x_3 = lean_unsigned_to_nat(1480u); x_4 = lean_unsigned_to_nat(20u); x_5 = l_Lean_Compiler_LCNF_Expr_mapFVarM___rarg___lambda__1___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -517,7 +517,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_Expr_mapFVarM___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; x_1 = l_Lean_Compiler_LCNF_Expr_mapFVarM___rarg___lambda__1___closed__1; x_2 = l_Lean_Compiler_LCNF_Expr_mapFVarM___rarg___lambda__2___closed__1; -x_3 = lean_unsigned_to_nat(1465u); +x_3 = lean_unsigned_to_nat(1469u); x_4 = lean_unsigned_to_nat(18u); x_5 = l_Lean_Compiler_LCNF_Expr_mapFVarM___rarg___lambda__2___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -633,7 +633,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_Expr_mapFVarM___rarg___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; x_1 = l_Lean_Compiler_LCNF_Expr_mapFVarM___rarg___lambda__1___closed__1; x_2 = l_Lean_Compiler_LCNF_Expr_mapFVarM___rarg___lambda__4___closed__1; -x_3 = lean_unsigned_to_nat(1560u); +x_3 = lean_unsigned_to_nat(1564u); x_4 = lean_unsigned_to_nat(20u); x_5 = l_Lean_Compiler_LCNF_Expr_mapFVarM___rarg___lambda__4___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -771,7 +771,7 @@ static lean_object* _init_l_Lean_Compiler_LCNF_Expr_mapFVarM___rarg___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; x_1 = l_Lean_Compiler_LCNF_Expr_mapFVarM___rarg___lambda__1___closed__1; x_2 = l_Lean_Compiler_LCNF_Expr_mapFVarM___rarg___lambda__6___closed__1; -x_3 = lean_unsigned_to_nat(1540u); +x_3 = lean_unsigned_to_nat(1544u); x_4 = lean_unsigned_to_nat(24u); x_5 = l_Lean_Compiler_LCNF_Expr_mapFVarM___rarg___lambda__6___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); diff --git a/stage0/stdlib/Lean/Compiler/LCNF/Main.c b/stage0/stdlib/Lean/Compiler/LCNF/Main.c index 44c631a813ed..99575740ff85 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/Main.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/Main.c @@ -13,202 +13,267 @@ #ifdef __cplusplus extern "C" { #endif -LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Compiler_LCNF_PassManager_run___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___spec__1(size_t, size_t, lean_object*); static lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___closed__19; uint8_t lean_is_matcher(lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__11; -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__4___closed__3; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__17; +LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Compiler_LCNF_PassManager_run___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__4; lean_object* l_Lean_Compiler_LCNF_toDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_main___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__21; +extern lean_object* l_Lean_profiler; 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_Compiler_LCNF_PassManager_run___lambda__3___boxed(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_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__7; static lean_object* l_Lean_Compiler_LCNF_compile___closed__1; +static lean_object* l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__3___closed__4; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__5; +lean_object* l_Lean_PersistentArray_toArray___rarg(lean_object*); lean_object* l_Lean_ConstantInfo_type(lean_object*); -lean_object* l_Lean_profileitM___at_Lean_addDecl___spec__7___rarg(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_getResetTraces___at_Lean_Compiler_LCNF_PassManager_run___spec__3___rarg(lean_object*, lean_object*); lean_object* l_Lean_Meta_isProp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +double lean_float_div(double, double); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___closed__28; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__18; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__2; -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_PassManager_run___spec__5(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__4(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_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +double l_Lean_trace_profiler_threshold_getSecs(lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__10___closed__3; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___closed__2; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__16; +lean_object* l_Lean_MessageData_ofList(lean_object*); +static lean_object* l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__4___closed__1; lean_object* l_Lean_PersistentArray_push___rarg(lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__3; lean_object* l_Lean_instBEqLocalInstance___boxed(lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__1___closed__1; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__4; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__20; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__17; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___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*); uint8_t l_Lean_Option_get___at_Lean_Compiler_LCNF_toConfigOptions___spec__2(lean_object*, lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); lean_object* l_Lean_Compiler_LCNF_checkDeadLocalDecls(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_PassManager_run___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__9; +lean_object* l_Lean_replaceRef(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_PassManager_run___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_Compiler_LCNF_PassManager_run(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_PassManager_run___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__7; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__12; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__20; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__18; +uint8_t lean_float_decLt(double, double); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Compiler_LCNF_PassManager_run___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___closed__21; uint8_t l_Lean_ConstantInfo_hasValue(lean_object*); lean_object* l_Lean_Compiler_LCNF_getPassManager___rarg(lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__8; lean_object* l_Lean_isTracingEnabledFor___at_Lean_Compiler_LCNF_UnreachableBranches_elimDead_go___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_main___lambda__1___closed__2; static lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___closed__13; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__3; static lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___closed__7; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___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_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__12(lean_object*, lean_object*, size_t, size_t, 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_Compiler_LCNF_checkpoint___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* lean_lcnf_compile_decls(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___closed__24; extern lean_object* l_Lean_Expr_instBEqExpr; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___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_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__16; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_compile(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__11; lean_object* l_instBEqProd___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__23; +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Compiler_LCNF_PassManager_run___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_Compiler_LCNF_LCtx_toLocalContext(lean_object*); -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___closed__15; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__13; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__10___closed__2; static lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___closed__20; static lean_object* l_Lean_Compiler_LCNF_showDecl___closed__1; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___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* l_Lean_addTrace___at_Lean_Compiler_LCNF_UnreachableBranches_elimDead_go___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___closed__27; lean_object* l_instHashableArray___rarg___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__1; lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__7; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__5; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__15; size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_main___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__1___boxed(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_Compiler_LCNF_main___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___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_EXPORT lean_object* l_Lean_Compiler_LCNF_main___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__1; static lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___closed__11; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__9; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__4; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__8; +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Compiler_LCNF_PassManager_run___spec__3___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___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* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__13; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__6; +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956_(lean_object*); static lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___closed__23; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_PassManager_run___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_PersistentArray_append___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_PassManager_run___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Compiler_LCNF_withPhase___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_panic___at_Lean_Compiler_LCNF_saveSpecParamInfo___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___closed__2; lean_object* l_Lean_Meta_isTypeFormerType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__6; lean_object* lean_st_ref_get(lean_object*, lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__2___closed__1; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__9; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__8; lean_object* lean_st_mk_ref(lean_object*, lean_object*); uint8_t l___private_Lean_Compiler_InlineAttrs_0__Lean_Compiler_hasInlineAttrCore(lean_object*, uint8_t, lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__2___closed__2; lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); lean_object* l_Lean_Name_append(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_PassManager_run___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__15; extern lean_object* l_Lean_Expr_instHashableExpr; +static lean_object* l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__3___closed__2; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___closed__1; +lean_object* lean_io_mono_nanos_now(lean_object*); +static lean_object* l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__3___closed__1; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__1; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__12; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_markRecDecls(lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_instHashableLocalInstance___boxed(lean_object*); lean_object* l_instHashableProd___rarg___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Compiler_LCNF_PassManager_run___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* l_Lean_Name_str___override(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_PassManager_run___lambda__2___closed__1; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_showDecl___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_isMatcher___at_Lean_Compiler_LCNF_shouldGenerateCode___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_isExtern(lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__14; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__23; lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Compiler_LCNF_PassManager_run___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___closed__16; static lean_object* l_Lean_Compiler_LCNF_compile___closed__2; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__6; static lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___closed__9; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__8; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__11; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__14; lean_object* l_Lean_KVMap_setBool(lean_object*, lean_object*, uint8_t); -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__17; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__22; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__19; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_PassManager_run___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___closed__18; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__4(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__2; +double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); extern lean_object* l_Lean_Compiler_compiler_check; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode_isCompIrrelevant(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__10___closed__1; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__15; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__4___closed__4; +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static double l_Lean_withSeconds___at_Lean_Compiler_LCNF_PassManager_run___spec__4___closed__1; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__13; static lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___closed__6; static lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___closed__17; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1(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_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___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* l_Lean_Compiler_LCNF_CompilerM_run___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Compiler_LCNF_PassManager_run___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___closed__22; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_PassManager_run___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___closed__10; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__6(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_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_float_to_string(double); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_PassManager_run___spec__11(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__19; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___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_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__9; +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Compiler_LCNF_PassManager_run___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_List_mapTR_loop___at_Lean_Compiler_LCNF_Simp_simpJpCases_x3f___spec__5(lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__3; static lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___closed__26; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___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_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___closed__3; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__7; lean_object* l_Lean_Compiler_LCNF_CompilerM_run___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___closed__8; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__19; +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__10(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_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__3___closed__3; static lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___closed__4; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__21; static lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___closed__14; +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Compiler_LCNF_PassManager_run___spec__3(lean_object*, lean_object*, lean_object*); +lean_object* lean_nat_sub(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__1; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__16; +lean_object* l_Lean_profileitM___at_Lean_addDecl___spec__14___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___closed__29; lean_object* l_Lean_Meta_InfoCacheKey_instHashableInfoCacheKey___boxed(lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__18; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__10; lean_object* l_Lean_Compiler_LCNF_Decl_check(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___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*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_getDeclInfo_x3f(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_checkpoint(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_Decl_size(lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__6; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__22; static lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___closed__3; static lean_object* l_Lean_Compiler_LCNF_showDecl___closed__2; lean_object* l_Lean_mkHashMapImp___rarg(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_PassManager_run___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__4___closed__1; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___closed__4; size_t lean_usize_add(size_t, size_t); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__12; lean_object* l_Lean_Compiler_LCNF_ppDecl_x27(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__11; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_showDecl(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__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_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__5; lean_object* lean_array_uget(lean_object*, size_t); -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__14; +extern lean_object* l_Lean_trace_profiler; lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__10; static lean_object* l_Lean_Compiler_LCNF_main___closed__1; lean_object* l_List_redLength___rarg(lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__10; +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___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_Compiler_LCNF_shouldGenerateCode___closed__25; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__2(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___closed__1; +static lean_object* l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__3___closed__5; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__10___closed__4; +LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Compiler_LCNF_PassManager_run___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__3; +static lean_object* l_Lean_Compiler_LCNF_main___lambda__1___closed__1; uint8_t lean_nat_dec_le(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, 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_Compiler_LCNF_shouldGenerateCode___closed__12; lean_object* l_Lean_Compiler_LCNF_getDeclAt_x3f(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_instBEqArray___rarg___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_compile___closed__3; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__12; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__4; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__10; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__2; +LEAN_EXPORT lean_object* l_Lean_withSeconds___at_Lean_Compiler_LCNF_PassManager_run___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__13; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Compiler_LCNF_PassManager_run___spec__3___rarg___boxed(lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofName(lean_object*); uint8_t l_Lean_isCasesOnRecursor(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___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*); static lean_object* l_Lean_Compiler_LCNF_PassManager_run___lambda__2___closed__2; static lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___closed__5; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__5; lean_object* l_Nat_repr(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797_(lean_object*); lean_object* l_List_toArrayAux___rarg(lean_object*, lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__4___closed__2; +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_isMatcher___at_Lean_Compiler_LCNF_shouldGenerateCode___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Array_isEmpty___rarg(lean_object*); static lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___lambda__2___closed__1; @@ -2285,252 +2350,1359 @@ return x_24; } } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___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_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Compiler_LCNF_PassManager_run___spec__3___rarg(lean_object* x_1, lean_object* x_2) { _start: { -uint8_t x_9; lean_object* x_10; uint8_t x_11; -x_9 = lean_ctor_get_uint8(x_1, sizeof(void*)*3); -x_10 = lean_ctor_get(x_1, 2); -lean_inc(x_10); -x_11 = !lean_is_exclusive(x_4); -if (x_11 == 0) -{ -lean_object* x_12; lean_object* x_13; -x_12 = lean_ctor_get(x_4, 0); -lean_inc(x_12); -lean_ctor_set_uint8(x_4, sizeof(void*)*1, x_9); -lean_inc(x_7); -lean_inc(x_6); +lean_object* x_3; 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_st_ref_get(x_1, x_2); +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +x_5 = lean_ctor_get(x_3, 1); lean_inc(x_5); -x_13 = lean_apply_6(x_10, x_2, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_13) == 0) -{ -lean_object* x_14; lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; -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_uint8(x_1, sizeof(void*)*3 + 1); -x_17 = lean_ctor_get(x_1, 1); -lean_inc(x_17); -lean_dec(x_1); -x_18 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_18, 0, x_12); -lean_ctor_set_uint8(x_18, sizeof(void*)*1, x_16); -lean_inc(x_14); -x_19 = l_Lean_Compiler_LCNF_checkpoint(x_17, x_14, x_18, x_5, x_6, x_7, x_15); -if (lean_obj_tag(x_19) == 0) +lean_dec(x_3); +x_6 = lean_ctor_get(x_4, 3); +lean_inc(x_6); +lean_dec(x_4); +x_7 = lean_st_ref_take(x_1, x_5); +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 = !lean_is_exclusive(x_8); +if (x_10 == 0) { -uint8_t x_20; -x_20 = !lean_is_exclusive(x_19); -if (x_20 == 0) +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_11 = lean_ctor_get(x_8, 3); +lean_dec(x_11); +x_12 = l_Lean_Compiler_LCNF_shouldGenerateCode___closed__7; +lean_ctor_set(x_8, 3, x_12); +x_13 = lean_st_ref_set(x_1, x_8, x_9); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) { -lean_object* x_21; lean_object* x_22; -x_21 = lean_ctor_get(x_19, 0); -lean_dec(x_21); -x_22 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_22, 0, x_14); -lean_ctor_set(x_19, 0, x_22); -return x_19; +lean_object* x_15; +x_15 = lean_ctor_get(x_13, 0); +lean_dec(x_15); +lean_ctor_set(x_13, 0, x_6); +return x_13; } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_23 = lean_ctor_get(x_19, 1); -lean_inc(x_23); -lean_dec(x_19); -x_24 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_24, 0, x_14); -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; -} +lean_object* x_16; lean_object* x_17; +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); +lean_dec(x_13); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_6); +lean_ctor_set(x_17, 1, x_16); +return x_17; } -else -{ -uint8_t x_26; -lean_dec(x_14); -x_26 = !lean_is_exclusive(x_19); -if (x_26 == 0) -{ -return x_19; } else { -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_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_18 = lean_ctor_get(x_8, 0); +x_19 = lean_ctor_get(x_8, 1); +x_20 = lean_ctor_get(x_8, 2); +x_21 = lean_ctor_get(x_8, 4); +x_22 = lean_ctor_get(x_8, 5); +x_23 = lean_ctor_get(x_8, 6); +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_dec(x_8); +x_24 = l_Lean_Compiler_LCNF_shouldGenerateCode___closed__7; +x_25 = lean_alloc_ctor(0, 7, 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_24); +lean_ctor_set(x_25, 4, x_21); +lean_ctor_set(x_25, 5, x_22); +lean_ctor_set(x_25, 6, x_23); +x_26 = lean_st_ref_set(x_1, x_25, x_9); +x_27 = lean_ctor_get(x_26, 1); lean_inc(x_27); -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); +if (lean_is_exclusive(x_26)) { + lean_ctor_release(x_26, 0); + lean_ctor_release(x_26, 1); + x_28 = x_26; +} else { + lean_dec_ref(x_26); + x_28 = lean_box(0); +} +if (lean_is_scalar(x_28)) { + x_29 = lean_alloc_ctor(0, 2, 0); +} else { + x_29 = x_28; +} +lean_ctor_set(x_29, 0, x_6); +lean_ctor_set(x_29, 1, x_27); return x_29; } } } -else +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Compiler_LCNF_PassManager_run___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -uint8_t x_30; -lean_dec(x_12); +lean_object* x_4; +x_4 = lean_alloc_closure((void*)(l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Compiler_LCNF_PassManager_run___spec__3___rarg___boxed), 2, 0); +return x_4; +} +} +static double _init_l_Lean_withSeconds___at_Lean_Compiler_LCNF_PassManager_run___spec__4___closed__1() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; double x_4; +x_1 = lean_unsigned_to_nat(1000000000u); +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_Lean_withSeconds___at_Lean_Compiler_LCNF_PassManager_run___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) { +_start: +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_7 = lean_io_mono_nanos_now(x_6); +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); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_1); -x_30 = !lean_is_exclusive(x_13); -if (x_30 == 0) +x_10 = lean_apply_5(x_1, x_2, x_3, x_4, x_5, x_9); +if (lean_obj_tag(x_10) == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +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_io_mono_nanos_now(x_12); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) { +lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; double x_19; double x_20; double x_21; lean_object* x_22; lean_object* x_23; +x_15 = lean_ctor_get(x_13, 0); +x_16 = lean_nat_sub(x_15, x_8); +lean_dec(x_8); +lean_dec(x_15); +x_17 = 0; +x_18 = lean_unsigned_to_nat(0u); +x_19 = l_Float_ofScientific(x_16, x_17, x_18); +lean_dec(x_16); +x_20 = l_Lean_withSeconds___at_Lean_Compiler_LCNF_PassManager_run___spec__4___closed__1; +x_21 = lean_float_div(x_19, x_20); +x_22 = lean_box_float(x_21); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_11); +lean_ctor_set(x_23, 1, x_22); +lean_ctor_set(x_13, 0, x_23); return x_13; } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_13, 0); -x_32 = lean_ctor_get(x_13, 1); -lean_inc(x_32); -lean_inc(x_31); +lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; lean_object* x_28; double x_29; double x_30; double x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_24 = lean_ctor_get(x_13, 0); +x_25 = lean_ctor_get(x_13, 1); +lean_inc(x_25); +lean_inc(x_24); lean_dec(x_13); -x_33 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_33, 0, x_31); +x_26 = lean_nat_sub(x_24, x_8); +lean_dec(x_8); +lean_dec(x_24); +x_27 = 0; +x_28 = lean_unsigned_to_nat(0u); +x_29 = l_Float_ofScientific(x_26, x_27, x_28); +lean_dec(x_26); +x_30 = l_Lean_withSeconds___at_Lean_Compiler_LCNF_PassManager_run___spec__4___closed__1; +x_31 = lean_float_div(x_29, x_30); +x_32 = lean_box_float(x_31); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_11); lean_ctor_set(x_33, 1, x_32); -return x_33; -} +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_25); +return x_34; } } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_34 = lean_ctor_get(x_4, 0); -lean_inc(x_34); -lean_dec(x_4); -lean_inc(x_34); -x_35 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_35, 0, x_34); -lean_ctor_set_uint8(x_35, sizeof(void*)*1, x_9); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_36 = lean_apply_6(x_10, x_2, x_35, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_36) == 0) +uint8_t x_35; +lean_dec(x_8); +x_35 = !lean_is_exclusive(x_10); +if (x_35 == 0) { -lean_object* x_37; lean_object* x_38; uint8_t x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; -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_uint8(x_1, sizeof(void*)*3 + 1); -x_40 = lean_ctor_get(x_1, 1); -lean_inc(x_40); -lean_dec(x_1); -x_41 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_41, 0, x_34); -lean_ctor_set_uint8(x_41, sizeof(void*)*1, x_39); +return x_10; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_10, 0); +x_37 = lean_ctor_get(x_10, 1); lean_inc(x_37); -x_42 = l_Lean_Compiler_LCNF_checkpoint(x_40, x_37, x_41, x_5, x_6, x_7, x_38); -if (lean_obj_tag(x_42) == 0) +lean_inc(x_36); +lean_dec(x_10); +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; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Compiler_LCNF_PassManager_run___spec__6(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) { +_start: { -lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_43 = lean_ctor_get(x_42, 1); -lean_inc(x_43); -if (lean_is_exclusive(x_42)) { - lean_ctor_release(x_42, 0); - lean_ctor_release(x_42, 1); - x_44 = x_42; -} else { - lean_dec_ref(x_42); - x_44 = lean_box(0); +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_11 = lean_st_ref_take(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_is_exclusive(x_12); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; size_t x_18; size_t 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_ctor_get(x_12, 3); +x_16 = l_Lean_PersistentArray_toArray___rarg(x_15); +x_17 = lean_array_get_size(x_16); +x_18 = lean_usize_of_nat(x_17); +lean_dec(x_17); +x_19 = 0; +x_20 = l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___spec__1(x_18, x_19, x_16); +x_21 = lean_alloc_ctor(9, 3, 1); +lean_ctor_set(x_21, 0, x_2); +lean_ctor_set(x_21, 1, x_4); +lean_ctor_set(x_21, 2, x_20); +lean_ctor_set_uint8(x_21, sizeof(void*)*3, x_5); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_3); +lean_ctor_set(x_22, 1, x_21); +x_23 = l_Lean_PersistentArray_push___rarg(x_1, x_22); +lean_ctor_set(x_12, 3, x_23); +x_24 = lean_st_ref_set(x_9, x_12, x_13); +x_25 = !lean_is_exclusive(x_24); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; +x_26 = lean_ctor_get(x_24, 0); +lean_dec(x_26); +x_27 = lean_box(0); +lean_ctor_set(x_24, 0, x_27); +return x_24; } -x_45 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_45, 0, x_37); -if (lean_is_scalar(x_44)) { - x_46 = lean_alloc_ctor(0, 2, 0); -} else { - x_46 = x_44; +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_24, 1); +lean_inc(x_28); +lean_dec(x_24); +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_ctor_set(x_46, 0, x_45); -lean_ctor_set(x_46, 1, x_43); -return x_46; } else { -lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; -lean_dec(x_37); -x_47 = lean_ctor_get(x_42, 0); -lean_inc(x_47); -x_48 = lean_ctor_get(x_42, 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; size_t x_40; size_t x_41; 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_31 = lean_ctor_get(x_12, 0); +x_32 = lean_ctor_get(x_12, 1); +x_33 = lean_ctor_get(x_12, 2); +x_34 = lean_ctor_get(x_12, 3); +x_35 = lean_ctor_get(x_12, 4); +x_36 = lean_ctor_get(x_12, 5); +x_37 = lean_ctor_get(x_12, 6); +lean_inc(x_37); +lean_inc(x_36); +lean_inc(x_35); +lean_inc(x_34); +lean_inc(x_33); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_12); +x_38 = l_Lean_PersistentArray_toArray___rarg(x_34); +x_39 = lean_array_get_size(x_38); +x_40 = lean_usize_of_nat(x_39); +lean_dec(x_39); +x_41 = 0; +x_42 = l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___spec__1(x_40, x_41, x_38); +x_43 = lean_alloc_ctor(9, 3, 1); +lean_ctor_set(x_43, 0, x_2); +lean_ctor_set(x_43, 1, x_4); +lean_ctor_set(x_43, 2, x_42); +lean_ctor_set_uint8(x_43, sizeof(void*)*3, x_5); +x_44 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_44, 0, x_3); +lean_ctor_set(x_44, 1, x_43); +x_45 = l_Lean_PersistentArray_push___rarg(x_1, x_44); +x_46 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_46, 0, x_31); +lean_ctor_set(x_46, 1, x_32); +lean_ctor_set(x_46, 2, x_33); +lean_ctor_set(x_46, 3, x_45); +lean_ctor_set(x_46, 4, x_35); +lean_ctor_set(x_46, 5, x_36); +lean_ctor_set(x_46, 6, x_37); +x_47 = lean_st_ref_set(x_9, x_46, x_13); +x_48 = lean_ctor_get(x_47, 1); lean_inc(x_48); -if (lean_is_exclusive(x_42)) { - lean_ctor_release(x_42, 0); - lean_ctor_release(x_42, 1); - x_49 = x_42; +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_42); + lean_dec_ref(x_47); x_49 = lean_box(0); } +x_50 = lean_box(0); if (lean_is_scalar(x_49)) { - x_50 = lean_alloc_ctor(1, 2, 0); + x_51 = lean_alloc_ctor(0, 2, 0); } else { - x_50 = x_49; + x_51 = x_49; +} +lean_ctor_set(x_51, 0, x_50); +lean_ctor_set(x_51, 1, x_48); +return x_51; } -lean_ctor_set(x_50, 0, x_47); -lean_ctor_set(x_50, 1, x_48); -return x_50; } } +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__5(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) { +_start: +{ +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_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_12 = lean_ctor_get(x_8, 2); +x_13 = lean_ctor_get(x_8, 5); +x_14 = l_Lean_replaceRef(x_3, x_13); +lean_dec(x_13); +lean_inc(x_12); +lean_ctor_set(x_8, 5, x_14); +x_15 = lean_st_ref_get(x_9, x_10); +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 = lean_st_ref_get(x_7, x_17); +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_20, 0); +lean_inc(x_22); +lean_dec(x_20); +x_23 = l_Lean_Compiler_LCNF_LCtx_toLocalContext(x_22); +x_24 = l_Lean_Compiler_LCNF_shouldGenerateCode___closed__14; +x_25 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_25, 0, x_18); +lean_ctor_set(x_25, 1, x_24); +lean_ctor_set(x_25, 2, x_23); +lean_ctor_set(x_25, 3, x_12); +x_26 = lean_alloc_ctor(3, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_4); +x_27 = l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Compiler_LCNF_PassManager_run___spec__6(x_1, x_2, x_3, x_26, x_5, x_6, x_7, x_8, x_9, x_21); +lean_dec(x_8); +return x_27; +} else { -lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; -lean_dec(x_34); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_1); -x_51 = lean_ctor_get(x_36, 0); -lean_inc(x_51); -x_52 = lean_ctor_get(x_36, 1); -lean_inc(x_52); -if (lean_is_exclusive(x_36)) { - lean_ctor_release(x_36, 0); - lean_ctor_release(x_36, 1); - x_53 = x_36; -} else { - lean_dec_ref(x_36); - x_53 = lean_box(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; +x_28 = lean_ctor_get(x_8, 0); +x_29 = lean_ctor_get(x_8, 1); +x_30 = lean_ctor_get(x_8, 2); +x_31 = lean_ctor_get(x_8, 3); +x_32 = lean_ctor_get(x_8, 4); +x_33 = lean_ctor_get(x_8, 5); +x_34 = lean_ctor_get(x_8, 6); +x_35 = lean_ctor_get(x_8, 7); +x_36 = lean_ctor_get(x_8, 8); +x_37 = lean_ctor_get(x_8, 9); +x_38 = lean_ctor_get(x_8, 10); +lean_inc(x_38); +lean_inc(x_37); +lean_inc(x_36); +lean_inc(x_35); +lean_inc(x_34); +lean_inc(x_33); +lean_inc(x_32); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_8); +x_39 = l_Lean_replaceRef(x_3, x_33); +lean_dec(x_33); +lean_inc(x_30); +x_40 = lean_alloc_ctor(0, 11, 0); +lean_ctor_set(x_40, 0, x_28); +lean_ctor_set(x_40, 1, x_29); +lean_ctor_set(x_40, 2, x_30); +lean_ctor_set(x_40, 3, x_31); +lean_ctor_set(x_40, 4, x_32); +lean_ctor_set(x_40, 5, x_39); +lean_ctor_set(x_40, 6, x_34); +lean_ctor_set(x_40, 7, x_35); +lean_ctor_set(x_40, 8, x_36); +lean_ctor_set(x_40, 9, x_37); +lean_ctor_set(x_40, 10, x_38); +x_41 = lean_st_ref_get(x_9, x_10); +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_ctor_get(x_42, 0); +lean_inc(x_44); +lean_dec(x_42); +x_45 = lean_st_ref_get(x_7, x_43); +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_ctor_get(x_46, 0); +lean_inc(x_48); +lean_dec(x_46); +x_49 = l_Lean_Compiler_LCNF_LCtx_toLocalContext(x_48); +x_50 = l_Lean_Compiler_LCNF_shouldGenerateCode___closed__14; +x_51 = lean_alloc_ctor(0, 4, 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); +lean_ctor_set(x_51, 3, x_30); +x_52 = lean_alloc_ctor(3, 2, 0); +lean_ctor_set(x_52, 0, x_51); +lean_ctor_set(x_52, 1, x_4); +x_53 = l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Compiler_LCNF_PassManager_run___spec__6(x_1, x_2, x_3, x_52, x_5, x_6, x_7, x_40, x_9, x_47); +lean_dec(x_40); +return x_53; +} } -if (lean_is_scalar(x_53)) { - x_54 = lean_alloc_ctor(1, 2, 0); -} else { - x_54 = x_53; } -lean_ctor_set(x_54, 0, x_51); -lean_ctor_set(x_54, 1, x_52); -return x_54; +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Compiler_LCNF_PassManager_run___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: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_7; lean_object* x_8; +x_7 = lean_ctor_get(x_1, 0); +lean_inc(x_7); +x_8 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_8, 0, x_7); +lean_ctor_set(x_8, 1, x_6); +return x_8; } +else +{ +lean_object* x_9; lean_object* x_10; +x_9 = lean_ctor_get(x_1, 0); +lean_inc(x_9); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_6); +return x_10; } } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__2___closed__1() { +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___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) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Running pass: ", 14); +lean_object* x_7; +x_7 = lean_apply_5(x_1, x_2, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_7) == 0) +{ +uint8_t x_8; +x_8 = !lean_is_exclusive(x_7); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; +x_9 = lean_ctor_get(x_7, 0); +x_10 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_7, 0, x_10); +return x_7; +} +else +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_11 = lean_ctor_get(x_7, 0); +x_12 = lean_ctor_get(x_7, 1); +lean_inc(x_12); +lean_inc(x_11); +lean_dec(x_7); +x_13 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_13, 0, x_11); +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; +} +} +else +{ +uint8_t x_15; +x_15 = !lean_is_exclusive(x_7); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; +x_16 = lean_ctor_get(x_7, 0); +x_17 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set_tag(x_7, 0); +lean_ctor_set(x_7, 0, x_17); +return x_7; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_18 = lean_ctor_get(x_7, 0); +x_19 = lean_ctor_get(x_7, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_7); +x_20 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_20, 0, 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_19); +return x_21; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___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) { +_start: +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; +lean_inc(x_10); +x_13 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__5(x_1, x_2, x_3, x_6, x_4, x_8, x_9, x_10, x_11, x_12); +x_14 = lean_ctor_get(x_13, 1); +lean_inc(x_14); +lean_dec(x_13); +x_15 = l_MonadExcept_ofExcept___at_Lean_Compiler_LCNF_PassManager_run___spec__7(x_5, x_8, x_9, x_10, x_11, x_14); +lean_dec(x_10); +return x_15; +} +} +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_profiler; +return x_1; +} +} +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__3___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("[", 1); +return x_1; +} +} +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__3___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__3___closed__2; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__3___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("s] ", 3); +return x_1; +} +} +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__3___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__3___closed__4; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, uint8_t x_7, double 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; lean_object* x_16; +lean_dec(x_9); +x_15 = lean_ctor_get(x_12, 5); +lean_inc(x_15); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_2); +x_16 = lean_apply_6(x_1, x_2, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_16) == 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); +x_19 = l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__3___closed__1; +x_20 = l_Lean_Option_get___at_Lean_Compiler_LCNF_toConfigOptions___spec__2(x_6, x_19); +lean_dec(x_6); +if (x_20 == 0) +{ +if (x_7 == 0) +{ +lean_object* x_21; lean_object* x_22; +x_21 = lean_box(0); +x_22 = l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__2(x_3, x_4, x_15, x_5, x_2, x_17, x_21, x_10, x_11, x_12, x_13, x_18); +lean_dec(x_13); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_2); +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; +x_23 = lean_float_to_string(x_8); +x_24 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_24, 0, x_23); +x_25 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_25, 0, x_24); +x_26 = l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__3___closed__3; +x_27 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_25); +x_28 = l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__3___closed__5; +x_29 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +x_30 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_17); +x_31 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__13; +x_32 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +x_33 = lean_box(0); +x_34 = l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__2(x_3, x_4, x_15, x_5, x_2, x_32, x_33, x_10, x_11, x_12, x_13, x_18); +lean_dec(x_13); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_2); +return x_34; +} +} +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; +x_35 = lean_float_to_string(x_8); +x_36 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_36, 0, x_35); +x_37 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_37, 0, x_36); +x_38 = l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__3___closed__3; +x_39 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_37); +x_40 = l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__3___closed__5; +x_41 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_40); +x_42 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_17); +x_43 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__13; +x_44 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_44, 0, x_42); +lean_ctor_set(x_44, 1, x_43); +x_45 = lean_box(0); +x_46 = l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__2(x_3, x_4, x_15, x_5, x_2, x_44, x_45, x_10, x_11, x_12, x_13, x_18); +lean_dec(x_13); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_2); +return x_46; +} +} +else +{ +uint8_t x_47; +lean_dec(x_15); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_47 = !lean_is_exclusive(x_16); +if (x_47 == 0) +{ +return x_16; +} +else +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_16, 0); +x_49 = lean_ctor_get(x_16, 1); +lean_inc(x_49); +lean_inc(x_48); +lean_dec(x_16); +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; +} +} +} +} +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__4___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_trace_profiler; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t 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, 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_dec(x_7); +x_13 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Compiler_LCNF_PassManager_run___spec__3___rarg(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_alloc_closure((void*)(l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__1), 6, 1); +lean_closure_set(x_16, 0, x_1); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_17 = l_Lean_withSeconds___at_Lean_Compiler_LCNF_PassManager_run___spec__4(x_16, x_8, x_9, x_10, x_11, x_15); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; lean_object* x_67; uint8_t x_68; +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_ctor_get(x_18, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_18, 1); +lean_inc(x_21); +lean_dec(x_18); +x_67 = l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__4___closed__1; +x_68 = l_Lean_Option_get___at_Lean_Compiler_LCNF_toConfigOptions___spec__2(x_5, x_67); +if (x_68 == 0) +{ +uint8_t x_69; +x_69 = 0; +x_22 = x_69; +goto block_66; +} +else +{ +double x_70; double x_71; uint8_t x_72; +x_70 = l_Lean_trace_profiler_threshold_getSecs(x_5); +x_71 = lean_unbox_float(x_21); +x_72 = lean_float_decLt(x_70, x_71); +x_22 = x_72; +goto block_66; +} +block_66: +{ +if (x_6 == 0) +{ +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; +lean_dec(x_21); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +x_23 = lean_st_ref_take(x_11, 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_is_exclusive(x_24); +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_24, 3); +x_28 = l_Lean_PersistentArray_append___rarg(x_14, x_27); +lean_ctor_set(x_24, 3, x_28); +x_29 = lean_st_ref_set(x_11, x_24, x_25); +x_30 = lean_ctor_get(x_29, 1); +lean_inc(x_30); +lean_dec(x_29); +x_31 = l_MonadExcept_ofExcept___at_Lean_Compiler_LCNF_PassManager_run___spec__7(x_20, x_8, x_9, x_10, x_11, x_30); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_20); +if (lean_obj_tag(x_31) == 0) +{ +uint8_t x_32; +x_32 = !lean_is_exclusive(x_31); +if (x_32 == 0) +{ +return x_31; +} +else +{ +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(0, 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_36; +x_36 = !lean_is_exclusive(x_31); +if (x_36 == 0) +{ +return x_31; +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_31, 0); +x_38 = lean_ctor_get(x_31, 1); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_31); +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_40; lean_object* x_41; 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_40 = lean_ctor_get(x_24, 0); +x_41 = lean_ctor_get(x_24, 1); +x_42 = lean_ctor_get(x_24, 2); +x_43 = lean_ctor_get(x_24, 3); +x_44 = lean_ctor_get(x_24, 4); +x_45 = lean_ctor_get(x_24, 5); +x_46 = lean_ctor_get(x_24, 6); +lean_inc(x_46); +lean_inc(x_45); +lean_inc(x_44); +lean_inc(x_43); +lean_inc(x_42); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_24); +x_47 = l_Lean_PersistentArray_append___rarg(x_14, x_43); +x_48 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_48, 0, x_40); +lean_ctor_set(x_48, 1, x_41); +lean_ctor_set(x_48, 2, x_42); +lean_ctor_set(x_48, 3, x_47); +lean_ctor_set(x_48, 4, x_44); +lean_ctor_set(x_48, 5, x_45); +lean_ctor_set(x_48, 6, x_46); +x_49 = lean_st_ref_set(x_11, x_48, x_25); +x_50 = lean_ctor_get(x_49, 1); +lean_inc(x_50); +lean_dec(x_49); +x_51 = l_MonadExcept_ofExcept___at_Lean_Compiler_LCNF_PassManager_run___spec__7(x_20, x_8, x_9, x_10, x_11, x_50); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_20); +if (lean_obj_tag(x_51) == 0) +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_52 = lean_ctor_get(x_51, 0); +lean_inc(x_52); +x_53 = lean_ctor_get(x_51, 1); +lean_inc(x_53); +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_51); + x_54 = lean_box(0); +} +if (lean_is_scalar(x_54)) { + x_55 = lean_alloc_ctor(0, 2, 0); +} else { + x_55 = x_54; +} +lean_ctor_set(x_55, 0, x_52); +lean_ctor_set(x_55, 1, x_53); +return x_55; +} +else +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_56 = lean_ctor_get(x_51, 0); +lean_inc(x_56); +x_57 = lean_ctor_get(x_51, 1); +lean_inc(x_57); +if (lean_is_exclusive(x_51)) { + lean_ctor_release(x_51, 0); + lean_ctor_release(x_51, 1); + x_58 = x_51; +} else { + lean_dec_ref(x_51); + x_58 = lean_box(0); +} +if (lean_is_scalar(x_58)) { + x_59 = lean_alloc_ctor(1, 2, 0); +} else { + x_59 = x_58; +} +lean_ctor_set(x_59, 0, x_56); +lean_ctor_set(x_59, 1, x_57); +return x_59; +} +} +} +else +{ +lean_object* x_60; double x_61; lean_object* x_62; +x_60 = lean_box(0); +x_61 = lean_unbox_float(x_21); +lean_dec(x_21); +x_62 = l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__3(x_2, x_20, x_14, x_3, x_4, x_5, x_22, x_61, x_60, x_8, x_9, x_10, x_11, x_19); +return x_62; +} +} +else +{ +lean_object* x_63; double x_64; lean_object* x_65; +x_63 = lean_box(0); +x_64 = lean_unbox_float(x_21); +lean_dec(x_21); +x_65 = l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__3(x_2, x_20, x_14, x_3, x_4, x_5, x_22, x_64, x_63, x_8, x_9, x_10, x_11, x_19); +return x_65; +} +} +} +else +{ +uint8_t x_73; +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +x_73 = !lean_is_exclusive(x_17); +if (x_73 == 0) +{ +return x_17; +} +else +{ +lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_74 = lean_ctor_get(x_17, 0); +x_75 = lean_ctor_get(x_17, 1); +lean_inc(x_75); +lean_inc(x_74); +lean_dec(x_17); +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_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__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) { +_start: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_10 = lean_ctor_get(x_7, 2); +lean_inc(x_10); +lean_inc(x_1); +x_11 = l_Lean_isTracingEnabledFor___at_Lean_Compiler_LCNF_UnreachableBranches_elimDead_go___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_unbox(x_12); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); +lean_dec(x_11); +x_15 = l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__4___closed__1; +x_16 = l_Lean_Option_get___at_Lean_Compiler_LCNF_toConfigOptions___spec__2(x_10, x_15); +if (x_16 == 0) +{ +lean_object* x_17; +lean_dec(x_12); +lean_dec(x_10); +lean_dec(x_2); +lean_dec(x_1); +x_17 = lean_apply_5(x_3, x_5, x_6, x_7, x_8, x_14); +if (lean_obj_tag(x_17) == 0) +{ +uint8_t x_18; +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) +{ +return x_17; +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_17, 0); +x_20 = lean_ctor_get(x_17, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_17); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); +return x_21; +} +} +else +{ +uint8_t x_22; +x_22 = !lean_is_exclusive(x_17); +if (x_22 == 0) +{ +return x_17; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_17, 0); +x_24 = lean_ctor_get(x_17, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_17); +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_26; uint8_t x_27; lean_object* x_28; +x_26 = lean_box(0); +x_27 = lean_unbox(x_12); +lean_dec(x_12); +x_28 = l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__4(x_3, x_2, x_1, x_4, x_10, x_27, x_26, x_5, x_6, x_7, x_8, x_14); +return x_28; +} +} +else +{ +lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; +x_29 = lean_ctor_get(x_11, 1); +lean_inc(x_29); +lean_dec(x_11); +x_30 = lean_box(0); +x_31 = lean_unbox(x_12); +lean_dec(x_12); +x_32 = l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__4(x_3, x_2, x_1, x_4, x_10, x_31, x_30, x_5, x_6, x_7, x_8, x_29); +return x_32; +} +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("new compiler phase: ", 20); +return x_1; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(", pass: ", 8); +return x_1; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("base", 4); +return x_1; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__5; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___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_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__2; +x_2 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__7; +x_3 = lean_alloc_ctor(7, 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_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__8; +x_2 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__4; +x_3 = lean_alloc_ctor(7, 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_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__10() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("mono", 4); +return x_1; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__10; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___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_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__2; +x_2 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__12; +x_3 = lean_alloc_ctor(7, 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_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__13; +x_2 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__4; +x_3 = lean_alloc_ctor(7, 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_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__15() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("impure", 6); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__2___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__2___closed__1; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__15; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___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_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__18() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__2; +x_2 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__17; +x_3 = lean_alloc_ctor(7, 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_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__19() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__18; +x_2 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__4; +x_3 = lean_alloc_ctor(7, 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_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___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: +{ +uint8_t x_8; lean_object* x_9; lean_object* x_10; +x_8 = lean_ctor_get_uint8(x_1, sizeof(void*)*3); +x_9 = lean_ctor_get(x_1, 1); +lean_inc(x_9); +lean_dec(x_1); +x_10 = l_Lean_MessageData_ofName(x_9); +switch (x_8) { +case 0: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_11 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__9; +x_12 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_10); +x_13 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__13; +x_14 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_14, 0, x_12); +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_7); +return x_15; +} +case 1: +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_16 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__14; +x_17 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_10); +x_18 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__13; +x_19 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_19, 0, x_17); +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_7); +return x_20; +} +default: +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_21 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__19; +x_22 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_10); +x_23 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__13; +x_24 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_24, 0, x_22); +lean_ctor_set(x_24, 1, x_23); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_7); +return x_25; +} +} } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___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_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8(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) { _start: { uint8_t x_10; @@ -2549,225 +3721,120 @@ 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; uint8_t 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; lean_object* x_21; x_12 = lean_array_uget(x_1, x_3); -x_13 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__2; -x_14 = l_Lean_isTracingEnabledFor___at_Lean_Compiler_LCNF_UnreachableBranches_elimDead_go___spec__1(x_13, x_5, x_6, x_7, x_8, x_9); -x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_12); +x_13 = lean_alloc_closure((void*)(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___boxed), 7, 1); +lean_closure_set(x_13, 0, x_12); +x_14 = lean_ctor_get_uint8(x_12, sizeof(void*)*3); +x_15 = lean_ctor_get(x_12, 2); lean_inc(x_15); -x_16 = lean_unbox(x_15); -lean_dec(x_15); -if (x_16 == 0) -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_17 = lean_ctor_get(x_14, 1); -lean_inc(x_17); -lean_dec(x_14); -x_18 = lean_box(0); +x_16 = lean_apply_1(x_15, x_4); +x_17 = lean_box(x_14); +x_18 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_withPhase___rarg___boxed), 7, 2); +lean_closure_set(x_18, 0, x_17); +lean_closure_set(x_18, 1, x_16); +x_19 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__2; +x_20 = 1; lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_19 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__1(x_12, x_4, x_18, x_5, x_6, x_7, x_8, x_17); -if (lean_obj_tag(x_19) == 0) -{ -lean_object* x_20; -x_20 = lean_ctor_get(x_19, 0); -lean_inc(x_20); -if (lean_obj_tag(x_20) == 0) -{ -uint8_t x_21; -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -x_21 = !lean_is_exclusive(x_19); -if (x_21 == 0) +x_21 = l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2(x_19, x_13, x_18, x_20, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_21) == 0) { -lean_object* x_22; lean_object* x_23; -x_22 = lean_ctor_get(x_19, 0); -lean_dec(x_22); -x_23 = lean_ctor_get(x_20, 0); +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; +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_20); -lean_ctor_set(x_19, 0, x_23); -return x_19; -} -else -{ -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_19, 1); -lean_inc(x_24); -lean_dec(x_19); -x_25 = lean_ctor_get(x_20, 0); +lean_dec(x_21); +x_24 = lean_ctor_get_uint8(x_12, sizeof(void*)*3 + 1); +x_25 = lean_ctor_get(x_12, 1); lean_inc(x_25); -lean_dec(x_20); -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_24); -return x_26; -} -} -else +lean_dec(x_12); +x_26 = lean_ctor_get(x_5, 0); +lean_inc(x_26); +x_27 = lean_alloc_ctor(0, 1, 1); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set_uint8(x_27, sizeof(void*)*1, x_24); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_22); +x_28 = l_Lean_Compiler_LCNF_checkpoint(x_25, x_22, x_27, x_6, x_7, x_8, x_23); +if (lean_obj_tag(x_28) == 0) { -lean_object* x_27; lean_object* x_28; size_t x_29; size_t x_30; -x_27 = lean_ctor_get(x_19, 1); -lean_inc(x_27); -lean_dec(x_19); -x_28 = lean_ctor_get(x_20, 0); -lean_inc(x_28); -lean_dec(x_20); -x_29 = 1; -x_30 = lean_usize_add(x_3, x_29); -x_3 = x_30; -x_4 = x_28; -x_9 = x_27; +lean_object* x_29; size_t x_30; size_t x_31; +x_29 = lean_ctor_get(x_28, 1); +lean_inc(x_29); +lean_dec(x_28); +x_30 = 1; +x_31 = lean_usize_add(x_3, x_30); +x_3 = x_31; +x_4 = x_22; +x_9 = x_29; goto _start; } -} else { -uint8_t x_32; +uint8_t x_33; +lean_dec(x_22); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_32 = !lean_is_exclusive(x_19); -if (x_32 == 0) +x_33 = !lean_is_exclusive(x_28); +if (x_33 == 0) { -return x_19; +return x_28; } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_19, 0); -x_34 = lean_ctor_get(x_19, 1); +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_inc(x_33); -lean_dec(x_19); -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_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; 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_36 = lean_ctor_get(x_14, 1); -lean_inc(x_36); -lean_dec(x_14); -x_37 = lean_ctor_get(x_12, 1); -lean_inc(x_37); -x_38 = l_Lean_MessageData_ofName(x_37); -x_39 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__2___closed__2; -x_40 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_40, 0, x_39); -lean_ctor_set(x_40, 1, x_38); -x_41 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__13; -x_42 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_42, 0, x_40); -lean_ctor_set(x_42, 1, x_41); -x_43 = l_Lean_addTrace___at_Lean_Compiler_LCNF_UnreachableBranches_elimDead_go___spec__2(x_13, x_42, x_5, x_6, x_7, x_8, 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); -lean_dec(x_43); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_46 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__1(x_12, x_4, x_44, x_5, x_6, x_7, x_8, x_45); -lean_dec(x_44); -if (lean_obj_tag(x_46) == 0) -{ -lean_object* x_47; -x_47 = lean_ctor_get(x_46, 0); -lean_inc(x_47); -if (lean_obj_tag(x_47) == 0) -{ -uint8_t x_48; -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -x_48 = !lean_is_exclusive(x_46); -if (x_48 == 0) -{ -lean_object* x_49; lean_object* x_50; -x_49 = lean_ctor_get(x_46, 0); -lean_dec(x_49); -x_50 = lean_ctor_get(x_47, 0); -lean_inc(x_50); -lean_dec(x_47); -lean_ctor_set(x_46, 0, x_50); -return x_46; -} -else -{ -lean_object* x_51; lean_object* x_52; lean_object* x_53; -x_51 = lean_ctor_get(x_46, 1); -lean_inc(x_51); -lean_dec(x_46); -x_52 = lean_ctor_get(x_47, 0); -lean_inc(x_52); -lean_dec(x_47); -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_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; } -else -{ -lean_object* x_54; lean_object* x_55; size_t x_56; size_t x_57; -x_54 = lean_ctor_get(x_46, 1); -lean_inc(x_54); -lean_dec(x_46); -x_55 = lean_ctor_get(x_47, 0); -lean_inc(x_55); -lean_dec(x_47); -x_56 = 1; -x_57 = lean_usize_add(x_3, x_56); -x_3 = x_57; -x_4 = x_55; -x_9 = x_54; -goto _start; } } else { -uint8_t x_59; +uint8_t x_37; +lean_dec(x_12); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_59 = !lean_is_exclusive(x_46); -if (x_59 == 0) +x_37 = !lean_is_exclusive(x_21); +if (x_37 == 0) { -return x_46; +return x_21; } else { -lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_60 = lean_ctor_get(x_46, 0); -x_61 = lean_ctor_get(x_46, 1); -lean_inc(x_61); -lean_inc(x_60); -lean_dec(x_46); -x_62 = 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_38; lean_object* x_39; lean_object* x_40; +x_38 = lean_ctor_get(x_21, 0); +x_39 = lean_ctor_get(x_21, 1); +lean_inc(x_39); +lean_inc(x_38); +lean_dec(x_21); +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_EXPORT lean_object* l_Lean_addTrace___at_Lean_Compiler_LCNF_PassManager_run___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_EXPORT lean_object* l_Lean_addTrace___at_Lean_Compiler_LCNF_PassManager_run___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) { _start: { 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; @@ -2911,7 +3978,7 @@ return x_56; } } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__4___closed__1() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__10___closed__1() { _start: { lean_object* x_1; @@ -2919,7 +3986,7 @@ x_1 = lean_mk_string_from_bytes("Lean.Compiler.LCNF.Main", 23); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__4___closed__2() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__10___closed__2() { _start: { lean_object* x_1; @@ -2927,7 +3994,7 @@ x_1 = lean_mk_string_from_bytes("Lean.Compiler.LCNF.PassManager.run", 34); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__4___closed__3() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__10___closed__3() { _start: { lean_object* x_1; @@ -2935,20 +4002,20 @@ x_1 = lean_mk_string_from_bytes("unreachable code has been reached", 33); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__4___closed__4() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__10___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_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__4___closed__1; -x_2 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__4___closed__2; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__10___closed__1; +x_2 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__10___closed__2; x_3 = lean_unsigned_to_nat(81u); x_4 = lean_unsigned_to_nat(52u); -x_5 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__4___closed__3; +x_5 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__10___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_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___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_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__10(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: { uint8_t x_11; @@ -2984,7 +4051,7 @@ lean_dec(x_13); x_25 = lean_ctor_get(x_23, 1); lean_inc(x_25); lean_dec(x_23); -x_26 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__4___closed__4; +x_26 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__10___closed__4; lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); @@ -3073,7 +4140,7 @@ x_50 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_50, 0, x_48); lean_ctor_set(x_50, 1, x_49); lean_inc(x_1); -x_51 = l_Lean_addTrace___at_Lean_Compiler_LCNF_PassManager_run___spec__3(x_1, x_50, x_6, x_7, x_8, x_9, x_38); +x_51 = l_Lean_addTrace___at_Lean_Compiler_LCNF_PassManager_run___spec__9(x_1, x_50, x_6, x_7, x_8, x_9, x_38); x_52 = lean_ctor_get(x_51, 1); lean_inc(x_52); lean_dec(x_51); @@ -3127,7 +4194,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_PassManager_run___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, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_PassManager_run___spec__11(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) { _start: { uint8_t x_10; @@ -3214,7 +4281,7 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__6(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_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__12(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: { uint8_t x_11; @@ -3250,7 +4317,7 @@ lean_dec(x_13); x_25 = lean_ctor_get(x_23, 1); lean_inc(x_25); lean_dec(x_23); -x_26 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__4___closed__4; +x_26 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__10___closed__4; lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); @@ -3339,7 +4406,7 @@ x_50 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_50, 0, x_48); lean_ctor_set(x_50, 1, x_49); lean_inc(x_1); -x_51 = l_Lean_addTrace___at_Lean_Compiler_LCNF_PassManager_run___spec__3(x_1, x_50, x_6, x_7, x_8, x_9, x_38); +x_51 = l_Lean_addTrace___at_Lean_Compiler_LCNF_PassManager_run___spec__9(x_1, x_50, x_6, x_7, x_8, x_9, x_38); x_52 = lean_ctor_get(x_51, 1); lean_inc(x_52); lean_dec(x_51); @@ -3456,7 +4523,7 @@ lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_20 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__2(x_16, x_19, x_10, x_14, x_3, x_4, x_5, x_6, x_17); +x_20 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8(x_16, x_19, x_10, x_14, x_3, x_4, x_5, x_6, x_17); lean_dec(x_16); if (lean_obj_tag(x_20) == 0) { @@ -3510,7 +4577,7 @@ x_32 = lean_array_get_size(x_21); x_33 = lean_usize_of_nat(x_32); lean_dec(x_32); x_34 = lean_box(0); -x_35 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__4(x_23, x_21, x_33, x_10, x_34, x_3, x_4, x_5, x_6, x_31); +x_35 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__10(x_23, x_21, x_33, x_10, x_34, x_3, x_4, x_5, x_6, x_31); if (lean_obj_tag(x_35) == 0) { uint8_t x_36; @@ -3650,7 +4717,7 @@ lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_20 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__2(x_16, x_19, x_10, x_14, x_3, x_4, x_5, x_6, x_17); +x_20 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8(x_16, x_19, x_10, x_14, x_3, x_4, x_5, x_6, x_17); lean_dec(x_16); if (lean_obj_tag(x_20) == 0) { @@ -3704,7 +4771,7 @@ x_32 = lean_array_get_size(x_21); x_33 = lean_usize_of_nat(x_32); lean_dec(x_32); x_34 = lean_box(0); -x_35 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__6(x_23, x_21, x_33, x_10, x_34, x_3, x_4, x_5, x_6, x_31); +x_35 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__12(x_23, x_21, x_33, x_10, x_34, x_3, x_4, x_5, x_6, x_31); if (lean_obj_tag(x_35) == 0) { uint8_t x_36; @@ -3861,7 +4928,7 @@ lean_dec(x_7); x_27 = l_Lean_Compiler_LCNF_shouldGenerateCode___closed__9; lean_inc(x_5); lean_inc(x_4); -x_28 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_PassManager_run___spec__5(x_1, x_25, x_26, x_27, x_2, x_3, x_4, x_5, x_6); +x_28 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_PassManager_run___spec__11(x_1, x_25, x_26, x_27, x_2, x_3, x_4, x_5, x_6); lean_dec(x_1); if (lean_obj_tag(x_28) == 0) { @@ -3966,7 +5033,7 @@ lean_dec(x_7); x_48 = l_Lean_Compiler_LCNF_shouldGenerateCode___closed__9; lean_inc(x_5); lean_inc(x_4); -x_49 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_PassManager_run___spec__5(x_1, x_46, x_47, x_48, x_2, x_3, x_4, x_5, x_6); +x_49 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_PassManager_run___spec__11(x_1, x_46, x_47, x_48, x_2, x_3, x_4, x_5, x_6); lean_dec(x_1); if (lean_obj_tag(x_49) == 0) { @@ -4113,7 +5180,7 @@ lean_dec(x_7); x_83 = l_Lean_Compiler_LCNF_shouldGenerateCode___closed__9; lean_inc(x_5); lean_inc(x_69); -x_84 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_PassManager_run___spec__5(x_1, x_81, x_82, x_83, x_2, x_3, x_69, x_5, x_6); +x_84 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_PassManager_run___spec__11(x_1, x_81, x_82, x_83, x_2, x_3, x_69, x_5, x_6); lean_dec(x_1); if (lean_obj_tag(x_84) == 0) { @@ -4232,7 +5299,7 @@ lean_dec(x_7); x_105 = l_Lean_Compiler_LCNF_shouldGenerateCode___closed__9; lean_inc(x_5); lean_inc(x_91); -x_106 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_PassManager_run___spec__5(x_1, x_103, x_104, x_105, x_2, x_3, x_91, x_5, x_6); +x_106 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_PassManager_run___spec__11(x_1, x_103, x_104, x_105, x_2, x_3, x_91, x_5, x_6); lean_dec(x_1); if (lean_obj_tag(x_106) == 0) { @@ -4318,16 +5385,131 @@ x_11 = l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_PassManager_run___spec__1( return x_11; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___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_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Compiler_LCNF_PassManager_run___spec__3___rarg___boxed(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_9; -x_9 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_object* x_3; +x_3 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Compiler_LCNF_PassManager_run___spec__3___rarg(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Compiler_LCNF_PassManager_run___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Compiler_LCNF_PassManager_run___spec__3(x_1, x_2, x_3); lean_dec(x_3); -return x_9; +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Compiler_LCNF_PassManager_run___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) { +_start: +{ +uint8_t x_11; lean_object* x_12; +x_11 = lean_unbox(x_5); +lean_dec(x_5); +x_12 = l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Compiler_LCNF_PassManager_run___spec__6(x_1, x_2, x_3, x_4, x_11, 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); +return x_12; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Compiler_LCNF_PassManager_run___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: +{ +uint8_t x_11; lean_object* x_12; +x_11 = lean_unbox(x_5); +lean_dec(x_5); +x_12 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__5(x_1, x_2, x_3, x_4, x_11, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_7); +lean_dec(x_6); +return x_12; +} +} +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Compiler_LCNF_PassManager_run___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: +{ +lean_object* x_7; +x_7 = l_MonadExcept_ofExcept___at_Lean_Compiler_LCNF_PassManager_run___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); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___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) { +_start: +{ +uint8_t x_13; lean_object* x_14; +x_13 = lean_unbox(x_4); +lean_dec(x_4); +x_14 = l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__2(x_1, x_2, x_3, x_13, 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_5); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___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_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +uint8_t x_15; uint8_t x_16; double x_17; lean_object* x_18; +x_15 = lean_unbox(x_5); +lean_dec(x_5); +x_16 = lean_unbox(x_7); +lean_dec(x_7); +x_17 = lean_unbox_float(x_8); +lean_dec(x_8); +x_18 = l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__3(x_1, x_2, x_3, x_4, x_15, x_6, x_16, x_17, x_9, x_10, x_11, x_12, x_13, x_14); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___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, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; uint8_t x_14; lean_object* x_15; +x_13 = lean_unbox(x_4); +lean_dec(x_4); +x_14 = lean_unbox(x_6); +lean_dec(x_6); +x_15 = l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__4(x_1, x_2, x_3, x_13, x_5, x_14, x_7, x_8, x_9, x_10, x_11, x_12); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___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_10; lean_object* x_11; +x_10 = lean_unbox(x_4); +lean_dec(x_4); +x_11 = l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2(x_1, x_2, x_3, x_10, x_5, x_6, x_7, x_8, x_9); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___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_Compiler_LCNF_PassManager_run___spec__8___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_8; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___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_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___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) { _start: { size_t x_10; size_t x_11; lean_object* x_12; @@ -4335,16 +5517,16 @@ x_10 = lean_unbox_usize(x_2); lean_dec(x_2); x_11 = lean_unbox_usize(x_3); lean_dec(x_3); -x_12 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__2(x_1, x_10, x_11, x_4, x_5, x_6, x_7, x_8, x_9); +x_12 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8(x_1, x_10, x_11, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_1); return x_12; } } -LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Compiler_LCNF_PassManager_run___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_EXPORT lean_object* l_Lean_addTrace___at_Lean_Compiler_LCNF_PassManager_run___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) { _start: { lean_object* x_8; -x_8 = l_Lean_addTrace___at_Lean_Compiler_LCNF_PassManager_run___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +x_8 = l_Lean_addTrace___at_Lean_Compiler_LCNF_PassManager_run___spec__9(x_1, x_2, x_3, x_4, x_5, x_6, x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); @@ -4352,7 +5534,7 @@ lean_dec(x_3); return x_8; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___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_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___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: { size_t x_11; size_t x_12; lean_object* x_13; @@ -4360,12 +5542,12 @@ x_11 = lean_unbox_usize(x_3); lean_dec(x_3); x_12 = lean_unbox_usize(x_4); lean_dec(x_4); -x_13 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__4(x_1, x_2, x_11, x_12, x_5, x_6, x_7, x_8, x_9, x_10); +x_13 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__10(x_1, x_2, x_11, x_12, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_2); return x_13; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_PassManager_run___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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_PassManager_run___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_10; size_t x_11; lean_object* x_12; @@ -4373,14 +5555,14 @@ x_10 = lean_unbox_usize(x_2); lean_dec(x_2); x_11 = lean_unbox_usize(x_3); lean_dec(x_3); -x_12 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_PassManager_run___spec__5(x_1, x_10, x_11, x_4, x_5, x_6, x_7, x_8, x_9); +x_12 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_PassManager_run___spec__11(x_1, x_10, x_11, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_6); lean_dec(x_5); lean_dec(x_1); return x_12; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___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_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___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: { size_t x_11; size_t x_12; lean_object* x_13; @@ -4388,7 +5570,7 @@ x_11 = lean_unbox_usize(x_3); lean_dec(x_3); x_12 = lean_unbox_usize(x_4); lean_dec(x_4); -x_13 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__6(x_1, x_2, x_11, x_12, x_5, x_6, x_7, x_8, x_9, x_10); +x_13 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__12(x_1, x_2, x_11, x_12, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_2); return x_13; } @@ -4546,7 +5728,46 @@ x_7 = l_Lean_Compiler_LCNF_showDecl(x_6, x_2, x_3, x_4, x_5); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_main___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) { +static lean_object* _init_l_Lean_Compiler_LCNF_main___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("compiling new: ", 15); +return x_1; +} +} +static lean_object* _init_l_Lean_Compiler_LCNF_main___lambda__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Compiler_LCNF_main___lambda__1___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_main___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; +x_6 = lean_box(0); +x_7 = l_List_mapTR_loop___at_Lean_Compiler_LCNF_Simp_simpJpCases_x3f___spec__5(x_1, x_6); +x_8 = l_Lean_MessageData_ofList(x_7); +lean_dec(x_7); +x_9 = l_Lean_Compiler_LCNF_main___lambda__1___closed__2; +x_10 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_8); +x_11 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__13; +x_12 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_12, 0, x_10); +lean_ctor_set(x_12, 1, 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; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_main___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; @@ -4612,30 +5833,52 @@ return x_1; LEAN_EXPORT lean_object* lean_lcnf_compile_decls(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; 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; uint8_t x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; x_5 = lean_ctor_get(x_2, 2); lean_inc(x_5); -x_6 = l_List_redLength___rarg(x_1); -x_7 = lean_mk_empty_array_with_capacity(x_6); -lean_dec(x_6); -x_8 = l_List_toArrayAux___rarg(x_1, x_7); -x_9 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_main___lambda__1), 6, 1); -lean_closure_set(x_9, 0, x_8); -x_10 = l_Lean_Compiler_LCNF_compile___closed__3; -x_11 = 0; -x_12 = lean_box(x_11); -x_13 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_CompilerM_run___rarg___boxed), 6, 3); -lean_closure_set(x_13, 0, x_9); -lean_closure_set(x_13, 1, x_10); -lean_closure_set(x_13, 2, x_12); -x_14 = l_Lean_Compiler_LCNF_main___closed__1; -x_15 = lean_box(0); -x_16 = l_Lean_profileitM___at_Lean_addDecl___spec__7___rarg(x_14, x_5, x_13, x_15, x_2, x_3, x_4); +lean_inc(x_1); +x_6 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_main___lambda__1___boxed), 5, 1); +lean_closure_set(x_6, 0, x_1); +x_7 = l_List_redLength___rarg(x_1); +x_8 = lean_mk_empty_array_with_capacity(x_7); +lean_dec(x_7); +x_9 = l_List_toArrayAux___rarg(x_1, x_8); +x_10 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_main___lambda__2), 6, 1); +lean_closure_set(x_10, 0, x_9); +x_11 = l_Lean_Compiler_LCNF_compile___closed__3; +x_12 = 0; +x_13 = lean_box(x_12); +x_14 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_CompilerM_run___rarg___boxed), 6, 3); +lean_closure_set(x_14, 0, x_10); +lean_closure_set(x_14, 1, x_11); +lean_closure_set(x_14, 2, x_13); +x_15 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__2; +x_16 = 1; +x_17 = lean_box(x_16); +x_18 = lean_alloc_closure((void*)(l_Lean_withTraceNode___at_Lean_addDecl___spec__7___boxed), 7, 4); +lean_closure_set(x_18, 0, x_15); +lean_closure_set(x_18, 1, x_6); +lean_closure_set(x_18, 2, x_14); +lean_closure_set(x_18, 3, x_17); +x_19 = l_Lean_Compiler_LCNF_main___closed__1; +x_20 = lean_box(0); +x_21 = l_Lean_profileitM___at_Lean_addDecl___spec__14___rarg(x_19, x_5, x_18, x_20, x_2, x_3, x_4); lean_dec(x_5); -return x_16; +return x_21; +} +} +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_main___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_Compiler_LCNF_main___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; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__1() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__1() { _start: { lean_object* x_1; @@ -4643,17 +5886,17 @@ x_1 = lean_mk_string_from_bytes("init", 4); return x_1; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__2() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__1; -x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__1; +x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__1; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__3() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__3() { _start: { lean_object* x_1; @@ -4661,27 +5904,27 @@ x_1 = lean_mk_string_from_bytes("Lean", 4); return x_1; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__4() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__3; +x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__3; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__5() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__4; +x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__4; x_2 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__6() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__6() { _start: { lean_object* x_1; @@ -4689,17 +5932,17 @@ x_1 = lean_mk_string_from_bytes("LCNF", 4); return x_1; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__7() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__5; -x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__6; +x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__5; +x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__8() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__8() { _start: { lean_object* x_1; @@ -4707,17 +5950,17 @@ x_1 = lean_mk_string_from_bytes("initFn", 6); return x_1; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__9() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__7; -x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__8; +x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__7; +x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__8; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__10() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__10() { _start: { lean_object* x_1; @@ -4725,47 +5968,47 @@ x_1 = lean_mk_string_from_bytes("_@", 2); return x_1; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__11() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__9; -x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__10; +x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__9; +x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__10; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__12() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__11; -x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__3; +x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__11; +x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__3; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__13() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__12; +x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__12; x_2 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__14() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__13; -x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__6; +x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__13; +x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__15() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__15() { _start: { lean_object* x_1; @@ -4773,17 +6016,17 @@ x_1 = lean_mk_string_from_bytes("Main", 4); return x_1; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__16() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__14; -x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__15; +x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__14; +x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__15; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__17() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__17() { _start: { lean_object* x_1; @@ -4791,27 +6034,27 @@ x_1 = lean_mk_string_from_bytes("_hyg", 4); return x_1; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__18() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__16; -x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__17; +x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__16; +x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__17; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__19() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__18; -x_2 = lean_unsigned_to_nat(1797u); +x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__18; +x_2 = lean_unsigned_to_nat(1956u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__20() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__20() { _start: { lean_object* x_1; @@ -4819,17 +6062,17 @@ x_1 = lean_mk_string_from_bytes("test", 4); return x_1; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__21() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__1; -x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__20; +x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__20; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__22() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__22() { _start: { lean_object* x_1; @@ -4837,23 +6080,23 @@ x_1 = lean_mk_string_from_bytes("jp", 2); return x_1; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__23() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__1; -x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__22; +x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__22; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__2; +x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__2; x_3 = 1; -x_4 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__19; +x_4 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__19; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); if (lean_obj_tag(x_5) == 0) { @@ -4861,7 +6104,7 @@ lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = lean_ctor_get(x_5, 1); lean_inc(x_6); lean_dec(x_5); -x_7 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__21; +x_7 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__21; x_8 = l_Lean_registerTraceClass(x_7, x_3, x_4, x_6); if (lean_obj_tag(x_8) == 0) { @@ -4877,7 +6120,7 @@ lean_object* x_12; lean_object* x_13; uint8_t 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_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__23; +x_13 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__23; x_14 = 0; x_15 = l_Lean_registerTraceClass(x_13, x_14, x_4, x_12); return x_15; @@ -5093,18 +6336,65 @@ l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___closed__ lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___closed__3); l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___closed__4 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___closed__4(); lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___closed__4); -l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__2___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__2___closed__1(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__2___closed__1); -l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__2___closed__2 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__2___closed__2(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__2___closed__2); -l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__4___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__4___closed__1(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__4___closed__1); -l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__4___closed__2 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__4___closed__2(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__4___closed__2); -l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__4___closed__3 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__4___closed__3(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__4___closed__3); -l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__4___closed__4 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__4___closed__4(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__4___closed__4); +l_Lean_withSeconds___at_Lean_Compiler_LCNF_PassManager_run___spec__4___closed__1 = _init_l_Lean_withSeconds___at_Lean_Compiler_LCNF_PassManager_run___spec__4___closed__1(); +l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__3___closed__1 = _init_l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__3___closed__1); +l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__3___closed__2 = _init_l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__3___closed__2); +l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__3___closed__3 = _init_l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__3___closed__3(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__3___closed__3); +l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__3___closed__4 = _init_l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__3___closed__4(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__3___closed__4); +l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__3___closed__5 = _init_l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__3___closed__5(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__3___closed__5); +l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__4___closed__1 = _init_l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__4___closed__1); +l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__1(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__1); +l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__2 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__2(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__2); +l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__3 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__3(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__3); +l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__4 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__4(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__4); +l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__5 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__5(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__5); +l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__6 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__6(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__6); +l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__7 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__7(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__7); +l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__8 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__8(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__8); +l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__9 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__9(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__9); +l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__10 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__10(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__10); +l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__11 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__11(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__11); +l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__12 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__12(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__12); +l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__13 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__13(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__13); +l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__14 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__14(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__14); +l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__15 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__15(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__15); +l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__16 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__16(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__16); +l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__17 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__17(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__17); +l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__18 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__18(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__18); +l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__19 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__19(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___lambda__1___closed__19); +l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__10___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__10___closed__1(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__10___closed__1); +l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__10___closed__2 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__10___closed__2(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__10___closed__2); +l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__10___closed__3 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__10___closed__3(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__10___closed__3); +l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__10___closed__4 = _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__10___closed__4(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__10___closed__4); l_Lean_Compiler_LCNF_PassManager_run___lambda__2___closed__1 = _init_l_Lean_Compiler_LCNF_PassManager_run___lambda__2___closed__1(); lean_mark_persistent(l_Lean_Compiler_LCNF_PassManager_run___lambda__2___closed__1); l_Lean_Compiler_LCNF_PassManager_run___lambda__2___closed__2 = _init_l_Lean_Compiler_LCNF_PassManager_run___lambda__2___closed__2(); @@ -5119,55 +6409,59 @@ l_Lean_Compiler_LCNF_showDecl___closed__1 = _init_l_Lean_Compiler_LCNF_showDecl_ lean_mark_persistent(l_Lean_Compiler_LCNF_showDecl___closed__1); l_Lean_Compiler_LCNF_showDecl___closed__2 = _init_l_Lean_Compiler_LCNF_showDecl___closed__2(); lean_mark_persistent(l_Lean_Compiler_LCNF_showDecl___closed__2); +l_Lean_Compiler_LCNF_main___lambda__1___closed__1 = _init_l_Lean_Compiler_LCNF_main___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Compiler_LCNF_main___lambda__1___closed__1); +l_Lean_Compiler_LCNF_main___lambda__1___closed__2 = _init_l_Lean_Compiler_LCNF_main___lambda__1___closed__2(); +lean_mark_persistent(l_Lean_Compiler_LCNF_main___lambda__1___closed__2); l_Lean_Compiler_LCNF_main___closed__1 = _init_l_Lean_Compiler_LCNF_main___closed__1(); lean_mark_persistent(l_Lean_Compiler_LCNF_main___closed__1); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__1 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__1(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__1); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__2 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__2(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__2); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__3 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__3(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__3); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__4 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__4(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__4); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__5 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__5(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__5); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__6 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__6(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__6); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__7 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__7(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__7); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__8 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__8(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__8); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__9 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__9(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__9); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__10 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__10(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__10); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__11 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__11(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__11); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__12 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__12(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__12); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__13 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__13(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__13); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__14 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__14(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__14); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__15 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__15(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__15); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__16 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__16(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__16); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__17 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__17(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__17); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__18 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__18(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__18); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__19 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__19(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__19); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__20 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__20(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__20); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__21 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__21(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__21); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__22 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__22(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__22); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__23 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__23(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797____closed__23); -res = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1797_(lean_io_mk_world()); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__1 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__1(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__1); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__2 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__2(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__2); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__3 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__3(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__3); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__4 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__4(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__4); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__5 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__5(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__5); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__6 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__6(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__6); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__7 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__7(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__7); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__8 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__8(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__8); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__9 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__9(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__9); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__10 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__10(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__10); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__11 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__11(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__11); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__12 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__12(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__12); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__13 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__13(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__13); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__14 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__14(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__14); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__15 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__15(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__15); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__16 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__16(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__16); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__17 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__17(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__17); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__18 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__18(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__18); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__19 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__19(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__19); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__20 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__20(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__20); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__21 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__21(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__21); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__22 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__22(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__22); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__23 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__23(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956____closed__23); +res = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1956_(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/Compiler/LCNF/Simp/Main.c b/stage0/stdlib/Lean/Compiler/LCNF/Simp/Main.c index f4dc5e1b4f84..330eeb963a73 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/Simp/Main.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/Simp/Main.c @@ -13,7 +13,7 @@ #ifdef __cplusplus extern "C" { #endif -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_simp___lambda__7(lean_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_Compiler_LCNF_Simp_simp___lambda__7(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* l_Lean_Compiler_LCNF_mkAuxLetDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Simp_specializePartialApp___closed__3; lean_object* l___private_Init_Util_0__outOfBounds___rarg(lean_object*); @@ -23,13 +23,12 @@ lean_object* l___private_Lean_Compiler_LCNF_CompilerM_0__Lean_Compiler_LCNF_norm lean_object* l_Lean_Compiler_LCNF_Simp_InlineCandidateInfo_arity(lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); lean_object* l_Lean_Compiler_LCNF_Simp_isUsed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_simp___lambda__3(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_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_simp___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_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_simp___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*); uint8_t l_Lean_Compiler_LCNF_Code_isFun(lean_object*); lean_object* l_Lean_Compiler_LCNF_Simp_eraseLetDecl(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_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_Simp_withDiscrCtorImp_updateCtx___spec__5(size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_simp___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_Compiler_LCNF_Simp_inlineApp_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_Array_anyMUnsafe_any___at_Lean_Compiler_LCNF_Simp_simp___spec__3(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_Compiler_LCNF_Simp_addDefaultAlt(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -48,8 +47,8 @@ LEAN_EXPORT lean_object* l___private_Init_Data_Array_BasicAux_0__mapMonoMImp___a lean_object* l_Lean_Compiler_LCNF_replaceExprFVars(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_etaPolyApp_x3f(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*); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_simp___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_Lean_Compiler_LCNF_Simp_simp___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_Lean_Compiler_LCNF_Simp_simp___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_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_simp___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_array_fset(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_instInhabitedAltCore___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_normParam___at_Lean_Compiler_LCNF_Simp_simpFunDecl___spec__2(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -90,15 +89,15 @@ lean_object* l_Lean_Compiler_LCNF_isInductiveWithNoCtors(lean_object*, lean_obje static lean_object* l_Lean_Compiler_LCNF_Simp_inlineApp_x3f___closed__2; lean_object* l_Lean_Compiler_LCNF_Simp_markUsedFunDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Simp_inlineApp_x3f___closed__1; -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_simp___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*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_simp___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* l_Lean_Compiler_LCNF_Simp_betaReduce(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_st_ref_get(lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_Simp_withDiscrCtorImp_updateCtx(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_normArgs___at_Lean_Compiler_LCNF_Simp_simp___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_Lean_Compiler_LCNF_Simp_simp___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_Compiler_LCNF_Simp_simp___lambda__4(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* l_Lean_Compiler_LCNF_Simp_CtorInfo_getName(lean_object*); lean_object* l_Lean_Compiler_LCNF_eraseCode(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_simp___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_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_simp___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* l_Lean_Compiler_LCNF_AltCore_getCode(lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_inlineApp_x3f___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_Compiler_LCNF_normParams___at_Lean_Compiler_LCNF_Simp_simpFunDecl___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -138,6 +137,7 @@ lean_object* lean_nat_sub(lean_object*, lean_object*); extern lean_object* l_Lean_Compiler_LCNF_instInhabitedParam; LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_normLetDecl___at_Lean_Compiler_LCNF_Simp_simp___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_Compiler_LCNF_Simp_isOnceOrMustInline(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetDecl____x40_Lean_Compiler_LCNF_Basic___hyg_1940_(lean_object*, lean_object*); lean_object* l_Array_ofSubarray___rarg(lean_object*); lean_object* l_Lean_mkHashMapImp___rarg(lean_object*); lean_object* l_Lean_Compiler_LCNF_Simp_ConstantFold_foldConstants(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -152,7 +152,8 @@ static lean_object* l_Lean_Compiler_LCNF_Simp_inlineApp_x3f___closed__3; lean_object* l_Lean_Compiler_LCNF_Simp_findCtor_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_simpFunDecl(lean_object*, 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_Compiler_LCNF_Simp_simp___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_Compiler_LCNF_Simp_simp___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_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_simp___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*); uint8_t l_Lean_Compiler_LCNF_Code_isReturnOf(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Simp_etaPolyApp_x3f___closed__1; lean_object* l_Lean_Expr_headBeta(lean_object*); @@ -161,6 +162,7 @@ lean_object* l___private_Lean_Compiler_LCNF_CompilerM_0__Lean_Compiler_LCNF_norm uint8_t lean_nat_dec_le(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_elimVar_x3f___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); +LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_simp___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_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Simp_markUsedLetValue___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_object* l_Lean_Compiler_LCNF_mkAuxParam(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_Simp_betaReduce___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*); @@ -10430,7 +10432,94 @@ return x_142; } } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_simp___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_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_simp___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) { +_start: +{ +lean_object* x_14; lean_object* x_15; +lean_dec(x_5); +x_14 = lean_ctor_get(x_4, 3); +lean_inc(x_14); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_15 = l_Lean_Compiler_LCNF_Simp_simpValue_x3f(x_14, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_18 = lean_box(0); +x_19 = l_Lean_Compiler_LCNF_Simp_simp___lambda__1(x_1, x_2, x_3, x_4, x_18, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_17); +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; +x_20 = lean_ctor_get(x_15, 1); +lean_inc(x_20); +lean_dec(x_15); +x_21 = lean_ctor_get(x_16, 0); +lean_inc(x_21); +lean_dec(x_16); +x_22 = l_Lean_Compiler_LCNF_Simp_markSimplified___rarg(x_7, x_8, x_9, x_10, x_11, x_12, x_20); +x_23 = lean_ctor_get(x_22, 1); +lean_inc(x_23); +lean_dec(x_22); +x_24 = l_Lean_Compiler_LCNF_LetDecl_updateValue(x_4, x_21, x_9, x_10, x_11, x_12, x_23); +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_box(0); +x_28 = l_Lean_Compiler_LCNF_Simp_simp___lambda__1(x_1, x_2, x_3, x_25, x_27, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_26); +return x_28; +} +} +else +{ +uint8_t x_29; +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_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_29 = !lean_is_exclusive(x_15); +if (x_29 == 0) +{ +return x_15; +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_15, 0); +x_31 = lean_ctor_get(x_15, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_15); +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_Compiler_LCNF_Simp_simp___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) { _start: { size_t x_15; size_t x_16; uint8_t x_17; @@ -10480,7 +10569,7 @@ return x_25; } } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_simp___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_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_simp___lambda__4(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; @@ -10557,7 +10646,7 @@ x_28 = lean_ctor_get(x_19, 1); lean_inc(x_28); lean_dec(x_19); x_29 = lean_box(0); -x_30 = l_Lean_Compiler_LCNF_Simp_simp___lambda__2(x_1, x_16, x_5, x_2, x_3, x_29, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_28); +x_30 = l_Lean_Compiler_LCNF_Simp_simp___lambda__3(x_1, x_16, x_5, x_2, x_3, x_29, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_28); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -10589,7 +10678,7 @@ lean_inc(x_33); x_34 = lean_ctor_get(x_32, 1); lean_inc(x_34); lean_dec(x_32); -x_35 = l_Lean_Compiler_LCNF_Simp_simp___lambda__2(x_1, x_16, x_5, x_2, x_3, x_33, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_34); +x_35 = l_Lean_Compiler_LCNF_Simp_simp___lambda__3(x_1, x_16, x_5, x_2, x_3, x_33, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_34); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -10639,7 +10728,7 @@ return x_39; } } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_simp___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_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_simp___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) { _start: { lean_object* x_12; @@ -10695,7 +10784,7 @@ return x_20; } } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_simp___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_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_simp___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) { _start: { size_t x_15; size_t x_16; uint8_t x_17; @@ -10745,7 +10834,7 @@ return x_25; } } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_simp___lambda__6(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_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_simp___lambda__7(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; @@ -10822,7 +10911,7 @@ x_28 = lean_ctor_get(x_19, 1); lean_inc(x_28); lean_dec(x_19); x_29 = lean_box(0); -x_30 = l_Lean_Compiler_LCNF_Simp_simp___lambda__5(x_1, x_16, x_5, x_2, x_3, x_29, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_28); +x_30 = l_Lean_Compiler_LCNF_Simp_simp___lambda__6(x_1, x_16, x_5, x_2, x_3, x_29, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_28); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -10854,7 +10943,7 @@ lean_inc(x_33); x_34 = lean_ctor_get(x_32, 1); lean_inc(x_34); lean_dec(x_32); -x_35 = l_Lean_Compiler_LCNF_Simp_simp___lambda__5(x_1, x_16, x_5, x_2, x_3, x_33, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_34); +x_35 = l_Lean_Compiler_LCNF_Simp_simp___lambda__6(x_1, x_16, x_5, x_2, x_3, x_33, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_34); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -10904,7 +10993,7 @@ return x_39; } } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_simp___lambda__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_Compiler_LCNF_Simp_simp___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: { if (lean_obj_tag(x_2) == 0) @@ -11384,7 +11473,7 @@ x_36 = l_Lean_Compiler_LCNF_Simp_incVisited___rarg(x_3, x_4, x_5, x_6, x_7, x_8, switch (lean_obj_tag(x_1)) { case 0: { -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_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; uint8_t x_44; x_37 = lean_ctor_get(x_36, 1); lean_inc(x_37); lean_dec(x_36); @@ -11400,194 +11489,134 @@ lean_inc(x_42); x_43 = lean_ctor_get(x_41, 1); lean_inc(x_43); lean_dec(x_41); -x_44 = lean_ctor_get(x_42, 3); -lean_inc(x_44); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_45 = l_Lean_Compiler_LCNF_Simp_simpValue_x3f(x_44, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_43); -if (lean_obj_tag(x_45) == 0) +x_44 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetDecl____x40_Lean_Compiler_LCNF_Basic___hyg_1940_(x_38, x_42); +if (x_44 == 0) { -lean_object* x_46; +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_45 = l_Lean_Compiler_LCNF_Simp_markSimplified___rarg(x_3, x_4, x_5, x_6, x_7, x_8, x_43); x_46 = lean_ctor_get(x_45, 0); lean_inc(x_46); -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_45, 1); lean_inc(x_47); lean_dec(x_45); -x_48 = lean_box(0); -x_49 = l_Lean_Compiler_LCNF_Simp_simp___lambda__1(x_39, x_38, x_1, x_42, x_48, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_47); -return x_49; -} -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; -x_50 = lean_ctor_get(x_45, 1); -lean_inc(x_50); -lean_dec(x_45); -x_51 = lean_ctor_get(x_46, 0); -lean_inc(x_51); -lean_dec(x_46); -x_52 = l_Lean_Compiler_LCNF_Simp_markSimplified___rarg(x_3, x_4, x_5, x_6, x_7, x_8, x_50); -x_53 = lean_ctor_get(x_52, 1); -lean_inc(x_53); -lean_dec(x_52); -x_54 = l_Lean_Compiler_LCNF_LetDecl_updateValue(x_42, x_51, x_5, x_6, x_7, x_8, 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 = lean_box(0); -x_58 = l_Lean_Compiler_LCNF_Simp_simp___lambda__1(x_39, x_38, x_1, x_55, x_57, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_56); -return x_58; -} -} -else -{ -uint8_t x_59; -lean_dec(x_42); -lean_dec(x_39); -lean_dec(x_38); -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_59 = !lean_is_exclusive(x_45); -if (x_59 == 0) -{ -return x_45; +x_48 = l_Lean_Compiler_LCNF_Simp_simp___lambda__2(x_39, x_38, x_1, x_42, x_46, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_47); +return x_48; } else { -lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_60 = lean_ctor_get(x_45, 0); -x_61 = lean_ctor_get(x_45, 1); -lean_inc(x_61); -lean_inc(x_60); -lean_dec(x_45); -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_49; lean_object* x_50; +x_49 = lean_box(0); +x_50 = l_Lean_Compiler_LCNF_Simp_simp___lambda__2(x_39, x_38, x_1, x_42, x_49, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_43); +return x_50; } } case 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; uint8_t x_69; -x_63 = lean_ctor_get(x_36, 1); -lean_inc(x_63); +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; +x_51 = lean_ctor_get(x_36, 1); +lean_inc(x_51); lean_dec(x_36); -x_64 = lean_ctor_get(x_1, 0); -lean_inc(x_64); -x_65 = lean_ctor_get(x_1, 1); -lean_inc(x_65); -x_66 = lean_ctor_get(x_64, 0); -lean_inc(x_66); -x_67 = l_Lean_Compiler_LCNF_Simp_isOnceOrMustInline(x_66, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_63); -x_68 = lean_ctor_get(x_67, 0); -lean_inc(x_68); -x_69 = lean_unbox(x_68); -if (x_69 == 0) +x_52 = lean_ctor_get(x_1, 0); +lean_inc(x_52); +x_53 = lean_ctor_get(x_1, 1); +lean_inc(x_53); +x_54 = lean_ctor_get(x_52, 0); +lean_inc(x_54); +x_55 = l_Lean_Compiler_LCNF_Simp_isOnceOrMustInline(x_54, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_51); +x_56 = lean_ctor_get(x_55, 0); +lean_inc(x_56); +x_57 = lean_unbox(x_56); +if (x_57 == 0) { -lean_object* x_70; lean_object* x_71; uint8_t x_72; -x_70 = lean_ctor_get(x_67, 1); -lean_inc(x_70); -lean_dec(x_67); +lean_object* x_58; lean_object* x_59; uint8_t x_60; +x_58 = lean_ctor_get(x_55, 1); +lean_inc(x_58); +lean_dec(x_55); lean_inc(x_1); -lean_inc(x_64); -x_71 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_Simp_simp___lambda__3___boxed), 14, 4); -lean_closure_set(x_71, 0, x_65); -lean_closure_set(x_71, 1, x_64); -lean_closure_set(x_71, 2, x_1); -lean_closure_set(x_71, 3, x_68); -x_72 = l_Lean_Compiler_LCNF_Code_isFun(x_1); +lean_inc(x_52); +x_59 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_Simp_simp___lambda__4___boxed), 14, 4); +lean_closure_set(x_59, 0, x_53); +lean_closure_set(x_59, 1, x_52); +lean_closure_set(x_59, 2, x_1); +lean_closure_set(x_59, 3, x_56); +x_60 = l_Lean_Compiler_LCNF_Code_isFun(x_1); lean_dec(x_1); -if (x_72 == 0) +if (x_60 == 0) { -lean_object* x_73; lean_object* x_74; -x_73 = lean_box(0); -x_74 = l_Lean_Compiler_LCNF_Simp_simp___lambda__4(x_71, x_64, x_73, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_70); -return x_74; +lean_object* x_61; lean_object* x_62; +x_61 = lean_box(0); +x_62 = l_Lean_Compiler_LCNF_Simp_simp___lambda__5(x_59, x_52, x_61, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_58); +return x_62; } else { -lean_object* x_75; lean_object* x_76; uint8_t x_77; -x_75 = lean_ctor_get(x_64, 3); -lean_inc(x_75); -x_76 = lean_ctor_get(x_64, 2); -lean_inc(x_76); -x_77 = l_Lean_Compiler_LCNF_isEtaExpandCandidateCore(x_75, x_76); -lean_dec(x_76); -if (x_77 == 0) +lean_object* x_63; lean_object* x_64; uint8_t x_65; +x_63 = lean_ctor_get(x_52, 3); +lean_inc(x_63); +x_64 = lean_ctor_get(x_52, 2); +lean_inc(x_64); +x_65 = l_Lean_Compiler_LCNF_isEtaExpandCandidateCore(x_63, x_64); +lean_dec(x_64); +if (x_65 == 0) { -lean_object* x_78; lean_object* x_79; -x_78 = lean_box(0); -x_79 = l_Lean_Compiler_LCNF_Simp_simp___lambda__4(x_71, x_64, x_78, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_70); -return x_79; +lean_object* x_66; lean_object* x_67; +x_66 = lean_box(0); +x_67 = l_Lean_Compiler_LCNF_Simp_simp___lambda__5(x_59, x_52, x_66, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_58); +return x_67; } 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; -x_80 = lean_st_ref_get(x_3, x_70); -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 = 0; +lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; uint8_t x_72; lean_object* x_73; +x_68 = lean_st_ref_get(x_3, x_58); +x_69 = lean_ctor_get(x_68, 0); +lean_inc(x_69); +x_70 = lean_ctor_get(x_68, 1); +lean_inc(x_70); +lean_dec(x_68); +x_71 = lean_ctor_get(x_69, 0); +lean_inc(x_71); +lean_dec(x_69); +x_72 = 0; lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_85 = l_Lean_Compiler_LCNF_normFunDeclImp(x_84, x_64, x_83, x_5, x_6, x_7, x_8, x_82); -if (lean_obj_tag(x_85) == 0) +x_73 = l_Lean_Compiler_LCNF_normFunDeclImp(x_72, x_52, x_71, x_5, x_6, x_7, x_8, x_70); +if (lean_obj_tag(x_73) == 0) { -lean_object* x_86; lean_object* x_87; lean_object* x_88; -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); +lean_object* x_74; lean_object* x_75; lean_object* x_76; +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); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_88 = l_Lean_Compiler_LCNF_FunDeclCore_etaExpand(x_86, x_5, x_6, x_7, x_8, x_87); -if (lean_obj_tag(x_88) == 0) +x_76 = l_Lean_Compiler_LCNF_FunDeclCore_etaExpand(x_74, x_5, x_6, x_7, x_8, x_75); +if (lean_obj_tag(x_76) == 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; -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 = l_Lean_Compiler_LCNF_Simp_markSimplified___rarg(x_3, x_4, x_5, x_6, x_7, x_8, x_90); -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 = l_Lean_Compiler_LCNF_Simp_simp___lambda__4(x_71, x_89, x_92, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_93); -lean_dec(x_92); -return x_94; +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_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 = l_Lean_Compiler_LCNF_Simp_markSimplified___rarg(x_3, x_4, x_5, x_6, x_7, x_8, x_78); +x_80 = lean_ctor_get(x_79, 0); +lean_inc(x_80); +x_81 = lean_ctor_get(x_79, 1); +lean_inc(x_81); +lean_dec(x_79); +x_82 = l_Lean_Compiler_LCNF_Simp_simp___lambda__5(x_59, x_77, x_80, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_81); +lean_dec(x_80); +return x_82; } else { -uint8_t x_95; -lean_dec(x_71); +uint8_t x_83; +lean_dec(x_59); lean_dec(x_7); lean_dec(x_8); lean_dec(x_6); @@ -11595,30 +11624,30 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_95 = !lean_is_exclusive(x_88); -if (x_95 == 0) +x_83 = !lean_is_exclusive(x_76); +if (x_83 == 0) { -return x_88; +return x_76; } else { -lean_object* x_96; lean_object* x_97; lean_object* x_98; -x_96 = lean_ctor_get(x_88, 0); -x_97 = lean_ctor_get(x_88, 1); -lean_inc(x_97); -lean_inc(x_96); -lean_dec(x_88); -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; +lean_object* x_84; lean_object* x_85; lean_object* x_86; +x_84 = lean_ctor_get(x_76, 0); +x_85 = lean_ctor_get(x_76, 1); +lean_inc(x_85); +lean_inc(x_84); +lean_dec(x_76); +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; } } } else { -uint8_t x_99; -lean_dec(x_71); +uint8_t x_87; +lean_dec(x_59); lean_dec(x_7); lean_dec(x_8); lean_dec(x_6); @@ -11626,23 +11655,23 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_99 = !lean_is_exclusive(x_85); -if (x_99 == 0) +x_87 = !lean_is_exclusive(x_73); +if (x_87 == 0) { -return x_85; +return x_73; } else { -lean_object* x_100; lean_object* x_101; lean_object* x_102; -x_100 = lean_ctor_get(x_85, 0); -x_101 = lean_ctor_get(x_85, 1); -lean_inc(x_101); -lean_inc(x_100); -lean_dec(x_85); -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; +lean_object* x_88; lean_object* x_89; lean_object* x_90; +x_88 = lean_ctor_get(x_73, 0); +x_89 = lean_ctor_get(x_73, 1); +lean_inc(x_89); +lean_inc(x_88); +lean_dec(x_73); +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; } } } @@ -11650,46 +11679,46 @@ return x_102; } else { -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_103 = lean_ctor_get(x_67, 1); -lean_inc(x_103); -lean_dec(x_67); -x_104 = lean_st_ref_get(x_3, x_103); -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_107 = lean_ctor_get(x_105, 0); -lean_inc(x_107); -lean_dec(x_105); -x_108 = 0; +lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; uint8_t x_96; lean_object* x_97; +x_91 = lean_ctor_get(x_55, 1); +lean_inc(x_91); +lean_dec(x_55); +x_92 = lean_st_ref_get(x_3, x_91); +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, 0); +lean_inc(x_95); +lean_dec(x_93); +x_96 = 0; lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_64); -x_109 = l_Lean_Compiler_LCNF_normFunDeclImp(x_108, x_64, x_107, x_5, x_6, x_7, x_8, x_106); -if (lean_obj_tag(x_109) == 0) +lean_inc(x_52); +x_97 = l_Lean_Compiler_LCNF_normFunDeclImp(x_96, x_52, x_95, x_5, x_6, x_7, x_8, x_94); +if (lean_obj_tag(x_97) == 0) { -lean_object* x_110; lean_object* x_111; lean_object* x_112; uint8_t x_113; lean_object* x_114; -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_box(0); -x_113 = lean_unbox(x_68); -lean_dec(x_68); -x_114 = l_Lean_Compiler_LCNF_Simp_simp___lambda__3(x_65, x_64, x_1, x_113, x_110, x_112, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_111); -return x_114; +lean_object* x_98; lean_object* x_99; lean_object* x_100; uint8_t x_101; lean_object* x_102; +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); +x_100 = lean_box(0); +x_101 = lean_unbox(x_56); +lean_dec(x_56); +x_102 = l_Lean_Compiler_LCNF_Simp_simp___lambda__4(x_53, x_52, x_1, x_101, x_98, x_100, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_99); +return x_102; } else { -uint8_t x_115; -lean_dec(x_68); -lean_dec(x_65); -lean_dec(x_64); +uint8_t x_103; +lean_dec(x_56); +lean_dec(x_53); +lean_dec(x_52); lean_dec(x_7); lean_dec(x_8); lean_dec(x_6); @@ -11698,134 +11727,134 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_115 = !lean_is_exclusive(x_109); -if (x_115 == 0) +x_103 = !lean_is_exclusive(x_97); +if (x_103 == 0) { -return x_109; +return x_97; } else { -lean_object* x_116; lean_object* x_117; lean_object* x_118; -x_116 = lean_ctor_get(x_109, 0); -x_117 = lean_ctor_get(x_109, 1); -lean_inc(x_117); -lean_inc(x_116); -lean_dec(x_109); -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; -} -} -} -} -case 2: +lean_object* x_104; lean_object* x_105; lean_object* x_106; +x_104 = lean_ctor_get(x_97, 0); +x_105 = lean_ctor_get(x_97, 1); +lean_inc(x_105); +lean_inc(x_104); +lean_dec(x_97); +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; +} +} +} +} +case 2: { -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_36, 1); -lean_inc(x_119); +lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; uint8_t x_113; +x_107 = lean_ctor_get(x_36, 1); +lean_inc(x_107); lean_dec(x_36); -x_120 = lean_ctor_get(x_1, 0); -lean_inc(x_120); -x_121 = lean_ctor_get(x_1, 1); -lean_inc(x_121); -x_122 = lean_ctor_get(x_120, 0); -lean_inc(x_122); -x_123 = l_Lean_Compiler_LCNF_Simp_isOnceOrMustInline(x_122, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_119); -x_124 = lean_ctor_get(x_123, 0); -lean_inc(x_124); -x_125 = lean_unbox(x_124); -if (x_125 == 0) +x_108 = lean_ctor_get(x_1, 0); +lean_inc(x_108); +x_109 = lean_ctor_get(x_1, 1); +lean_inc(x_109); +x_110 = lean_ctor_get(x_108, 0); +lean_inc(x_110); +x_111 = l_Lean_Compiler_LCNF_Simp_isOnceOrMustInline(x_110, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_107); +x_112 = lean_ctor_get(x_111, 0); +lean_inc(x_112); +x_113 = lean_unbox(x_112); +if (x_113 == 0) { -lean_object* x_126; lean_object* x_127; uint8_t x_128; -x_126 = lean_ctor_get(x_123, 1); -lean_inc(x_126); -lean_dec(x_123); +lean_object* x_114; lean_object* x_115; uint8_t x_116; +x_114 = lean_ctor_get(x_111, 1); +lean_inc(x_114); +lean_dec(x_111); lean_inc(x_1); -lean_inc(x_120); -x_127 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_Simp_simp___lambda__6___boxed), 14, 4); -lean_closure_set(x_127, 0, x_121); -lean_closure_set(x_127, 1, x_120); -lean_closure_set(x_127, 2, x_1); -lean_closure_set(x_127, 3, x_124); -x_128 = l_Lean_Compiler_LCNF_Code_isFun(x_1); +lean_inc(x_108); +x_115 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_Simp_simp___lambda__7___boxed), 14, 4); +lean_closure_set(x_115, 0, x_109); +lean_closure_set(x_115, 1, x_108); +lean_closure_set(x_115, 2, x_1); +lean_closure_set(x_115, 3, x_112); +x_116 = l_Lean_Compiler_LCNF_Code_isFun(x_1); lean_dec(x_1); -if (x_128 == 0) +if (x_116 == 0) { -lean_object* x_129; lean_object* x_130; -x_129 = lean_box(0); -x_130 = l_Lean_Compiler_LCNF_Simp_simp___lambda__4(x_127, x_120, x_129, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_126); -return x_130; +lean_object* x_117; lean_object* x_118; +x_117 = lean_box(0); +x_118 = l_Lean_Compiler_LCNF_Simp_simp___lambda__5(x_115, x_108, x_117, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_114); +return x_118; } else { -lean_object* x_131; lean_object* x_132; uint8_t x_133; -x_131 = lean_ctor_get(x_120, 3); -lean_inc(x_131); -x_132 = lean_ctor_get(x_120, 2); -lean_inc(x_132); -x_133 = l_Lean_Compiler_LCNF_isEtaExpandCandidateCore(x_131, x_132); -lean_dec(x_132); -if (x_133 == 0) +lean_object* x_119; lean_object* x_120; uint8_t x_121; +x_119 = lean_ctor_get(x_108, 3); +lean_inc(x_119); +x_120 = lean_ctor_get(x_108, 2); +lean_inc(x_120); +x_121 = l_Lean_Compiler_LCNF_isEtaExpandCandidateCore(x_119, x_120); +lean_dec(x_120); +if (x_121 == 0) { -lean_object* x_134; lean_object* x_135; -x_134 = lean_box(0); -x_135 = l_Lean_Compiler_LCNF_Simp_simp___lambda__4(x_127, x_120, x_134, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_126); -return x_135; +lean_object* x_122; lean_object* x_123; +x_122 = lean_box(0); +x_123 = l_Lean_Compiler_LCNF_Simp_simp___lambda__5(x_115, x_108, x_122, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_114); +return x_123; } else { -lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; uint8_t x_140; lean_object* x_141; -x_136 = lean_st_ref_get(x_3, x_126); -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_137, 0); -lean_inc(x_139); -lean_dec(x_137); -x_140 = 0; +lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; uint8_t x_128; lean_object* x_129; +x_124 = lean_st_ref_get(x_3, x_114); +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_125, 0); +lean_inc(x_127); +lean_dec(x_125); +x_128 = 0; lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_141 = l_Lean_Compiler_LCNF_normFunDeclImp(x_140, x_120, x_139, x_5, x_6, x_7, x_8, x_138); -if (lean_obj_tag(x_141) == 0) +x_129 = l_Lean_Compiler_LCNF_normFunDeclImp(x_128, x_108, x_127, x_5, x_6, x_7, x_8, x_126); +if (lean_obj_tag(x_129) == 0) { -lean_object* x_142; lean_object* x_143; lean_object* x_144; -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_object* x_130; lean_object* x_131; lean_object* x_132; +x_130 = lean_ctor_get(x_129, 0); +lean_inc(x_130); +x_131 = lean_ctor_get(x_129, 1); +lean_inc(x_131); +lean_dec(x_129); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_144 = l_Lean_Compiler_LCNF_FunDeclCore_etaExpand(x_142, x_5, x_6, x_7, x_8, x_143); -if (lean_obj_tag(x_144) == 0) +x_132 = l_Lean_Compiler_LCNF_FunDeclCore_etaExpand(x_130, x_5, x_6, x_7, x_8, x_131); +if (lean_obj_tag(x_132) == 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; -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_144); -x_147 = l_Lean_Compiler_LCNF_Simp_markSimplified___rarg(x_3, x_4, x_5, x_6, x_7, x_8, x_146); -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 = l_Lean_Compiler_LCNF_Simp_simp___lambda__4(x_127, x_145, x_148, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_149); -lean_dec(x_148); -return x_150; +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); +lean_dec(x_132); +x_135 = l_Lean_Compiler_LCNF_Simp_markSimplified___rarg(x_3, x_4, x_5, x_6, x_7, x_8, x_134); +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_138 = l_Lean_Compiler_LCNF_Simp_simp___lambda__5(x_115, x_133, x_136, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_137); +lean_dec(x_136); +return x_138; } else { -uint8_t x_151; -lean_dec(x_127); +uint8_t x_139; +lean_dec(x_115); lean_dec(x_7); lean_dec(x_8); lean_dec(x_6); @@ -11833,30 +11862,30 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_151 = !lean_is_exclusive(x_144); -if (x_151 == 0) +x_139 = !lean_is_exclusive(x_132); +if (x_139 == 0) { -return x_144; +return x_132; } else { -lean_object* x_152; lean_object* x_153; lean_object* x_154; -x_152 = lean_ctor_get(x_144, 0); -x_153 = lean_ctor_get(x_144, 1); -lean_inc(x_153); -lean_inc(x_152); -lean_dec(x_144); -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_140; lean_object* x_141; lean_object* x_142; +x_140 = lean_ctor_get(x_132, 0); +x_141 = lean_ctor_get(x_132, 1); +lean_inc(x_141); +lean_inc(x_140); +lean_dec(x_132); +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_155; -lean_dec(x_127); +uint8_t x_143; +lean_dec(x_115); lean_dec(x_7); lean_dec(x_8); lean_dec(x_6); @@ -11864,23 +11893,23 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_155 = !lean_is_exclusive(x_141); -if (x_155 == 0) +x_143 = !lean_is_exclusive(x_129); +if (x_143 == 0) { -return x_141; +return x_129; } else { -lean_object* x_156; lean_object* x_157; lean_object* x_158; -x_156 = lean_ctor_get(x_141, 0); -x_157 = lean_ctor_get(x_141, 1); -lean_inc(x_157); -lean_inc(x_156); -lean_dec(x_141); -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_144; lean_object* x_145; lean_object* x_146; +x_144 = lean_ctor_get(x_129, 0); +x_145 = lean_ctor_get(x_129, 1); +lean_inc(x_145); +lean_inc(x_144); +lean_dec(x_129); +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; } } } @@ -11888,46 +11917,46 @@ return x_158; } else { -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_object* x_165; -x_159 = lean_ctor_get(x_123, 1); -lean_inc(x_159); -lean_dec(x_123); -x_160 = lean_st_ref_get(x_3, x_159); -x_161 = lean_ctor_get(x_160, 0); -lean_inc(x_161); -x_162 = lean_ctor_get(x_160, 1); -lean_inc(x_162); -lean_dec(x_160); -x_163 = lean_ctor_get(x_161, 0); -lean_inc(x_163); -lean_dec(x_161); -x_164 = 0; +lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; uint8_t x_152; lean_object* x_153; +x_147 = lean_ctor_get(x_111, 1); +lean_inc(x_147); +lean_dec(x_111); +x_148 = lean_st_ref_get(x_3, x_147); +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_ctor_get(x_149, 0); +lean_inc(x_151); +lean_dec(x_149); +x_152 = 0; lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_120); -x_165 = l_Lean_Compiler_LCNF_normFunDeclImp(x_164, x_120, x_163, x_5, x_6, x_7, x_8, x_162); -if (lean_obj_tag(x_165) == 0) +lean_inc(x_108); +x_153 = l_Lean_Compiler_LCNF_normFunDeclImp(x_152, x_108, x_151, x_5, x_6, x_7, x_8, x_150); +if (lean_obj_tag(x_153) == 0) { -lean_object* x_166; lean_object* x_167; lean_object* x_168; uint8_t x_169; lean_object* x_170; -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 = lean_box(0); -x_169 = lean_unbox(x_124); -lean_dec(x_124); -x_170 = l_Lean_Compiler_LCNF_Simp_simp___lambda__6(x_121, x_120, x_1, x_169, x_166, x_168, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_167); -return x_170; +lean_object* x_154; lean_object* x_155; lean_object* x_156; uint8_t x_157; lean_object* x_158; +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_box(0); +x_157 = lean_unbox(x_112); +lean_dec(x_112); +x_158 = l_Lean_Compiler_LCNF_Simp_simp___lambda__7(x_109, x_108, x_1, x_157, x_154, x_156, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_155); +return x_158; } else { -uint8_t x_171; -lean_dec(x_124); -lean_dec(x_121); -lean_dec(x_120); +uint8_t x_159; +lean_dec(x_112); +lean_dec(x_109); +lean_dec(x_108); lean_dec(x_7); lean_dec(x_8); lean_dec(x_6); @@ -11936,68 +11965,68 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_171 = !lean_is_exclusive(x_165); -if (x_171 == 0) +x_159 = !lean_is_exclusive(x_153); +if (x_159 == 0) { -return x_165; +return x_153; } else { -lean_object* x_172; lean_object* x_173; lean_object* x_174; -x_172 = lean_ctor_get(x_165, 0); -x_173 = lean_ctor_get(x_165, 1); -lean_inc(x_173); -lean_inc(x_172); -lean_dec(x_165); -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_160; lean_object* x_161; lean_object* x_162; +x_160 = lean_ctor_get(x_153, 0); +x_161 = lean_ctor_get(x_153, 1); +lean_inc(x_161); +lean_inc(x_160); +lean_dec(x_153); +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; } } } } case 3: { -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; uint8_t x_182; lean_object* x_183; -x_175 = lean_ctor_get(x_36, 1); -lean_inc(x_175); +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; uint8_t x_170; lean_object* x_171; +x_163 = lean_ctor_get(x_36, 1); +lean_inc(x_163); lean_dec(x_36); -x_176 = lean_ctor_get(x_1, 0); -lean_inc(x_176); -x_177 = lean_ctor_get(x_1, 1); -lean_inc(x_177); -x_178 = lean_st_ref_get(x_3, x_175); -x_179 = lean_ctor_get(x_178, 0); -lean_inc(x_179); -x_180 = lean_ctor_get(x_178, 1); -lean_inc(x_180); -lean_dec(x_178); -x_181 = lean_ctor_get(x_179, 0); -lean_inc(x_181); -lean_dec(x_179); -x_182 = 0; -lean_inc(x_176); -x_183 = l___private_Lean_Compiler_LCNF_CompilerM_0__Lean_Compiler_LCNF_normFVarImp(x_181, x_176, x_182); -if (lean_obj_tag(x_183) == 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_208; -x_184 = lean_ctor_get(x_183, 0); -lean_inc(x_184); -lean_dec(x_183); -lean_inc(x_177); -x_185 = l_Lean_Compiler_LCNF_normArgs___at_Lean_Compiler_LCNF_Simp_simp___spec__2(x_182, x_177, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_180); -x_186 = lean_ctor_get(x_185, 0); -lean_inc(x_186); -x_187 = lean_ctor_get(x_185, 1); -lean_inc(x_187); -if (lean_is_exclusive(x_185)) { - lean_ctor_release(x_185, 0); - lean_ctor_release(x_185, 1); - x_188 = x_185; +x_164 = lean_ctor_get(x_1, 0); +lean_inc(x_164); +x_165 = lean_ctor_get(x_1, 1); +lean_inc(x_165); +x_166 = lean_st_ref_get(x_3, x_163); +x_167 = lean_ctor_get(x_166, 0); +lean_inc(x_167); +x_168 = lean_ctor_get(x_166, 1); +lean_inc(x_168); +lean_dec(x_166); +x_169 = lean_ctor_get(x_167, 0); +lean_inc(x_169); +lean_dec(x_167); +x_170 = 0; +lean_inc(x_164); +x_171 = l___private_Lean_Compiler_LCNF_CompilerM_0__Lean_Compiler_LCNF_normFVarImp(x_169, x_164, x_170); +if (lean_obj_tag(x_171) == 0) +{ +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_196; +x_172 = lean_ctor_get(x_171, 0); +lean_inc(x_172); +lean_dec(x_171); +lean_inc(x_165); +x_173 = l_Lean_Compiler_LCNF_normArgs___at_Lean_Compiler_LCNF_Simp_simp___spec__2(x_170, x_165, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_168); +x_174 = lean_ctor_get(x_173, 0); +lean_inc(x_174); +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_185); - x_188 = lean_box(0); + lean_dec_ref(x_173); + x_176 = lean_box(0); } lean_inc(x_8); lean_inc(x_7); @@ -12006,31 +12035,31 @@ lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -lean_inc(x_186); -lean_inc(x_184); -x_208 = l_Lean_Compiler_LCNF_Simp_inlineJp_x3f(x_184, x_186, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_187); -if (lean_obj_tag(x_208) == 0) +lean_inc(x_174); +lean_inc(x_172); +x_196 = l_Lean_Compiler_LCNF_Simp_inlineJp_x3f(x_172, x_174, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_175); +if (lean_obj_tag(x_196) == 0) { -lean_object* x_209; -x_209 = lean_ctor_get(x_208, 0); -lean_inc(x_209); -if (lean_obj_tag(x_209) == 0) +lean_object* x_197; +x_197 = lean_ctor_get(x_196, 0); +lean_inc(x_197); +if (lean_obj_tag(x_197) == 0) { -lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; uint8_t x_215; -x_210 = lean_ctor_get(x_208, 1); -lean_inc(x_210); -lean_dec(x_208); -lean_inc(x_184); -x_211 = l_Lean_Compiler_LCNF_Simp_markUsedFVar(x_184, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_210); -x_212 = lean_ctor_get(x_211, 1); -lean_inc(x_212); -lean_dec(x_211); -x_213 = lean_array_get_size(x_186); -x_214 = lean_unsigned_to_nat(0u); -x_215 = lean_nat_dec_lt(x_214, x_213); -if (x_215 == 0) +lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; uint8_t x_203; +x_198 = lean_ctor_get(x_196, 1); +lean_inc(x_198); +lean_dec(x_196); +lean_inc(x_172); +x_199 = l_Lean_Compiler_LCNF_Simp_markUsedFVar(x_172, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_198); +x_200 = lean_ctor_get(x_199, 1); +lean_inc(x_200); +lean_dec(x_199); +x_201 = lean_array_get_size(x_174); +x_202 = lean_unsigned_to_nat(0u); +x_203 = lean_nat_dec_lt(x_202, x_201); +if (x_203 == 0) { -lean_dec(x_213); +lean_dec(x_201); lean_dec(x_7); lean_dec(x_8); lean_dec(x_6); @@ -12038,16 +12067,16 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_189 = x_212; -goto block_207; +x_177 = x_200; +goto block_195; } else { -uint8_t x_216; -x_216 = lean_nat_dec_le(x_213, x_213); -if (x_216 == 0) +uint8_t x_204; +x_204 = lean_nat_dec_le(x_201, x_201); +if (x_204 == 0) { -lean_dec(x_213); +lean_dec(x_201); lean_dec(x_7); lean_dec(x_8); lean_dec(x_6); @@ -12055,17 +12084,17 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_189 = x_212; -goto block_207; +x_177 = x_200; +goto block_195; } else { -size_t x_217; size_t x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; -x_217 = 0; -x_218 = lean_usize_of_nat(x_213); -lean_dec(x_213); -x_219 = lean_box(0); -x_220 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Simp_markUsedLetValue___spec__1(x_186, x_217, x_218, x_219, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_212); +size_t x_205; size_t x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; +x_205 = 0; +x_206 = lean_usize_of_nat(x_201); +lean_dec(x_201); +x_207 = lean_box(0); +x_208 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Simp_markUsedLetValue___spec__1(x_174, x_205, x_206, x_207, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_200); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -12073,42 +12102,42 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_221 = lean_ctor_get(x_220, 1); -lean_inc(x_221); -lean_dec(x_220); -x_189 = x_221; -goto block_207; +x_209 = lean_ctor_get(x_208, 1); +lean_inc(x_209); +lean_dec(x_208); +x_177 = x_209; +goto block_195; } } } else { -lean_object* x_222; lean_object* x_223; -lean_dec(x_188); -lean_dec(x_186); -lean_dec(x_184); -lean_dec(x_177); +lean_object* x_210; lean_object* x_211; lean_dec(x_176); +lean_dec(x_174); +lean_dec(x_172); +lean_dec(x_165); +lean_dec(x_164); lean_dec(x_1); -x_222 = lean_ctor_get(x_208, 1); -lean_inc(x_222); -lean_dec(x_208); -x_223 = lean_ctor_get(x_209, 0); -lean_inc(x_223); -lean_dec(x_209); -x_1 = x_223; -x_9 = x_222; +x_210 = lean_ctor_get(x_196, 1); +lean_inc(x_210); +lean_dec(x_196); +x_211 = lean_ctor_get(x_197, 0); +lean_inc(x_211); +lean_dec(x_197); +x_1 = x_211; +x_9 = x_210; goto _start; } } else { -uint8_t x_225; -lean_dec(x_188); -lean_dec(x_186); -lean_dec(x_184); -lean_dec(x_177); +uint8_t x_213; lean_dec(x_176); +lean_dec(x_174); +lean_dec(x_172); +lean_dec(x_165); +lean_dec(x_164); lean_dec(x_7); lean_dec(x_8); lean_dec(x_6); @@ -12117,158 +12146,158 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_225 = !lean_is_exclusive(x_208); -if (x_225 == 0) +x_213 = !lean_is_exclusive(x_196); +if (x_213 == 0) { -return x_208; +return x_196; } else { -lean_object* x_226; lean_object* x_227; lean_object* x_228; -x_226 = lean_ctor_get(x_208, 0); -x_227 = lean_ctor_get(x_208, 1); -lean_inc(x_227); -lean_inc(x_226); -lean_dec(x_208); -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; +lean_object* x_214; lean_object* x_215; lean_object* x_216; +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 = 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_207: +block_195: { -uint8_t x_190; -x_190 = lean_name_eq(x_176, x_184); -lean_dec(x_176); -if (x_190 == 0) +uint8_t x_178; +x_178 = lean_name_eq(x_164, x_172); +lean_dec(x_164); +if (x_178 == 0) { -uint8_t x_191; -lean_dec(x_177); -x_191 = !lean_is_exclusive(x_1); -if (x_191 == 0) -{ -lean_object* x_192; lean_object* x_193; lean_object* x_194; -x_192 = lean_ctor_get(x_1, 1); -lean_dec(x_192); -x_193 = lean_ctor_get(x_1, 0); -lean_dec(x_193); -lean_ctor_set(x_1, 1, x_186); -lean_ctor_set(x_1, 0, x_184); -if (lean_is_scalar(x_188)) { - x_194 = lean_alloc_ctor(0, 2, 0); +uint8_t x_179; +lean_dec(x_165); +x_179 = !lean_is_exclusive(x_1); +if (x_179 == 0) +{ +lean_object* x_180; lean_object* x_181; lean_object* x_182; +x_180 = lean_ctor_get(x_1, 1); +lean_dec(x_180); +x_181 = lean_ctor_get(x_1, 0); +lean_dec(x_181); +lean_ctor_set(x_1, 1, x_174); +lean_ctor_set(x_1, 0, x_172); +if (lean_is_scalar(x_176)) { + x_182 = lean_alloc_ctor(0, 2, 0); } else { - x_194 = x_188; + x_182 = x_176; } -lean_ctor_set(x_194, 0, x_1); -lean_ctor_set(x_194, 1, x_189); -return x_194; +lean_ctor_set(x_182, 0, x_1); +lean_ctor_set(x_182, 1, x_177); +return x_182; } else { -lean_object* x_195; lean_object* x_196; +lean_object* x_183; lean_object* x_184; lean_dec(x_1); -x_195 = lean_alloc_ctor(3, 2, 0); -lean_ctor_set(x_195, 0, x_184); -lean_ctor_set(x_195, 1, x_186); -if (lean_is_scalar(x_188)) { - x_196 = lean_alloc_ctor(0, 2, 0); +x_183 = lean_alloc_ctor(3, 2, 0); +lean_ctor_set(x_183, 0, x_172); +lean_ctor_set(x_183, 1, x_174); +if (lean_is_scalar(x_176)) { + x_184 = lean_alloc_ctor(0, 2, 0); } else { - x_196 = x_188; + x_184 = x_176; } -lean_ctor_set(x_196, 0, x_195); -lean_ctor_set(x_196, 1, x_189); -return x_196; +lean_ctor_set(x_184, 0, x_183); +lean_ctor_set(x_184, 1, x_177); +return x_184; } } else { -size_t x_197; size_t x_198; uint8_t x_199; -x_197 = lean_ptr_addr(x_177); -lean_dec(x_177); -x_198 = lean_ptr_addr(x_186); -x_199 = lean_usize_dec_eq(x_197, x_198); -if (x_199 == 0) +size_t x_185; size_t x_186; uint8_t x_187; +x_185 = lean_ptr_addr(x_165); +lean_dec(x_165); +x_186 = lean_ptr_addr(x_174); +x_187 = lean_usize_dec_eq(x_185, x_186); +if (x_187 == 0) { -uint8_t x_200; -x_200 = !lean_is_exclusive(x_1); -if (x_200 == 0) +uint8_t x_188; +x_188 = !lean_is_exclusive(x_1); +if (x_188 == 0) { -lean_object* x_201; lean_object* x_202; lean_object* x_203; -x_201 = lean_ctor_get(x_1, 1); -lean_dec(x_201); -x_202 = lean_ctor_get(x_1, 0); -lean_dec(x_202); -lean_ctor_set(x_1, 1, x_186); -lean_ctor_set(x_1, 0, x_184); -if (lean_is_scalar(x_188)) { - x_203 = lean_alloc_ctor(0, 2, 0); +lean_object* x_189; lean_object* x_190; lean_object* x_191; +x_189 = lean_ctor_get(x_1, 1); +lean_dec(x_189); +x_190 = lean_ctor_get(x_1, 0); +lean_dec(x_190); +lean_ctor_set(x_1, 1, x_174); +lean_ctor_set(x_1, 0, x_172); +if (lean_is_scalar(x_176)) { + x_191 = lean_alloc_ctor(0, 2, 0); } else { - x_203 = x_188; + x_191 = x_176; } -lean_ctor_set(x_203, 0, x_1); -lean_ctor_set(x_203, 1, x_189); -return x_203; +lean_ctor_set(x_191, 0, x_1); +lean_ctor_set(x_191, 1, x_177); +return x_191; } else { -lean_object* x_204; lean_object* x_205; +lean_object* x_192; lean_object* x_193; lean_dec(x_1); -x_204 = lean_alloc_ctor(3, 2, 0); -lean_ctor_set(x_204, 0, x_184); -lean_ctor_set(x_204, 1, x_186); -if (lean_is_scalar(x_188)) { - x_205 = lean_alloc_ctor(0, 2, 0); +x_192 = lean_alloc_ctor(3, 2, 0); +lean_ctor_set(x_192, 0, x_172); +lean_ctor_set(x_192, 1, x_174); +if (lean_is_scalar(x_176)) { + x_193 = lean_alloc_ctor(0, 2, 0); } else { - x_205 = x_188; + x_193 = x_176; } -lean_ctor_set(x_205, 0, x_204); -lean_ctor_set(x_205, 1, x_189); -return x_205; +lean_ctor_set(x_193, 0, x_192); +lean_ctor_set(x_193, 1, x_177); +return x_193; } } else { -lean_object* x_206; -lean_dec(x_186); -lean_dec(x_184); -if (lean_is_scalar(x_188)) { - x_206 = lean_alloc_ctor(0, 2, 0); +lean_object* x_194; +lean_dec(x_174); +lean_dec(x_172); +if (lean_is_scalar(x_176)) { + x_194 = lean_alloc_ctor(0, 2, 0); } else { - x_206 = x_188; + x_194 = x_176; } -lean_ctor_set(x_206, 0, x_1); -lean_ctor_set(x_206, 1, x_189); -return x_206; +lean_ctor_set(x_194, 0, x_1); +lean_ctor_set(x_194, 1, x_177); +return x_194; } } } } else { -lean_object* x_229; -lean_dec(x_177); -lean_dec(x_176); +lean_object* x_217; +lean_dec(x_165); +lean_dec(x_164); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_229 = l_Lean_Compiler_LCNF_mkReturnErased(x_5, x_6, x_7, x_8, x_180); +x_217 = l_Lean_Compiler_LCNF_mkReturnErased(x_5, x_6, x_7, x_8, x_168); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -return x_229; +return x_217; } } case 4: { -lean_object* x_230; lean_object* x_231; lean_object* x_232; -x_230 = lean_ctor_get(x_36, 1); -lean_inc(x_230); +lean_object* x_218; lean_object* x_219; lean_object* x_220; +x_218 = lean_ctor_get(x_36, 1); +lean_inc(x_218); lean_dec(x_36); -x_231 = lean_ctor_get(x_1, 0); -lean_inc(x_231); +x_219 = lean_ctor_get(x_1, 0); +lean_inc(x_219); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); @@ -12276,74 +12305,74 @@ lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -lean_inc(x_231); -x_232 = l_Lean_Compiler_LCNF_Simp_simpCasesOnCtor_x3f(x_231, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_230); -if (lean_obj_tag(x_232) == 0) +lean_inc(x_219); +x_220 = l_Lean_Compiler_LCNF_Simp_simpCasesOnCtor_x3f(x_219, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_218); +if (lean_obj_tag(x_220) == 0) { -lean_object* x_233; -x_233 = lean_ctor_get(x_232, 0); -lean_inc(x_233); +lean_object* x_221; +x_221 = lean_ctor_get(x_220, 0); +lean_inc(x_221); +if (lean_obj_tag(x_221) == 0) +{ +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; uint8_t x_232; lean_object* x_233; +x_222 = lean_ctor_get(x_220, 1); +lean_inc(x_222); +lean_dec(x_220); +x_223 = lean_ctor_get(x_219, 0); +lean_inc(x_223); +x_224 = lean_ctor_get(x_219, 1); +lean_inc(x_224); +x_225 = lean_ctor_get(x_219, 2); +lean_inc(x_225); +x_226 = lean_ctor_get(x_219, 3); +lean_inc(x_226); +if (lean_is_exclusive(x_219)) { + lean_ctor_release(x_219, 0); + lean_ctor_release(x_219, 1); + lean_ctor_release(x_219, 2); + lean_ctor_release(x_219, 3); + x_227 = x_219; +} else { + lean_dec_ref(x_219); + x_227 = lean_box(0); +} +x_228 = lean_st_ref_get(x_3, x_222); +x_229 = lean_ctor_get(x_228, 0); +lean_inc(x_229); +x_230 = lean_ctor_get(x_228, 1); +lean_inc(x_230); +lean_dec(x_228); +x_231 = lean_ctor_get(x_229, 0); +lean_inc(x_231); +lean_dec(x_229); +x_232 = 0; +lean_inc(x_225); +x_233 = l___private_Lean_Compiler_LCNF_CompilerM_0__Lean_Compiler_LCNF_normFVarImp(x_231, x_225, x_232); 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; 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; -x_234 = lean_ctor_get(x_232, 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; +x_234 = lean_ctor_get(x_233, 0); lean_inc(x_234); -lean_dec(x_232); -x_235 = lean_ctor_get(x_231, 0); -lean_inc(x_235); -x_236 = lean_ctor_get(x_231, 1); +lean_dec(x_233); +x_235 = lean_st_ref_get(x_3, x_230); +x_236 = lean_ctor_get(x_235, 0); lean_inc(x_236); -x_237 = lean_ctor_get(x_231, 2); +x_237 = lean_ctor_get(x_235, 1); lean_inc(x_237); -x_238 = lean_ctor_get(x_231, 3); +lean_dec(x_235); +x_238 = lean_ctor_get(x_236, 0); lean_inc(x_238); -if (lean_is_exclusive(x_231)) { - lean_ctor_release(x_231, 0); - lean_ctor_release(x_231, 1); - lean_ctor_release(x_231, 2); - lean_ctor_release(x_231, 3); - x_239 = x_231; -} else { - lean_dec_ref(x_231); - x_239 = lean_box(0); -} -x_240 = lean_st_ref_get(x_3, x_234); -x_241 = lean_ctor_get(x_240, 0); +lean_dec(x_236); +lean_inc(x_224); +x_239 = l___private_Lean_Compiler_LCNF_CompilerM_0__Lean_Compiler_LCNF_normExprImp_go(x_238, x_232, x_224); +lean_inc(x_234); +x_240 = l_Lean_Compiler_LCNF_Simp_markUsedFVar(x_234, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_237); +x_241 = lean_ctor_get(x_240, 1); lean_inc(x_241); -x_242 = lean_ctor_get(x_240, 1); -lean_inc(x_242); lean_dec(x_240); -x_243 = lean_ctor_get(x_241, 0); -lean_inc(x_243); -lean_dec(x_241); -x_244 = 0; -lean_inc(x_237); -x_245 = l___private_Lean_Compiler_LCNF_CompilerM_0__Lean_Compiler_LCNF_normFVarImp(x_243, x_237, x_244); -if (lean_obj_tag(x_245) == 0) -{ -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; -x_246 = lean_ctor_get(x_245, 0); -lean_inc(x_246); -lean_dec(x_245); -x_247 = lean_st_ref_get(x_3, x_242); -x_248 = lean_ctor_get(x_247, 0); -lean_inc(x_248); -x_249 = lean_ctor_get(x_247, 1); -lean_inc(x_249); -lean_dec(x_247); -x_250 = lean_ctor_get(x_248, 0); -lean_inc(x_250); -lean_dec(x_248); -lean_inc(x_236); -x_251 = l___private_Lean_Compiler_LCNF_CompilerM_0__Lean_Compiler_LCNF_normExprImp_go(x_250, x_244, x_236); -lean_inc(x_246); -x_252 = l_Lean_Compiler_LCNF_Simp_markUsedFVar(x_246, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_249); -x_253 = lean_ctor_get(x_252, 1); -lean_inc(x_253); -lean_dec(x_252); -lean_inc(x_246); -x_254 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_Simp_simp___lambda__7), 10, 1); -lean_closure_set(x_254, 0, x_246); +lean_inc(x_234); +x_242 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_Simp_simp___lambda__8), 10, 1); +lean_closure_set(x_242, 0, x_234); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); @@ -12351,407 +12380,407 @@ lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -lean_inc(x_238); -x_255 = l___private_Init_Data_Array_BasicAux_0__mapMonoMImp___at_Lean_Compiler_LCNF_Simp_simp___spec__4(x_238, x_254, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_253); -if (lean_obj_tag(x_255) == 0) +lean_inc(x_226); +x_243 = l___private_Init_Data_Array_BasicAux_0__mapMonoMImp___at_Lean_Compiler_LCNF_Simp_simp___spec__4(x_226, x_242, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_241); +if (lean_obj_tag(x_243) == 0) { -lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; -x_256 = lean_ctor_get(x_255, 0); -lean_inc(x_256); -x_257 = lean_ctor_get(x_255, 1); -lean_inc(x_257); -if (lean_is_exclusive(x_255)) { - lean_ctor_release(x_255, 0); - lean_ctor_release(x_255, 1); - x_258 = x_255; +lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; +x_244 = lean_ctor_get(x_243, 0); +lean_inc(x_244); +x_245 = lean_ctor_get(x_243, 1); +lean_inc(x_245); +if (lean_is_exclusive(x_243)) { + lean_ctor_release(x_243, 0); + lean_ctor_release(x_243, 1); + x_246 = x_243; } else { - lean_dec_ref(x_255); - x_258 = lean_box(0); + lean_dec_ref(x_243); + x_246 = lean_box(0); } -x_259 = l_Lean_Compiler_LCNF_Simp_addDefaultAlt(x_256, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_257); -if (lean_obj_tag(x_259) == 0) +x_247 = l_Lean_Compiler_LCNF_Simp_addDefaultAlt(x_244, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_245); +if (lean_obj_tag(x_247) == 0) { -lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_294; lean_object* x_295; uint8_t x_306; -x_260 = lean_ctor_get(x_259, 0); -lean_inc(x_260); -x_261 = lean_ctor_get(x_259, 1); -lean_inc(x_261); -if (lean_is_exclusive(x_259)) { - lean_ctor_release(x_259, 0); - lean_ctor_release(x_259, 1); - x_262 = x_259; +lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_282; lean_object* x_283; uint8_t x_294; +x_248 = lean_ctor_get(x_247, 0); +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; } else { - lean_dec_ref(x_259); - x_262 = lean_box(0); -} -x_294 = lean_array_get_size(x_260); -x_306 = lean_nat_dec_eq(x_294, x_34); -if (x_306 == 0) -{ -lean_object* x_307; -lean_dec(x_294); -lean_dec(x_258); -x_307 = lean_box(0); -x_263 = x_307; -goto block_293; + lean_dec_ref(x_247); + x_250 = lean_box(0); +} +x_282 = lean_array_get_size(x_248); +x_294 = lean_nat_dec_eq(x_282, x_34); +if (x_294 == 0) +{ +lean_object* x_295; +lean_dec(x_282); +lean_dec(x_246); +x_295 = lean_box(0); +x_251 = x_295; +goto block_281; } else { -lean_object* x_308; uint8_t x_309; -x_308 = lean_unsigned_to_nat(0u); -x_309 = lean_nat_dec_lt(x_308, x_294); -if (x_309 == 0) +lean_object* x_296; uint8_t x_297; +x_296 = lean_unsigned_to_nat(0u); +x_297 = lean_nat_dec_lt(x_296, x_282); +if (x_297 == 0) { -lean_object* x_310; lean_object* x_311; -x_310 = l___private_Lean_Compiler_LCNF_Simp_Main_0__Lean_Compiler_LCNF_Simp_oneExitPointQuick_go___closed__1; -x_311 = l___private_Init_Util_0__outOfBounds___rarg(x_310); -if (lean_obj_tag(x_311) == 0) -{ -lean_object* x_312; -lean_dec(x_311); -lean_dec(x_294); -lean_dec(x_258); -x_312 = lean_box(0); -x_263 = x_312; -goto block_293; +lean_object* x_298; lean_object* x_299; +x_298 = l___private_Lean_Compiler_LCNF_Simp_Main_0__Lean_Compiler_LCNF_Simp_oneExitPointQuick_go___closed__1; +x_299 = l___private_Init_Util_0__outOfBounds___rarg(x_298); +if (lean_obj_tag(x_299) == 0) +{ +lean_object* x_300; +lean_dec(x_299); +lean_dec(x_282); +lean_dec(x_246); +x_300 = lean_box(0); +x_251 = x_300; +goto block_281; } else { -lean_object* x_313; -lean_dec(x_311); -lean_dec(x_262); -lean_dec(x_251); -lean_dec(x_246); +lean_object* x_301; +lean_dec(x_299); +lean_dec(x_250); lean_dec(x_239); -lean_dec(x_238); -lean_dec(x_237); -lean_dec(x_236); -lean_dec(x_235); +lean_dec(x_234); +lean_dec(x_227); +lean_dec(x_226); +lean_dec(x_225); +lean_dec(x_224); +lean_dec(x_223); lean_dec(x_1); -x_313 = lean_box(0); -x_295 = x_313; -goto block_305; +x_301 = lean_box(0); +x_283 = x_301; +goto block_293; } } else { -lean_object* x_314; -x_314 = lean_array_fget(x_260, x_308); -if (lean_obj_tag(x_314) == 0) +lean_object* x_302; +x_302 = lean_array_fget(x_248, x_296); +if (lean_obj_tag(x_302) == 0) { -lean_object* x_315; -lean_dec(x_314); -lean_dec(x_294); -lean_dec(x_258); -x_315 = lean_box(0); -x_263 = x_315; -goto block_293; +lean_object* x_303; +lean_dec(x_302); +lean_dec(x_282); +lean_dec(x_246); +x_303 = lean_box(0); +x_251 = x_303; +goto block_281; } else { -lean_object* x_316; -lean_dec(x_314); -lean_dec(x_262); -lean_dec(x_251); -lean_dec(x_246); +lean_object* x_304; +lean_dec(x_302); +lean_dec(x_250); lean_dec(x_239); -lean_dec(x_238); -lean_dec(x_237); -lean_dec(x_236); -lean_dec(x_235); +lean_dec(x_234); +lean_dec(x_227); +lean_dec(x_226); +lean_dec(x_225); +lean_dec(x_224); +lean_dec(x_223); lean_dec(x_1); -x_316 = lean_box(0); -x_295 = x_316; -goto block_305; +x_304 = lean_box(0); +x_283 = x_304; +goto block_293; } } } -block_293: +block_281: { -size_t x_264; size_t x_265; uint8_t x_266; -lean_dec(x_263); -x_264 = lean_ptr_addr(x_238); -lean_dec(x_238); -x_265 = lean_ptr_addr(x_260); -x_266 = lean_usize_dec_eq(x_264, x_265); -if (x_266 == 0) +size_t x_252; size_t x_253; uint8_t x_254; +lean_dec(x_251); +x_252 = lean_ptr_addr(x_226); +lean_dec(x_226); +x_253 = lean_ptr_addr(x_248); +x_254 = lean_usize_dec_eq(x_252, x_253); +if (x_254 == 0) { -uint8_t x_267; -lean_dec(x_237); -lean_dec(x_236); -x_267 = !lean_is_exclusive(x_1); -if (x_267 == 0) +uint8_t x_255; +lean_dec(x_225); +lean_dec(x_224); +x_255 = !lean_is_exclusive(x_1); +if (x_255 == 0) { -lean_object* x_268; lean_object* x_269; lean_object* x_270; -x_268 = lean_ctor_get(x_1, 0); -lean_dec(x_268); -if (lean_is_scalar(x_239)) { - x_269 = lean_alloc_ctor(0, 4, 0); +lean_object* x_256; lean_object* x_257; lean_object* x_258; +x_256 = lean_ctor_get(x_1, 0); +lean_dec(x_256); +if (lean_is_scalar(x_227)) { + x_257 = lean_alloc_ctor(0, 4, 0); } else { - x_269 = x_239; -} -lean_ctor_set(x_269, 0, x_235); -lean_ctor_set(x_269, 1, x_251); -lean_ctor_set(x_269, 2, x_246); -lean_ctor_set(x_269, 3, x_260); -lean_ctor_set(x_1, 0, x_269); -if (lean_is_scalar(x_262)) { - x_270 = lean_alloc_ctor(0, 2, 0); + x_257 = x_227; +} +lean_ctor_set(x_257, 0, x_223); +lean_ctor_set(x_257, 1, x_239); +lean_ctor_set(x_257, 2, x_234); +lean_ctor_set(x_257, 3, x_248); +lean_ctor_set(x_1, 0, x_257); +if (lean_is_scalar(x_250)) { + x_258 = lean_alloc_ctor(0, 2, 0); } else { - x_270 = x_262; + x_258 = x_250; } -lean_ctor_set(x_270, 0, x_1); -lean_ctor_set(x_270, 1, x_261); -return x_270; +lean_ctor_set(x_258, 0, x_1); +lean_ctor_set(x_258, 1, x_249); +return x_258; } else { -lean_object* x_271; lean_object* x_272; lean_object* x_273; +lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_dec(x_1); -if (lean_is_scalar(x_239)) { - x_271 = lean_alloc_ctor(0, 4, 0); +if (lean_is_scalar(x_227)) { + x_259 = lean_alloc_ctor(0, 4, 0); } else { - x_271 = x_239; -} -lean_ctor_set(x_271, 0, x_235); -lean_ctor_set(x_271, 1, x_251); -lean_ctor_set(x_271, 2, x_246); -lean_ctor_set(x_271, 3, x_260); -x_272 = lean_alloc_ctor(4, 1, 0); -lean_ctor_set(x_272, 0, x_271); -if (lean_is_scalar(x_262)) { - x_273 = lean_alloc_ctor(0, 2, 0); + x_259 = x_227; +} +lean_ctor_set(x_259, 0, x_223); +lean_ctor_set(x_259, 1, x_239); +lean_ctor_set(x_259, 2, x_234); +lean_ctor_set(x_259, 3, x_248); +x_260 = lean_alloc_ctor(4, 1, 0); +lean_ctor_set(x_260, 0, x_259); +if (lean_is_scalar(x_250)) { + x_261 = lean_alloc_ctor(0, 2, 0); } else { - x_273 = x_262; + x_261 = x_250; } -lean_ctor_set(x_273, 0, x_272); -lean_ctor_set(x_273, 1, x_261); -return x_273; +lean_ctor_set(x_261, 0, x_260); +lean_ctor_set(x_261, 1, x_249); +return x_261; } } else { -size_t x_274; size_t x_275; uint8_t x_276; -x_274 = lean_ptr_addr(x_236); -lean_dec(x_236); -x_275 = lean_ptr_addr(x_251); -x_276 = lean_usize_dec_eq(x_274, x_275); -if (x_276 == 0) -{ -uint8_t x_277; -lean_dec(x_237); -x_277 = !lean_is_exclusive(x_1); -if (x_277 == 0) -{ -lean_object* x_278; lean_object* x_279; lean_object* x_280; -x_278 = lean_ctor_get(x_1, 0); -lean_dec(x_278); -if (lean_is_scalar(x_239)) { - x_279 = lean_alloc_ctor(0, 4, 0); +size_t x_262; size_t x_263; uint8_t x_264; +x_262 = lean_ptr_addr(x_224); +lean_dec(x_224); +x_263 = lean_ptr_addr(x_239); +x_264 = lean_usize_dec_eq(x_262, x_263); +if (x_264 == 0) +{ +uint8_t x_265; +lean_dec(x_225); +x_265 = !lean_is_exclusive(x_1); +if (x_265 == 0) +{ +lean_object* x_266; lean_object* x_267; lean_object* x_268; +x_266 = lean_ctor_get(x_1, 0); +lean_dec(x_266); +if (lean_is_scalar(x_227)) { + x_267 = lean_alloc_ctor(0, 4, 0); } else { - x_279 = x_239; -} -lean_ctor_set(x_279, 0, x_235); -lean_ctor_set(x_279, 1, x_251); -lean_ctor_set(x_279, 2, x_246); -lean_ctor_set(x_279, 3, x_260); -lean_ctor_set(x_1, 0, x_279); -if (lean_is_scalar(x_262)) { - x_280 = lean_alloc_ctor(0, 2, 0); + x_267 = x_227; +} +lean_ctor_set(x_267, 0, x_223); +lean_ctor_set(x_267, 1, x_239); +lean_ctor_set(x_267, 2, x_234); +lean_ctor_set(x_267, 3, x_248); +lean_ctor_set(x_1, 0, x_267); +if (lean_is_scalar(x_250)) { + x_268 = lean_alloc_ctor(0, 2, 0); } else { - x_280 = x_262; + x_268 = x_250; } -lean_ctor_set(x_280, 0, x_1); -lean_ctor_set(x_280, 1, x_261); -return x_280; +lean_ctor_set(x_268, 0, x_1); +lean_ctor_set(x_268, 1, x_249); +return x_268; } else { -lean_object* x_281; lean_object* x_282; lean_object* x_283; +lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_dec(x_1); -if (lean_is_scalar(x_239)) { - x_281 = lean_alloc_ctor(0, 4, 0); +if (lean_is_scalar(x_227)) { + x_269 = lean_alloc_ctor(0, 4, 0); } else { - x_281 = x_239; -} -lean_ctor_set(x_281, 0, x_235); -lean_ctor_set(x_281, 1, x_251); -lean_ctor_set(x_281, 2, x_246); -lean_ctor_set(x_281, 3, x_260); -x_282 = lean_alloc_ctor(4, 1, 0); -lean_ctor_set(x_282, 0, x_281); -if (lean_is_scalar(x_262)) { - x_283 = lean_alloc_ctor(0, 2, 0); + x_269 = x_227; +} +lean_ctor_set(x_269, 0, x_223); +lean_ctor_set(x_269, 1, x_239); +lean_ctor_set(x_269, 2, x_234); +lean_ctor_set(x_269, 3, x_248); +x_270 = lean_alloc_ctor(4, 1, 0); +lean_ctor_set(x_270, 0, x_269); +if (lean_is_scalar(x_250)) { + x_271 = lean_alloc_ctor(0, 2, 0); } else { - x_283 = x_262; + x_271 = x_250; } -lean_ctor_set(x_283, 0, x_282); -lean_ctor_set(x_283, 1, x_261); -return x_283; +lean_ctor_set(x_271, 0, x_270); +lean_ctor_set(x_271, 1, x_249); +return x_271; } } else { -uint8_t x_284; -x_284 = lean_name_eq(x_237, x_246); -lean_dec(x_237); -if (x_284 == 0) +uint8_t x_272; +x_272 = lean_name_eq(x_225, x_234); +lean_dec(x_225); +if (x_272 == 0) { -uint8_t x_285; -x_285 = !lean_is_exclusive(x_1); -if (x_285 == 0) +uint8_t x_273; +x_273 = !lean_is_exclusive(x_1); +if (x_273 == 0) { -lean_object* x_286; lean_object* x_287; lean_object* x_288; -x_286 = lean_ctor_get(x_1, 0); -lean_dec(x_286); -if (lean_is_scalar(x_239)) { - x_287 = lean_alloc_ctor(0, 4, 0); +lean_object* x_274; lean_object* x_275; lean_object* x_276; +x_274 = lean_ctor_get(x_1, 0); +lean_dec(x_274); +if (lean_is_scalar(x_227)) { + x_275 = lean_alloc_ctor(0, 4, 0); } else { - x_287 = x_239; -} -lean_ctor_set(x_287, 0, x_235); -lean_ctor_set(x_287, 1, x_251); -lean_ctor_set(x_287, 2, x_246); -lean_ctor_set(x_287, 3, x_260); -lean_ctor_set(x_1, 0, x_287); -if (lean_is_scalar(x_262)) { - x_288 = lean_alloc_ctor(0, 2, 0); + x_275 = x_227; +} +lean_ctor_set(x_275, 0, x_223); +lean_ctor_set(x_275, 1, x_239); +lean_ctor_set(x_275, 2, x_234); +lean_ctor_set(x_275, 3, x_248); +lean_ctor_set(x_1, 0, x_275); +if (lean_is_scalar(x_250)) { + x_276 = lean_alloc_ctor(0, 2, 0); } else { - x_288 = x_262; + x_276 = x_250; } -lean_ctor_set(x_288, 0, x_1); -lean_ctor_set(x_288, 1, x_261); -return x_288; +lean_ctor_set(x_276, 0, x_1); +lean_ctor_set(x_276, 1, x_249); +return x_276; } else { -lean_object* x_289; lean_object* x_290; lean_object* x_291; +lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_dec(x_1); -if (lean_is_scalar(x_239)) { - x_289 = lean_alloc_ctor(0, 4, 0); +if (lean_is_scalar(x_227)) { + x_277 = lean_alloc_ctor(0, 4, 0); } else { - x_289 = x_239; -} -lean_ctor_set(x_289, 0, x_235); -lean_ctor_set(x_289, 1, x_251); -lean_ctor_set(x_289, 2, x_246); -lean_ctor_set(x_289, 3, x_260); -x_290 = lean_alloc_ctor(4, 1, 0); -lean_ctor_set(x_290, 0, x_289); -if (lean_is_scalar(x_262)) { - x_291 = lean_alloc_ctor(0, 2, 0); + x_277 = x_227; +} +lean_ctor_set(x_277, 0, x_223); +lean_ctor_set(x_277, 1, x_239); +lean_ctor_set(x_277, 2, x_234); +lean_ctor_set(x_277, 3, x_248); +x_278 = lean_alloc_ctor(4, 1, 0); +lean_ctor_set(x_278, 0, x_277); +if (lean_is_scalar(x_250)) { + x_279 = lean_alloc_ctor(0, 2, 0); } else { - x_291 = x_262; + x_279 = x_250; } -lean_ctor_set(x_291, 0, x_290); -lean_ctor_set(x_291, 1, x_261); -return x_291; +lean_ctor_set(x_279, 0, x_278); +lean_ctor_set(x_279, 1, x_249); +return x_279; } } else { -lean_object* x_292; -lean_dec(x_260); -lean_dec(x_251); -lean_dec(x_246); +lean_object* x_280; +lean_dec(x_248); lean_dec(x_239); -lean_dec(x_235); -if (lean_is_scalar(x_262)) { - x_292 = lean_alloc_ctor(0, 2, 0); +lean_dec(x_234); +lean_dec(x_227); +lean_dec(x_223); +if (lean_is_scalar(x_250)) { + x_280 = lean_alloc_ctor(0, 2, 0); } else { - x_292 = x_262; + x_280 = x_250; } -lean_ctor_set(x_292, 0, x_1); -lean_ctor_set(x_292, 1, x_261); -return x_292; +lean_ctor_set(x_280, 0, x_1); +lean_ctor_set(x_280, 1, x_249); +return x_280; } } } } -block_305: +block_293: { -lean_object* x_296; uint8_t x_297; -lean_dec(x_295); -x_296 = lean_unsigned_to_nat(0u); -x_297 = lean_nat_dec_lt(x_296, x_294); -lean_dec(x_294); -if (x_297 == 0) +lean_object* x_284; uint8_t x_285; +lean_dec(x_283); +x_284 = lean_unsigned_to_nat(0u); +x_285 = lean_nat_dec_lt(x_284, x_282); +lean_dec(x_282); +if (x_285 == 0) { -lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; -lean_dec(x_260); -x_298 = l___private_Lean_Compiler_LCNF_Simp_Main_0__Lean_Compiler_LCNF_Simp_oneExitPointQuick_go___closed__1; -x_299 = l___private_Init_Util_0__outOfBounds___rarg(x_298); -x_300 = l_Lean_Compiler_LCNF_AltCore_getCode(x_299); -lean_dec(x_299); -if (lean_is_scalar(x_258)) { - x_301 = lean_alloc_ctor(0, 2, 0); +lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; +lean_dec(x_248); +x_286 = l___private_Lean_Compiler_LCNF_Simp_Main_0__Lean_Compiler_LCNF_Simp_oneExitPointQuick_go___closed__1; +x_287 = l___private_Init_Util_0__outOfBounds___rarg(x_286); +x_288 = l_Lean_Compiler_LCNF_AltCore_getCode(x_287); +lean_dec(x_287); +if (lean_is_scalar(x_246)) { + x_289 = lean_alloc_ctor(0, 2, 0); } else { - x_301 = x_258; + x_289 = x_246; } -lean_ctor_set(x_301, 0, x_300); -lean_ctor_set(x_301, 1, x_261); -return x_301; +lean_ctor_set(x_289, 0, x_288); +lean_ctor_set(x_289, 1, x_249); +return x_289; } else { -lean_object* x_302; lean_object* x_303; lean_object* x_304; -x_302 = lean_array_fget(x_260, x_296); -lean_dec(x_260); -x_303 = l_Lean_Compiler_LCNF_AltCore_getCode(x_302); -lean_dec(x_302); -if (lean_is_scalar(x_258)) { - x_304 = lean_alloc_ctor(0, 2, 0); +lean_object* x_290; lean_object* x_291; lean_object* x_292; +x_290 = lean_array_fget(x_248, x_284); +lean_dec(x_248); +x_291 = l_Lean_Compiler_LCNF_AltCore_getCode(x_290); +lean_dec(x_290); +if (lean_is_scalar(x_246)) { + x_292 = lean_alloc_ctor(0, 2, 0); } else { - x_304 = x_258; + x_292 = x_246; } -lean_ctor_set(x_304, 0, x_303); -lean_ctor_set(x_304, 1, x_261); -return x_304; +lean_ctor_set(x_292, 0, x_291); +lean_ctor_set(x_292, 1, x_249); +return x_292; } } } else { -uint8_t x_317; -lean_dec(x_258); -lean_dec(x_251); +uint8_t x_305; lean_dec(x_246); lean_dec(x_239); -lean_dec(x_238); -lean_dec(x_237); -lean_dec(x_236); -lean_dec(x_235); +lean_dec(x_234); +lean_dec(x_227); +lean_dec(x_226); +lean_dec(x_225); +lean_dec(x_224); +lean_dec(x_223); lean_dec(x_1); -x_317 = !lean_is_exclusive(x_259); -if (x_317 == 0) +x_305 = !lean_is_exclusive(x_247); +if (x_305 == 0) { -return x_259; +return x_247; } else { -lean_object* x_318; lean_object* x_319; lean_object* x_320; -x_318 = lean_ctor_get(x_259, 0); -x_319 = lean_ctor_get(x_259, 1); -lean_inc(x_319); -lean_inc(x_318); -lean_dec(x_259); -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_306; lean_object* x_307; lean_object* x_308; +x_306 = lean_ctor_get(x_247, 0); +x_307 = lean_ctor_get(x_247, 1); +lean_inc(x_307); +lean_inc(x_306); +lean_dec(x_247); +x_308 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_308, 0, x_306); +lean_ctor_set(x_308, 1, x_307); +return x_308; } } } else { -uint8_t x_321; -lean_dec(x_251); -lean_dec(x_246); +uint8_t x_309; lean_dec(x_239); -lean_dec(x_238); -lean_dec(x_237); -lean_dec(x_236); -lean_dec(x_235); +lean_dec(x_234); +lean_dec(x_227); +lean_dec(x_226); +lean_dec(x_225); +lean_dec(x_224); +lean_dec(x_223); lean_dec(x_7); lean_dec(x_8); lean_dec(x_6); @@ -12760,50 +12789,50 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_321 = !lean_is_exclusive(x_255); -if (x_321 == 0) +x_309 = !lean_is_exclusive(x_243); +if (x_309 == 0) { -return x_255; +return x_243; } else { -lean_object* x_322; lean_object* x_323; lean_object* x_324; -x_322 = lean_ctor_get(x_255, 0); -x_323 = lean_ctor_get(x_255, 1); -lean_inc(x_323); -lean_inc(x_322); -lean_dec(x_255); -x_324 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_324, 0, x_322); -lean_ctor_set(x_324, 1, x_323); -return x_324; +lean_object* x_310; lean_object* x_311; lean_object* x_312; +x_310 = lean_ctor_get(x_243, 0); +x_311 = lean_ctor_get(x_243, 1); +lean_inc(x_311); +lean_inc(x_310); +lean_dec(x_243); +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; } } } else { -lean_object* x_325; -lean_dec(x_239); -lean_dec(x_238); -lean_dec(x_237); -lean_dec(x_236); -lean_dec(x_235); +lean_object* x_313; +lean_dec(x_227); +lean_dec(x_226); +lean_dec(x_225); +lean_dec(x_224); +lean_dec(x_223); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_325 = l_Lean_Compiler_LCNF_mkReturnErased(x_5, x_6, x_7, x_8, x_242); +x_313 = l_Lean_Compiler_LCNF_mkReturnErased(x_5, x_6, x_7, x_8, x_230); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -return x_325; +return x_313; } } else { -uint8_t x_326; -lean_dec(x_231); +uint8_t x_314; +lean_dec(x_219); lean_dec(x_7); lean_dec(x_8); lean_dec(x_6); @@ -12812,38 +12841,38 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_326 = !lean_is_exclusive(x_232); -if (x_326 == 0) +x_314 = !lean_is_exclusive(x_220); +if (x_314 == 0) { -lean_object* x_327; lean_object* x_328; -x_327 = lean_ctor_get(x_232, 0); -lean_dec(x_327); -x_328 = lean_ctor_get(x_233, 0); -lean_inc(x_328); -lean_dec(x_233); -lean_ctor_set(x_232, 0, x_328); -return x_232; +lean_object* x_315; lean_object* x_316; +x_315 = lean_ctor_get(x_220, 0); +lean_dec(x_315); +x_316 = lean_ctor_get(x_221, 0); +lean_inc(x_316); +lean_dec(x_221); +lean_ctor_set(x_220, 0, x_316); +return x_220; } else { -lean_object* x_329; lean_object* x_330; lean_object* x_331; -x_329 = lean_ctor_get(x_232, 1); -lean_inc(x_329); -lean_dec(x_232); -x_330 = lean_ctor_get(x_233, 0); -lean_inc(x_330); -lean_dec(x_233); -x_331 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_331, 0, x_330); -lean_ctor_set(x_331, 1, x_329); -return x_331; +lean_object* x_317; lean_object* x_318; lean_object* x_319; +x_317 = lean_ctor_get(x_220, 1); +lean_inc(x_317); +lean_dec(x_220); +x_318 = lean_ctor_get(x_221, 0); +lean_inc(x_318); +lean_dec(x_221); +x_319 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_319, 0, x_318); +lean_ctor_set(x_319, 1, x_317); +return x_319; } } } else { -uint8_t x_332; -lean_dec(x_231); +uint8_t x_320; +lean_dec(x_219); lean_dec(x_7); lean_dec(x_8); lean_dec(x_6); @@ -12852,54 +12881,54 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_332 = !lean_is_exclusive(x_232); -if (x_332 == 0) +x_320 = !lean_is_exclusive(x_220); +if (x_320 == 0) { -return x_232; +return x_220; } else { -lean_object* x_333; lean_object* x_334; lean_object* x_335; -x_333 = lean_ctor_get(x_232, 0); -x_334 = lean_ctor_get(x_232, 1); -lean_inc(x_334); -lean_inc(x_333); -lean_dec(x_232); -x_335 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_335, 0, x_333); -lean_ctor_set(x_335, 1, x_334); -return x_335; +lean_object* x_321; lean_object* x_322; lean_object* x_323; +x_321 = lean_ctor_get(x_220, 0); +x_322 = lean_ctor_get(x_220, 1); +lean_inc(x_322); +lean_inc(x_321); +lean_dec(x_220); +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; } } } case 5: { -lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; uint8_t x_342; lean_object* x_343; -x_336 = lean_ctor_get(x_36, 1); -lean_inc(x_336); +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_324 = lean_ctor_get(x_36, 1); +lean_inc(x_324); lean_dec(x_36); -x_337 = lean_ctor_get(x_1, 0); -lean_inc(x_337); -x_338 = lean_st_ref_get(x_3, 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_341 = lean_ctor_get(x_339, 0); -lean_inc(x_341); -lean_dec(x_339); -x_342 = 0; -lean_inc(x_337); -x_343 = l___private_Lean_Compiler_LCNF_CompilerM_0__Lean_Compiler_LCNF_normFVarImp(x_341, x_337, x_342); -if (lean_obj_tag(x_343) == 0) -{ -lean_object* x_344; lean_object* x_345; uint8_t x_346; -x_344 = lean_ctor_get(x_343, 0); -lean_inc(x_344); -lean_dec(x_343); -lean_inc(x_344); -x_345 = l_Lean_Compiler_LCNF_Simp_markUsedFVar(x_344, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_340); +x_325 = lean_ctor_get(x_1, 0); +lean_inc(x_325); +x_326 = lean_st_ref_get(x_3, x_324); +x_327 = lean_ctor_get(x_326, 0); +lean_inc(x_327); +x_328 = lean_ctor_get(x_326, 1); +lean_inc(x_328); +lean_dec(x_326); +x_329 = lean_ctor_get(x_327, 0); +lean_inc(x_329); +lean_dec(x_327); +x_330 = 0; +lean_inc(x_325); +x_331 = l___private_Lean_Compiler_LCNF_CompilerM_0__Lean_Compiler_LCNF_normFVarImp(x_329, x_325, x_330); +if (lean_obj_tag(x_331) == 0) +{ +lean_object* x_332; lean_object* x_333; uint8_t x_334; +x_332 = lean_ctor_get(x_331, 0); +lean_inc(x_332); +lean_dec(x_331); +lean_inc(x_332); +x_333 = l_Lean_Compiler_LCNF_Simp_markUsedFVar(x_332, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_328); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -12907,208 +12936,208 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_346 = !lean_is_exclusive(x_345); -if (x_346 == 0) -{ -lean_object* x_347; uint8_t x_348; -x_347 = lean_ctor_get(x_345, 0); -lean_dec(x_347); -x_348 = lean_name_eq(x_337, x_344); -lean_dec(x_337); -if (x_348 == 0) -{ -uint8_t x_349; -x_349 = !lean_is_exclusive(x_1); -if (x_349 == 0) -{ -lean_object* x_350; -x_350 = lean_ctor_get(x_1, 0); -lean_dec(x_350); -lean_ctor_set(x_1, 0, x_344); -lean_ctor_set(x_345, 0, x_1); -return x_345; +x_334 = !lean_is_exclusive(x_333); +if (x_334 == 0) +{ +lean_object* x_335; uint8_t x_336; +x_335 = lean_ctor_get(x_333, 0); +lean_dec(x_335); +x_336 = lean_name_eq(x_325, x_332); +lean_dec(x_325); +if (x_336 == 0) +{ +uint8_t x_337; +x_337 = !lean_is_exclusive(x_1); +if (x_337 == 0) +{ +lean_object* x_338; +x_338 = lean_ctor_get(x_1, 0); +lean_dec(x_338); +lean_ctor_set(x_1, 0, x_332); +lean_ctor_set(x_333, 0, x_1); +return x_333; } else { -lean_object* x_351; +lean_object* x_339; lean_dec(x_1); -x_351 = lean_alloc_ctor(5, 1, 0); -lean_ctor_set(x_351, 0, x_344); -lean_ctor_set(x_345, 0, x_351); -return x_345; +x_339 = lean_alloc_ctor(5, 1, 0); +lean_ctor_set(x_339, 0, x_332); +lean_ctor_set(x_333, 0, x_339); +return x_333; } } else { -lean_dec(x_344); -lean_ctor_set(x_345, 0, x_1); -return x_345; +lean_dec(x_332); +lean_ctor_set(x_333, 0, x_1); +return x_333; } } else { -lean_object* x_352; uint8_t x_353; -x_352 = lean_ctor_get(x_345, 1); -lean_inc(x_352); -lean_dec(x_345); -x_353 = lean_name_eq(x_337, x_344); -lean_dec(x_337); -if (x_353 == 0) +lean_object* x_340; uint8_t x_341; +x_340 = lean_ctor_get(x_333, 1); +lean_inc(x_340); +lean_dec(x_333); +x_341 = lean_name_eq(x_325, x_332); +lean_dec(x_325); +if (x_341 == 0) { -lean_object* x_354; lean_object* x_355; lean_object* x_356; +lean_object* x_342; lean_object* x_343; lean_object* x_344; if (lean_is_exclusive(x_1)) { lean_ctor_release(x_1, 0); - x_354 = x_1; + x_342 = x_1; } else { lean_dec_ref(x_1); - x_354 = lean_box(0); + x_342 = lean_box(0); } -if (lean_is_scalar(x_354)) { - x_355 = lean_alloc_ctor(5, 1, 0); +if (lean_is_scalar(x_342)) { + x_343 = lean_alloc_ctor(5, 1, 0); } else { - x_355 = x_354; + x_343 = x_342; } -lean_ctor_set(x_355, 0, x_344); -x_356 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_356, 0, x_355); -lean_ctor_set(x_356, 1, x_352); -return x_356; +lean_ctor_set(x_343, 0, x_332); +x_344 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_344, 0, x_343); +lean_ctor_set(x_344, 1, x_340); +return x_344; } else { -lean_object* x_357; -lean_dec(x_344); -x_357 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_357, 0, x_1); -lean_ctor_set(x_357, 1, x_352); -return x_357; +lean_object* x_345; +lean_dec(x_332); +x_345 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_345, 0, x_1); +lean_ctor_set(x_345, 1, x_340); +return x_345; } } } else { -lean_object* x_358; -lean_dec(x_337); +lean_object* x_346; +lean_dec(x_325); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_358 = l_Lean_Compiler_LCNF_mkReturnErased(x_5, x_6, x_7, x_8, x_340); +x_346 = l_Lean_Compiler_LCNF_mkReturnErased(x_5, x_6, x_7, x_8, x_328); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -return x_358; +return x_346; } } default: { -lean_object* x_359; lean_object* x_360; lean_object* x_361; uint8_t x_362; +lean_object* x_347; lean_object* x_348; lean_object* x_349; uint8_t x_350; lean_dec(x_7); lean_dec(x_8); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); -x_359 = lean_ctor_get(x_36, 1); -lean_inc(x_359); +x_347 = lean_ctor_get(x_36, 1); +lean_inc(x_347); lean_dec(x_36); -x_360 = lean_ctor_get(x_1, 0); -lean_inc(x_360); -x_361 = lean_st_ref_get(x_3, x_359); +x_348 = lean_ctor_get(x_1, 0); +lean_inc(x_348); +x_349 = lean_st_ref_get(x_3, x_347); lean_dec(x_3); -x_362 = !lean_is_exclusive(x_361); -if (x_362 == 0) +x_350 = !lean_is_exclusive(x_349); +if (x_350 == 0) { -lean_object* x_363; lean_object* x_364; uint8_t x_365; lean_object* x_366; size_t x_367; size_t x_368; uint8_t x_369; -x_363 = lean_ctor_get(x_361, 0); -x_364 = lean_ctor_get(x_363, 0); -lean_inc(x_364); -lean_dec(x_363); -x_365 = 0; -lean_inc(x_360); -x_366 = l___private_Lean_Compiler_LCNF_CompilerM_0__Lean_Compiler_LCNF_normExprImp_go(x_364, x_365, x_360); -x_367 = lean_ptr_addr(x_360); -lean_dec(x_360); -x_368 = lean_ptr_addr(x_366); -x_369 = lean_usize_dec_eq(x_367, x_368); -if (x_369 == 0) -{ -uint8_t x_370; -x_370 = !lean_is_exclusive(x_1); -if (x_370 == 0) -{ -lean_object* x_371; -x_371 = lean_ctor_get(x_1, 0); -lean_dec(x_371); -lean_ctor_set(x_1, 0, x_366); -lean_ctor_set(x_361, 0, x_1); -return x_361; +lean_object* x_351; lean_object* x_352; uint8_t x_353; lean_object* x_354; size_t x_355; size_t x_356; uint8_t x_357; +x_351 = lean_ctor_get(x_349, 0); +x_352 = lean_ctor_get(x_351, 0); +lean_inc(x_352); +lean_dec(x_351); +x_353 = 0; +lean_inc(x_348); +x_354 = l___private_Lean_Compiler_LCNF_CompilerM_0__Lean_Compiler_LCNF_normExprImp_go(x_352, x_353, x_348); +x_355 = lean_ptr_addr(x_348); +lean_dec(x_348); +x_356 = lean_ptr_addr(x_354); +x_357 = lean_usize_dec_eq(x_355, x_356); +if (x_357 == 0) +{ +uint8_t x_358; +x_358 = !lean_is_exclusive(x_1); +if (x_358 == 0) +{ +lean_object* x_359; +x_359 = lean_ctor_get(x_1, 0); +lean_dec(x_359); +lean_ctor_set(x_1, 0, x_354); +lean_ctor_set(x_349, 0, x_1); +return x_349; } else { -lean_object* x_372; +lean_object* x_360; lean_dec(x_1); -x_372 = lean_alloc_ctor(6, 1, 0); -lean_ctor_set(x_372, 0, x_366); -lean_ctor_set(x_361, 0, x_372); -return x_361; +x_360 = lean_alloc_ctor(6, 1, 0); +lean_ctor_set(x_360, 0, x_354); +lean_ctor_set(x_349, 0, x_360); +return x_349; } } else { -lean_dec(x_366); -lean_ctor_set(x_361, 0, x_1); -return x_361; +lean_dec(x_354); +lean_ctor_set(x_349, 0, x_1); +return x_349; } } else { -lean_object* x_373; lean_object* x_374; lean_object* x_375; uint8_t x_376; lean_object* x_377; size_t x_378; size_t x_379; uint8_t x_380; -x_373 = lean_ctor_get(x_361, 0); -x_374 = lean_ctor_get(x_361, 1); -lean_inc(x_374); -lean_inc(x_373); +lean_object* x_361; lean_object* x_362; lean_object* x_363; uint8_t x_364; lean_object* x_365; size_t x_366; size_t x_367; uint8_t x_368; +x_361 = lean_ctor_get(x_349, 0); +x_362 = lean_ctor_get(x_349, 1); +lean_inc(x_362); +lean_inc(x_361); +lean_dec(x_349); +x_363 = lean_ctor_get(x_361, 0); +lean_inc(x_363); lean_dec(x_361); -x_375 = lean_ctor_get(x_373, 0); -lean_inc(x_375); -lean_dec(x_373); -x_376 = 0; -lean_inc(x_360); -x_377 = l___private_Lean_Compiler_LCNF_CompilerM_0__Lean_Compiler_LCNF_normExprImp_go(x_375, x_376, x_360); -x_378 = lean_ptr_addr(x_360); -lean_dec(x_360); -x_379 = lean_ptr_addr(x_377); -x_380 = lean_usize_dec_eq(x_378, x_379); -if (x_380 == 0) -{ -lean_object* x_381; lean_object* x_382; lean_object* x_383; +x_364 = 0; +lean_inc(x_348); +x_365 = l___private_Lean_Compiler_LCNF_CompilerM_0__Lean_Compiler_LCNF_normExprImp_go(x_363, x_364, x_348); +x_366 = lean_ptr_addr(x_348); +lean_dec(x_348); +x_367 = lean_ptr_addr(x_365); +x_368 = lean_usize_dec_eq(x_366, x_367); +if (x_368 == 0) +{ +lean_object* x_369; lean_object* x_370; lean_object* x_371; if (lean_is_exclusive(x_1)) { lean_ctor_release(x_1, 0); - x_381 = x_1; + x_369 = x_1; } else { lean_dec_ref(x_1); - x_381 = lean_box(0); + x_369 = lean_box(0); } -if (lean_is_scalar(x_381)) { - x_382 = lean_alloc_ctor(6, 1, 0); +if (lean_is_scalar(x_369)) { + x_370 = lean_alloc_ctor(6, 1, 0); } else { - x_382 = x_381; + x_370 = x_369; } -lean_ctor_set(x_382, 0, x_377); -x_383 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_383, 0, x_382); -lean_ctor_set(x_383, 1, x_374); -return x_383; +lean_ctor_set(x_370, 0, x_365); +x_371 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_371, 0, x_370); +lean_ctor_set(x_371, 1, x_362); +return x_371; } else { -lean_object* x_384; -lean_dec(x_377); -x_384 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_384, 0, x_1); -lean_ctor_set(x_384, 1, x_374); -return x_384; +lean_object* x_372; +lean_dec(x_365); +x_372 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_372, 0, x_1); +lean_ctor_set(x_372, 1, x_362); +return x_372; } } } @@ -13116,340 +13145,278 @@ return x_384; } else { -lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; +lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_dec(x_7); -x_385 = lean_unsigned_to_nat(1u); -x_386 = lean_nat_add(x_13, x_385); +x_373 = lean_unsigned_to_nat(1u); +x_374 = lean_nat_add(x_13, x_373); lean_dec(x_13); -x_387 = lean_alloc_ctor(0, 11, 0); -lean_ctor_set(x_387, 0, x_10); -lean_ctor_set(x_387, 1, x_11); -lean_ctor_set(x_387, 2, x_12); -lean_ctor_set(x_387, 3, x_386); -lean_ctor_set(x_387, 4, x_14); -lean_ctor_set(x_387, 5, x_15); -lean_ctor_set(x_387, 6, x_16); -lean_ctor_set(x_387, 7, x_17); -lean_ctor_set(x_387, 8, x_18); -lean_ctor_set(x_387, 9, x_19); -lean_ctor_set(x_387, 10, x_20); -x_388 = l_Lean_Compiler_LCNF_Simp_incVisited___rarg(x_3, x_4, x_5, x_6, x_387, x_8, x_9); +x_375 = lean_alloc_ctor(0, 11, 0); +lean_ctor_set(x_375, 0, x_10); +lean_ctor_set(x_375, 1, x_11); +lean_ctor_set(x_375, 2, x_12); +lean_ctor_set(x_375, 3, x_374); +lean_ctor_set(x_375, 4, x_14); +lean_ctor_set(x_375, 5, x_15); +lean_ctor_set(x_375, 6, x_16); +lean_ctor_set(x_375, 7, x_17); +lean_ctor_set(x_375, 8, x_18); +lean_ctor_set(x_375, 9, x_19); +lean_ctor_set(x_375, 10, x_20); +x_376 = l_Lean_Compiler_LCNF_Simp_incVisited___rarg(x_3, x_4, x_5, x_6, x_375, x_8, x_9); switch (lean_obj_tag(x_1)) { case 0: { -lean_object* x_389; lean_object* x_390; lean_object* x_391; uint8_t x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; -x_389 = lean_ctor_get(x_388, 1); -lean_inc(x_389); -lean_dec(x_388); -x_390 = lean_ctor_get(x_1, 0); -lean_inc(x_390); -x_391 = lean_ctor_get(x_1, 1); +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; uint8_t x_384; +x_377 = lean_ctor_get(x_376, 1); +lean_inc(x_377); +lean_dec(x_376); +x_378 = lean_ctor_get(x_1, 0); +lean_inc(x_378); +x_379 = lean_ctor_get(x_1, 1); +lean_inc(x_379); +x_380 = 0; +lean_inc(x_378); +x_381 = l_Lean_Compiler_LCNF_normLetDecl___at_Lean_Compiler_LCNF_Simp_simp___spec__1(x_380, x_378, x_2, x_3, x_4, x_5, x_6, x_375, x_8, x_377); +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_384 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetDecl____x40_Lean_Compiler_LCNF_Basic___hyg_1940_(x_378, x_382); +if (x_384 == 0) +{ +lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; +x_385 = l_Lean_Compiler_LCNF_Simp_markSimplified___rarg(x_3, x_4, x_5, x_6, x_375, x_8, x_383); +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_Compiler_LCNF_Simp_simp___lambda__2(x_379, x_378, x_1, x_382, x_386, x_2, x_3, x_4, x_5, x_6, x_375, x_8, x_387); +return x_388; +} +else +{ +lean_object* x_389; lean_object* x_390; +x_389 = lean_box(0); +x_390 = l_Lean_Compiler_LCNF_Simp_simp___lambda__2(x_379, x_378, x_1, x_382, x_389, x_2, x_3, x_4, x_5, x_6, x_375, x_8, x_383); +return x_390; +} +} +case 1: +{ +lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; uint8_t x_397; +x_391 = lean_ctor_get(x_376, 1); lean_inc(x_391); -x_392 = 0; -lean_inc(x_390); -x_393 = l_Lean_Compiler_LCNF_normLetDecl___at_Lean_Compiler_LCNF_Simp_simp___spec__1(x_392, x_390, x_2, x_3, x_4, x_5, x_6, x_387, x_8, x_389); -x_394 = lean_ctor_get(x_393, 0); +lean_dec(x_376); +x_392 = lean_ctor_get(x_1, 0); +lean_inc(x_392); +x_393 = lean_ctor_get(x_1, 1); +lean_inc(x_393); +x_394 = lean_ctor_get(x_392, 0); 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_394, 3); +x_395 = l_Lean_Compiler_LCNF_Simp_isOnceOrMustInline(x_394, x_2, x_3, x_4, x_5, x_6, x_375, x_8, x_391); +x_396 = lean_ctor_get(x_395, 0); lean_inc(x_396); -lean_inc(x_8); -lean_inc(x_387); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_397 = l_Lean_Compiler_LCNF_Simp_simpValue_x3f(x_396, x_2, x_3, x_4, x_5, x_6, x_387, x_8, x_395); -if (lean_obj_tag(x_397) == 0) +x_397 = lean_unbox(x_396); +if (x_397 == 0) { -lean_object* x_398; -x_398 = lean_ctor_get(x_397, 0); +lean_object* x_398; lean_object* x_399; uint8_t x_400; +x_398 = lean_ctor_get(x_395, 1); lean_inc(x_398); -if (lean_obj_tag(x_398) == 0) +lean_dec(x_395); +lean_inc(x_1); +lean_inc(x_392); +x_399 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_Simp_simp___lambda__4___boxed), 14, 4); +lean_closure_set(x_399, 0, x_393); +lean_closure_set(x_399, 1, x_392); +lean_closure_set(x_399, 2, x_1); +lean_closure_set(x_399, 3, x_396); +x_400 = l_Lean_Compiler_LCNF_Code_isFun(x_1); +lean_dec(x_1); +if (x_400 == 0) { -lean_object* x_399; lean_object* x_400; lean_object* x_401; -x_399 = lean_ctor_get(x_397, 1); -lean_inc(x_399); -lean_dec(x_397); -x_400 = lean_box(0); -x_401 = l_Lean_Compiler_LCNF_Simp_simp___lambda__1(x_391, x_390, x_1, x_394, x_400, x_2, x_3, x_4, x_5, x_6, x_387, x_8, x_399); -return x_401; +lean_object* x_401; lean_object* x_402; +x_401 = lean_box(0); +x_402 = l_Lean_Compiler_LCNF_Simp_simp___lambda__5(x_399, x_392, x_401, x_2, x_3, x_4, x_5, x_6, x_375, x_8, x_398); +return x_402; } 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_397, 1); -lean_inc(x_402); -lean_dec(x_397); -x_403 = lean_ctor_get(x_398, 0); +lean_object* x_403; lean_object* x_404; uint8_t x_405; +x_403 = lean_ctor_get(x_392, 3); lean_inc(x_403); -lean_dec(x_398); -x_404 = l_Lean_Compiler_LCNF_Simp_markSimplified___rarg(x_3, x_4, x_5, x_6, x_387, x_8, x_402); -x_405 = lean_ctor_get(x_404, 1); -lean_inc(x_405); +x_404 = lean_ctor_get(x_392, 2); +lean_inc(x_404); +x_405 = l_Lean_Compiler_LCNF_isEtaExpandCandidateCore(x_403, x_404); lean_dec(x_404); -x_406 = l_Lean_Compiler_LCNF_LetDecl_updateValue(x_394, x_403, x_5, x_6, x_387, x_8, x_405); -x_407 = lean_ctor_get(x_406, 0); -lean_inc(x_407); -x_408 = lean_ctor_get(x_406, 1); -lean_inc(x_408); -lean_dec(x_406); -x_409 = lean_box(0); -x_410 = l_Lean_Compiler_LCNF_Simp_simp___lambda__1(x_391, x_390, x_1, x_407, x_409, x_2, x_3, x_4, x_5, x_6, x_387, x_8, x_408); -return x_410; -} +if (x_405 == 0) +{ +lean_object* x_406; lean_object* x_407; +x_406 = lean_box(0); +x_407 = l_Lean_Compiler_LCNF_Simp_simp___lambda__5(x_399, x_392, x_406, x_2, x_3, x_4, x_5, x_6, x_375, x_8, x_398); +return x_407; } else { -lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; -lean_dec(x_394); -lean_dec(x_391); -lean_dec(x_390); -lean_dec(x_387); -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_411 = lean_ctor_get(x_397, 0); +lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; uint8_t x_412; lean_object* x_413; +x_408 = lean_st_ref_get(x_3, x_398); +x_409 = lean_ctor_get(x_408, 0); +lean_inc(x_409); +x_410 = lean_ctor_get(x_408, 1); +lean_inc(x_410); +lean_dec(x_408); +x_411 = lean_ctor_get(x_409, 0); lean_inc(x_411); -x_412 = lean_ctor_get(x_397, 1); -lean_inc(x_412); -if (lean_is_exclusive(x_397)) { - lean_ctor_release(x_397, 0); - lean_ctor_release(x_397, 1); - x_413 = x_397; -} else { - lean_dec_ref(x_397); - x_413 = lean_box(0); -} -if (lean_is_scalar(x_413)) { - x_414 = lean_alloc_ctor(1, 2, 0); -} else { - x_414 = x_413; -} -lean_ctor_set(x_414, 0, x_411); -lean_ctor_set(x_414, 1, x_412); -return x_414; -} -} -case 1: -{ -lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; uint8_t x_421; -x_415 = lean_ctor_get(x_388, 1); -lean_inc(x_415); -lean_dec(x_388); -x_416 = lean_ctor_get(x_1, 0); -lean_inc(x_416); -x_417 = lean_ctor_get(x_1, 1); -lean_inc(x_417); -x_418 = lean_ctor_get(x_416, 0); -lean_inc(x_418); -x_419 = l_Lean_Compiler_LCNF_Simp_isOnceOrMustInline(x_418, x_2, x_3, x_4, x_5, x_6, x_387, x_8, x_415); -x_420 = lean_ctor_get(x_419, 0); -lean_inc(x_420); -x_421 = lean_unbox(x_420); -if (x_421 == 0) -{ -lean_object* x_422; lean_object* x_423; uint8_t x_424; -x_422 = lean_ctor_get(x_419, 1); -lean_inc(x_422); -lean_dec(x_419); -lean_inc(x_1); -lean_inc(x_416); -x_423 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_Simp_simp___lambda__3___boxed), 14, 4); -lean_closure_set(x_423, 0, x_417); -lean_closure_set(x_423, 1, x_416); -lean_closure_set(x_423, 2, x_1); -lean_closure_set(x_423, 3, x_420); -x_424 = l_Lean_Compiler_LCNF_Code_isFun(x_1); -lean_dec(x_1); -if (x_424 == 0) -{ -lean_object* x_425; lean_object* x_426; -x_425 = lean_box(0); -x_426 = l_Lean_Compiler_LCNF_Simp_simp___lambda__4(x_423, x_416, x_425, x_2, x_3, x_4, x_5, x_6, x_387, x_8, x_422); -return x_426; -} -else -{ -lean_object* x_427; lean_object* x_428; uint8_t x_429; -x_427 = lean_ctor_get(x_416, 3); -lean_inc(x_427); -x_428 = lean_ctor_get(x_416, 2); -lean_inc(x_428); -x_429 = l_Lean_Compiler_LCNF_isEtaExpandCandidateCore(x_427, x_428); -lean_dec(x_428); -if (x_429 == 0) -{ -lean_object* x_430; lean_object* x_431; -x_430 = lean_box(0); -x_431 = l_Lean_Compiler_LCNF_Simp_simp___lambda__4(x_423, x_416, x_430, x_2, x_3, x_4, x_5, x_6, x_387, x_8, x_422); -return x_431; -} -else -{ -lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; uint8_t x_436; lean_object* x_437; -x_432 = lean_st_ref_get(x_3, x_422); -x_433 = lean_ctor_get(x_432, 0); -lean_inc(x_433); -x_434 = lean_ctor_get(x_432, 1); -lean_inc(x_434); -lean_dec(x_432); -x_435 = lean_ctor_get(x_433, 0); -lean_inc(x_435); -lean_dec(x_433); -x_436 = 0; +lean_dec(x_409); +x_412 = 0; lean_inc(x_8); -lean_inc(x_387); +lean_inc(x_375); lean_inc(x_6); lean_inc(x_5); -x_437 = l_Lean_Compiler_LCNF_normFunDeclImp(x_436, x_416, x_435, x_5, x_6, x_387, x_8, x_434); -if (lean_obj_tag(x_437) == 0) +x_413 = l_Lean_Compiler_LCNF_normFunDeclImp(x_412, x_392, x_411, x_5, x_6, x_375, x_8, x_410); +if (lean_obj_tag(x_413) == 0) { -lean_object* x_438; lean_object* x_439; lean_object* x_440; -x_438 = lean_ctor_get(x_437, 0); -lean_inc(x_438); -x_439 = lean_ctor_get(x_437, 1); -lean_inc(x_439); -lean_dec(x_437); +lean_object* x_414; lean_object* x_415; lean_object* x_416; +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); lean_inc(x_8); -lean_inc(x_387); +lean_inc(x_375); lean_inc(x_6); lean_inc(x_5); -x_440 = l_Lean_Compiler_LCNF_FunDeclCore_etaExpand(x_438, x_5, x_6, x_387, x_8, x_439); -if (lean_obj_tag(x_440) == 0) -{ -lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; -x_441 = lean_ctor_get(x_440, 0); -lean_inc(x_441); -x_442 = lean_ctor_get(x_440, 1); -lean_inc(x_442); -lean_dec(x_440); -x_443 = l_Lean_Compiler_LCNF_Simp_markSimplified___rarg(x_3, x_4, x_5, x_6, x_387, x_8, x_442); -x_444 = lean_ctor_get(x_443, 0); -lean_inc(x_444); -x_445 = lean_ctor_get(x_443, 1); -lean_inc(x_445); -lean_dec(x_443); -x_446 = l_Lean_Compiler_LCNF_Simp_simp___lambda__4(x_423, x_441, x_444, x_2, x_3, x_4, x_5, x_6, x_387, x_8, x_445); -lean_dec(x_444); -return x_446; +x_416 = l_Lean_Compiler_LCNF_FunDeclCore_etaExpand(x_414, x_5, x_6, x_375, x_8, x_415); +if (lean_obj_tag(x_416) == 0) +{ +lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; +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 = l_Lean_Compiler_LCNF_Simp_markSimplified___rarg(x_3, x_4, x_5, x_6, x_375, x_8, x_418); +x_420 = lean_ctor_get(x_419, 0); +lean_inc(x_420); +x_421 = lean_ctor_get(x_419, 1); +lean_inc(x_421); +lean_dec(x_419); +x_422 = l_Lean_Compiler_LCNF_Simp_simp___lambda__5(x_399, x_417, x_420, x_2, x_3, x_4, x_5, x_6, x_375, x_8, x_421); +lean_dec(x_420); +return x_422; } else { -lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; -lean_dec(x_423); -lean_dec(x_387); +lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; +lean_dec(x_399); +lean_dec(x_375); 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_447 = lean_ctor_get(x_440, 0); -lean_inc(x_447); -x_448 = lean_ctor_get(x_440, 1); -lean_inc(x_448); -if (lean_is_exclusive(x_440)) { - lean_ctor_release(x_440, 0); - lean_ctor_release(x_440, 1); - x_449 = x_440; +x_423 = lean_ctor_get(x_416, 0); +lean_inc(x_423); +x_424 = lean_ctor_get(x_416, 1); +lean_inc(x_424); +if (lean_is_exclusive(x_416)) { + lean_ctor_release(x_416, 0); + lean_ctor_release(x_416, 1); + x_425 = x_416; } else { - lean_dec_ref(x_440); - x_449 = lean_box(0); + lean_dec_ref(x_416); + x_425 = lean_box(0); } -if (lean_is_scalar(x_449)) { - x_450 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_425)) { + x_426 = lean_alloc_ctor(1, 2, 0); } else { - x_450 = x_449; + x_426 = x_425; } -lean_ctor_set(x_450, 0, x_447); -lean_ctor_set(x_450, 1, x_448); -return x_450; +lean_ctor_set(x_426, 0, x_423); +lean_ctor_set(x_426, 1, x_424); +return x_426; } } else { -lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; -lean_dec(x_423); -lean_dec(x_387); +lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; +lean_dec(x_399); +lean_dec(x_375); 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_451 = lean_ctor_get(x_437, 0); -lean_inc(x_451); -x_452 = lean_ctor_get(x_437, 1); -lean_inc(x_452); -if (lean_is_exclusive(x_437)) { - lean_ctor_release(x_437, 0); - lean_ctor_release(x_437, 1); - x_453 = x_437; +x_427 = lean_ctor_get(x_413, 0); +lean_inc(x_427); +x_428 = lean_ctor_get(x_413, 1); +lean_inc(x_428); +if (lean_is_exclusive(x_413)) { + lean_ctor_release(x_413, 0); + lean_ctor_release(x_413, 1); + x_429 = x_413; } else { - lean_dec_ref(x_437); - x_453 = lean_box(0); + lean_dec_ref(x_413); + x_429 = lean_box(0); } -if (lean_is_scalar(x_453)) { - x_454 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_429)) { + x_430 = lean_alloc_ctor(1, 2, 0); } else { - x_454 = x_453; + x_430 = x_429; } -lean_ctor_set(x_454, 0, x_451); -lean_ctor_set(x_454, 1, x_452); -return x_454; +lean_ctor_set(x_430, 0, x_427); +lean_ctor_set(x_430, 1, x_428); +return x_430; } } } } else { -lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; uint8_t x_460; lean_object* x_461; -x_455 = lean_ctor_get(x_419, 1); -lean_inc(x_455); -lean_dec(x_419); -x_456 = lean_st_ref_get(x_3, x_455); -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_457, 0); -lean_inc(x_459); -lean_dec(x_457); -x_460 = 0; +lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; lean_object* x_435; uint8_t x_436; lean_object* x_437; +x_431 = lean_ctor_get(x_395, 1); +lean_inc(x_431); +lean_dec(x_395); +x_432 = lean_st_ref_get(x_3, x_431); +x_433 = lean_ctor_get(x_432, 0); +lean_inc(x_433); +x_434 = lean_ctor_get(x_432, 1); +lean_inc(x_434); +lean_dec(x_432); +x_435 = lean_ctor_get(x_433, 0); +lean_inc(x_435); +lean_dec(x_433); +x_436 = 0; lean_inc(x_8); -lean_inc(x_387); +lean_inc(x_375); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_416); -x_461 = l_Lean_Compiler_LCNF_normFunDeclImp(x_460, x_416, x_459, x_5, x_6, x_387, x_8, x_458); -if (lean_obj_tag(x_461) == 0) +lean_inc(x_392); +x_437 = l_Lean_Compiler_LCNF_normFunDeclImp(x_436, x_392, x_435, x_5, x_6, x_375, x_8, x_434); +if (lean_obj_tag(x_437) == 0) { -lean_object* x_462; lean_object* x_463; lean_object* x_464; uint8_t x_465; lean_object* x_466; -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_464 = lean_box(0); -x_465 = lean_unbox(x_420); -lean_dec(x_420); -x_466 = l_Lean_Compiler_LCNF_Simp_simp___lambda__3(x_417, x_416, x_1, x_465, x_462, x_464, x_2, x_3, x_4, x_5, x_6, x_387, x_8, x_463); -return x_466; +lean_object* x_438; lean_object* x_439; lean_object* x_440; uint8_t x_441; lean_object* x_442; +x_438 = lean_ctor_get(x_437, 0); +lean_inc(x_438); +x_439 = lean_ctor_get(x_437, 1); +lean_inc(x_439); +lean_dec(x_437); +x_440 = lean_box(0); +x_441 = lean_unbox(x_396); +lean_dec(x_396); +x_442 = l_Lean_Compiler_LCNF_Simp_simp___lambda__4(x_393, x_392, x_1, x_441, x_438, x_440, x_2, x_3, x_4, x_5, x_6, x_375, x_8, x_439); +return x_442; } else { -lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; -lean_dec(x_420); -lean_dec(x_417); -lean_dec(x_416); -lean_dec(x_387); +lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; +lean_dec(x_396); +lean_dec(x_393); +lean_dec(x_392); +lean_dec(x_375); lean_dec(x_8); lean_dec(x_6); lean_dec(x_5); @@ -13457,243 +13424,243 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_467 = lean_ctor_get(x_461, 0); -lean_inc(x_467); -x_468 = lean_ctor_get(x_461, 1); -lean_inc(x_468); -if (lean_is_exclusive(x_461)) { - lean_ctor_release(x_461, 0); - lean_ctor_release(x_461, 1); - x_469 = x_461; +x_443 = lean_ctor_get(x_437, 0); +lean_inc(x_443); +x_444 = lean_ctor_get(x_437, 1); +lean_inc(x_444); +if (lean_is_exclusive(x_437)) { + lean_ctor_release(x_437, 0); + lean_ctor_release(x_437, 1); + x_445 = x_437; } else { - lean_dec_ref(x_461); - x_469 = lean_box(0); + lean_dec_ref(x_437); + x_445 = lean_box(0); } -if (lean_is_scalar(x_469)) { - x_470 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_445)) { + x_446 = lean_alloc_ctor(1, 2, 0); } else { - x_470 = x_469; + x_446 = x_445; } -lean_ctor_set(x_470, 0, x_467); -lean_ctor_set(x_470, 1, x_468); -return x_470; +lean_ctor_set(x_446, 0, x_443); +lean_ctor_set(x_446, 1, x_444); +return x_446; } } } case 2: { -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; -x_471 = lean_ctor_get(x_388, 1); -lean_inc(x_471); -lean_dec(x_388); -x_472 = lean_ctor_get(x_1, 0); -lean_inc(x_472); -x_473 = lean_ctor_get(x_1, 1); -lean_inc(x_473); -x_474 = lean_ctor_get(x_472, 0); -lean_inc(x_474); -x_475 = l_Lean_Compiler_LCNF_Simp_isOnceOrMustInline(x_474, x_2, x_3, x_4, x_5, x_6, x_387, x_8, x_471); -x_476 = lean_ctor_get(x_475, 0); -lean_inc(x_476); -x_477 = lean_unbox(x_476); -if (x_477 == 0) +lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; uint8_t x_453; +x_447 = lean_ctor_get(x_376, 1); +lean_inc(x_447); +lean_dec(x_376); +x_448 = lean_ctor_get(x_1, 0); +lean_inc(x_448); +x_449 = lean_ctor_get(x_1, 1); +lean_inc(x_449); +x_450 = lean_ctor_get(x_448, 0); +lean_inc(x_450); +x_451 = l_Lean_Compiler_LCNF_Simp_isOnceOrMustInline(x_450, x_2, x_3, x_4, x_5, x_6, x_375, x_8, x_447); +x_452 = lean_ctor_get(x_451, 0); +lean_inc(x_452); +x_453 = lean_unbox(x_452); +if (x_453 == 0) { -lean_object* x_478; lean_object* x_479; uint8_t x_480; -x_478 = lean_ctor_get(x_475, 1); -lean_inc(x_478); -lean_dec(x_475); +lean_object* x_454; lean_object* x_455; uint8_t x_456; +x_454 = lean_ctor_get(x_451, 1); +lean_inc(x_454); +lean_dec(x_451); lean_inc(x_1); -lean_inc(x_472); -x_479 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_Simp_simp___lambda__6___boxed), 14, 4); -lean_closure_set(x_479, 0, x_473); -lean_closure_set(x_479, 1, x_472); -lean_closure_set(x_479, 2, x_1); -lean_closure_set(x_479, 3, x_476); -x_480 = l_Lean_Compiler_LCNF_Code_isFun(x_1); +lean_inc(x_448); +x_455 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_Simp_simp___lambda__7___boxed), 14, 4); +lean_closure_set(x_455, 0, x_449); +lean_closure_set(x_455, 1, x_448); +lean_closure_set(x_455, 2, x_1); +lean_closure_set(x_455, 3, x_452); +x_456 = l_Lean_Compiler_LCNF_Code_isFun(x_1); lean_dec(x_1); -if (x_480 == 0) +if (x_456 == 0) { -lean_object* x_481; lean_object* x_482; -x_481 = lean_box(0); -x_482 = l_Lean_Compiler_LCNF_Simp_simp___lambda__4(x_479, x_472, x_481, x_2, x_3, x_4, x_5, x_6, x_387, x_8, x_478); -return x_482; +lean_object* x_457; lean_object* x_458; +x_457 = lean_box(0); +x_458 = l_Lean_Compiler_LCNF_Simp_simp___lambda__5(x_455, x_448, x_457, x_2, x_3, x_4, x_5, x_6, x_375, x_8, x_454); +return x_458; } else { -lean_object* x_483; lean_object* x_484; uint8_t x_485; -x_483 = lean_ctor_get(x_472, 3); -lean_inc(x_483); -x_484 = lean_ctor_get(x_472, 2); -lean_inc(x_484); -x_485 = l_Lean_Compiler_LCNF_isEtaExpandCandidateCore(x_483, x_484); -lean_dec(x_484); -if (x_485 == 0) +lean_object* x_459; lean_object* x_460; uint8_t x_461; +x_459 = lean_ctor_get(x_448, 3); +lean_inc(x_459); +x_460 = lean_ctor_get(x_448, 2); +lean_inc(x_460); +x_461 = l_Lean_Compiler_LCNF_isEtaExpandCandidateCore(x_459, x_460); +lean_dec(x_460); +if (x_461 == 0) { -lean_object* x_486; lean_object* x_487; -x_486 = lean_box(0); -x_487 = l_Lean_Compiler_LCNF_Simp_simp___lambda__4(x_479, x_472, x_486, x_2, x_3, x_4, x_5, x_6, x_387, x_8, x_478); -return x_487; +lean_object* x_462; lean_object* x_463; +x_462 = lean_box(0); +x_463 = l_Lean_Compiler_LCNF_Simp_simp___lambda__5(x_455, x_448, x_462, x_2, x_3, x_4, x_5, x_6, x_375, x_8, x_454); +return x_463; } else { -lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; uint8_t x_492; lean_object* x_493; -x_488 = lean_st_ref_get(x_3, x_478); -x_489 = lean_ctor_get(x_488, 0); -lean_inc(x_489); -x_490 = lean_ctor_get(x_488, 1); -lean_inc(x_490); -lean_dec(x_488); -x_491 = lean_ctor_get(x_489, 0); -lean_inc(x_491); -lean_dec(x_489); -x_492 = 0; +lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; uint8_t x_468; lean_object* x_469; +x_464 = lean_st_ref_get(x_3, x_454); +x_465 = lean_ctor_get(x_464, 0); +lean_inc(x_465); +x_466 = lean_ctor_get(x_464, 1); +lean_inc(x_466); +lean_dec(x_464); +x_467 = lean_ctor_get(x_465, 0); +lean_inc(x_467); +lean_dec(x_465); +x_468 = 0; lean_inc(x_8); -lean_inc(x_387); +lean_inc(x_375); lean_inc(x_6); lean_inc(x_5); -x_493 = l_Lean_Compiler_LCNF_normFunDeclImp(x_492, x_472, x_491, x_5, x_6, x_387, x_8, x_490); -if (lean_obj_tag(x_493) == 0) +x_469 = l_Lean_Compiler_LCNF_normFunDeclImp(x_468, x_448, x_467, x_5, x_6, x_375, x_8, x_466); +if (lean_obj_tag(x_469) == 0) { -lean_object* x_494; lean_object* x_495; lean_object* x_496; -x_494 = lean_ctor_get(x_493, 0); -lean_inc(x_494); -x_495 = lean_ctor_get(x_493, 1); -lean_inc(x_495); -lean_dec(x_493); +lean_object* x_470; lean_object* x_471; lean_object* x_472; +x_470 = lean_ctor_get(x_469, 0); +lean_inc(x_470); +x_471 = lean_ctor_get(x_469, 1); +lean_inc(x_471); +lean_dec(x_469); lean_inc(x_8); -lean_inc(x_387); +lean_inc(x_375); lean_inc(x_6); lean_inc(x_5); -x_496 = l_Lean_Compiler_LCNF_FunDeclCore_etaExpand(x_494, x_5, x_6, x_387, x_8, x_495); -if (lean_obj_tag(x_496) == 0) +x_472 = l_Lean_Compiler_LCNF_FunDeclCore_etaExpand(x_470, x_5, x_6, x_375, x_8, x_471); +if (lean_obj_tag(x_472) == 0) { -lean_object* x_497; lean_object* x_498; lean_object* x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; -x_497 = lean_ctor_get(x_496, 0); -lean_inc(x_497); -x_498 = lean_ctor_get(x_496, 1); -lean_inc(x_498); -lean_dec(x_496); -x_499 = l_Lean_Compiler_LCNF_Simp_markSimplified___rarg(x_3, x_4, x_5, x_6, x_387, x_8, x_498); -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_Compiler_LCNF_Simp_simp___lambda__4(x_479, x_497, x_500, x_2, x_3, x_4, x_5, x_6, x_387, x_8, x_501); -lean_dec(x_500); -return x_502; +lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; +x_473 = lean_ctor_get(x_472, 0); +lean_inc(x_473); +x_474 = lean_ctor_get(x_472, 1); +lean_inc(x_474); +lean_dec(x_472); +x_475 = l_Lean_Compiler_LCNF_Simp_markSimplified___rarg(x_3, x_4, x_5, x_6, x_375, x_8, x_474); +x_476 = lean_ctor_get(x_475, 0); +lean_inc(x_476); +x_477 = lean_ctor_get(x_475, 1); +lean_inc(x_477); +lean_dec(x_475); +x_478 = l_Lean_Compiler_LCNF_Simp_simp___lambda__5(x_455, x_473, x_476, x_2, x_3, x_4, x_5, x_6, x_375, x_8, x_477); +lean_dec(x_476); +return x_478; } else { -lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_object* x_506; -lean_dec(x_479); -lean_dec(x_387); +lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; +lean_dec(x_455); +lean_dec(x_375); 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_503 = lean_ctor_get(x_496, 0); -lean_inc(x_503); -x_504 = lean_ctor_get(x_496, 1); -lean_inc(x_504); -if (lean_is_exclusive(x_496)) { - lean_ctor_release(x_496, 0); - lean_ctor_release(x_496, 1); - x_505 = x_496; +x_479 = lean_ctor_get(x_472, 0); +lean_inc(x_479); +x_480 = lean_ctor_get(x_472, 1); +lean_inc(x_480); +if (lean_is_exclusive(x_472)) { + lean_ctor_release(x_472, 0); + lean_ctor_release(x_472, 1); + x_481 = x_472; } else { - lean_dec_ref(x_496); - x_505 = lean_box(0); + lean_dec_ref(x_472); + x_481 = lean_box(0); } -if (lean_is_scalar(x_505)) { - x_506 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_481)) { + x_482 = lean_alloc_ctor(1, 2, 0); } else { - x_506 = x_505; + x_482 = x_481; } -lean_ctor_set(x_506, 0, x_503); -lean_ctor_set(x_506, 1, x_504); -return x_506; +lean_ctor_set(x_482, 0, x_479); +lean_ctor_set(x_482, 1, x_480); +return x_482; } } else { -lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; -lean_dec(x_479); -lean_dec(x_387); +lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; +lean_dec(x_455); +lean_dec(x_375); 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_507 = lean_ctor_get(x_493, 0); -lean_inc(x_507); -x_508 = lean_ctor_get(x_493, 1); -lean_inc(x_508); -if (lean_is_exclusive(x_493)) { - lean_ctor_release(x_493, 0); - lean_ctor_release(x_493, 1); - x_509 = x_493; +x_483 = lean_ctor_get(x_469, 0); +lean_inc(x_483); +x_484 = lean_ctor_get(x_469, 1); +lean_inc(x_484); +if (lean_is_exclusive(x_469)) { + lean_ctor_release(x_469, 0); + lean_ctor_release(x_469, 1); + x_485 = x_469; } else { - lean_dec_ref(x_493); - x_509 = lean_box(0); + lean_dec_ref(x_469); + x_485 = lean_box(0); } -if (lean_is_scalar(x_509)) { - x_510 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_485)) { + x_486 = lean_alloc_ctor(1, 2, 0); } else { - x_510 = x_509; + x_486 = x_485; } -lean_ctor_set(x_510, 0, x_507); -lean_ctor_set(x_510, 1, x_508); -return x_510; +lean_ctor_set(x_486, 0, x_483); +lean_ctor_set(x_486, 1, x_484); +return x_486; } } } } else { -lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; lean_object* x_515; uint8_t x_516; lean_object* x_517; -x_511 = lean_ctor_get(x_475, 1); -lean_inc(x_511); -lean_dec(x_475); -x_512 = lean_st_ref_get(x_3, x_511); -x_513 = lean_ctor_get(x_512, 0); -lean_inc(x_513); -x_514 = lean_ctor_get(x_512, 1); -lean_inc(x_514); -lean_dec(x_512); -x_515 = lean_ctor_get(x_513, 0); -lean_inc(x_515); -lean_dec(x_513); -x_516 = 0; +lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; uint8_t x_492; lean_object* x_493; +x_487 = lean_ctor_get(x_451, 1); +lean_inc(x_487); +lean_dec(x_451); +x_488 = lean_st_ref_get(x_3, x_487); +x_489 = lean_ctor_get(x_488, 0); +lean_inc(x_489); +x_490 = lean_ctor_get(x_488, 1); +lean_inc(x_490); +lean_dec(x_488); +x_491 = lean_ctor_get(x_489, 0); +lean_inc(x_491); +lean_dec(x_489); +x_492 = 0; lean_inc(x_8); -lean_inc(x_387); +lean_inc(x_375); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_472); -x_517 = l_Lean_Compiler_LCNF_normFunDeclImp(x_516, x_472, x_515, x_5, x_6, x_387, x_8, x_514); -if (lean_obj_tag(x_517) == 0) +lean_inc(x_448); +x_493 = l_Lean_Compiler_LCNF_normFunDeclImp(x_492, x_448, x_491, x_5, x_6, x_375, x_8, x_490); +if (lean_obj_tag(x_493) == 0) { -lean_object* x_518; lean_object* x_519; lean_object* x_520; uint8_t x_521; lean_object* x_522; -x_518 = lean_ctor_get(x_517, 0); -lean_inc(x_518); -x_519 = lean_ctor_get(x_517, 1); -lean_inc(x_519); -lean_dec(x_517); -x_520 = lean_box(0); -x_521 = lean_unbox(x_476); -lean_dec(x_476); -x_522 = l_Lean_Compiler_LCNF_Simp_simp___lambda__6(x_473, x_472, x_1, x_521, x_518, x_520, x_2, x_3, x_4, x_5, x_6, x_387, x_8, x_519); -return x_522; +lean_object* x_494; lean_object* x_495; lean_object* x_496; uint8_t x_497; lean_object* x_498; +x_494 = lean_ctor_get(x_493, 0); +lean_inc(x_494); +x_495 = lean_ctor_get(x_493, 1); +lean_inc(x_495); +lean_dec(x_493); +x_496 = lean_box(0); +x_497 = lean_unbox(x_452); +lean_dec(x_452); +x_498 = l_Lean_Compiler_LCNF_Simp_simp___lambda__7(x_449, x_448, x_1, x_497, x_494, x_496, x_2, x_3, x_4, x_5, x_6, x_375, x_8, x_495); +return x_498; } else { -lean_object* x_523; lean_object* x_524; lean_object* x_525; lean_object* x_526; -lean_dec(x_476); -lean_dec(x_473); -lean_dec(x_472); -lean_dec(x_387); +lean_object* x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; +lean_dec(x_452); +lean_dec(x_449); +lean_dec(x_448); +lean_dec(x_375); lean_dec(x_8); lean_dec(x_6); lean_dec(x_5); @@ -13701,183 +13668,183 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_523 = lean_ctor_get(x_517, 0); -lean_inc(x_523); -x_524 = lean_ctor_get(x_517, 1); -lean_inc(x_524); -if (lean_is_exclusive(x_517)) { - lean_ctor_release(x_517, 0); - lean_ctor_release(x_517, 1); - x_525 = x_517; +x_499 = lean_ctor_get(x_493, 0); +lean_inc(x_499); +x_500 = lean_ctor_get(x_493, 1); +lean_inc(x_500); +if (lean_is_exclusive(x_493)) { + lean_ctor_release(x_493, 0); + lean_ctor_release(x_493, 1); + x_501 = x_493; } else { - lean_dec_ref(x_517); - x_525 = lean_box(0); + lean_dec_ref(x_493); + x_501 = lean_box(0); } -if (lean_is_scalar(x_525)) { - x_526 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_501)) { + x_502 = lean_alloc_ctor(1, 2, 0); } else { - x_526 = x_525; + x_502 = x_501; } -lean_ctor_set(x_526, 0, x_523); -lean_ctor_set(x_526, 1, x_524); -return x_526; +lean_ctor_set(x_502, 0, x_499); +lean_ctor_set(x_502, 1, x_500); +return x_502; } } } case 3: { -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; uint8_t x_534; lean_object* x_535; -x_527 = lean_ctor_get(x_388, 1); -lean_inc(x_527); -lean_dec(x_388); -x_528 = lean_ctor_get(x_1, 0); -lean_inc(x_528); -x_529 = lean_ctor_get(x_1, 1); -lean_inc(x_529); -x_530 = lean_st_ref_get(x_3, x_527); -x_531 = lean_ctor_get(x_530, 0); -lean_inc(x_531); -x_532 = lean_ctor_get(x_530, 1); -lean_inc(x_532); -lean_dec(x_530); -x_533 = lean_ctor_get(x_531, 0); -lean_inc(x_533); -lean_dec(x_531); -x_534 = 0; -lean_inc(x_528); -x_535 = l___private_Lean_Compiler_LCNF_CompilerM_0__Lean_Compiler_LCNF_normFVarImp(x_533, x_528, x_534); -if (lean_obj_tag(x_535) == 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; uint8_t x_510; lean_object* x_511; +x_503 = lean_ctor_get(x_376, 1); +lean_inc(x_503); +lean_dec(x_376); +x_504 = lean_ctor_get(x_1, 0); +lean_inc(x_504); +x_505 = lean_ctor_get(x_1, 1); +lean_inc(x_505); +x_506 = lean_st_ref_get(x_3, x_503); +x_507 = lean_ctor_get(x_506, 0); +lean_inc(x_507); +x_508 = lean_ctor_get(x_506, 1); +lean_inc(x_508); +lean_dec(x_506); +x_509 = lean_ctor_get(x_507, 0); +lean_inc(x_509); +lean_dec(x_507); +x_510 = 0; +lean_inc(x_504); +x_511 = l___private_Lean_Compiler_LCNF_CompilerM_0__Lean_Compiler_LCNF_normFVarImp(x_509, x_504, x_510); +if (lean_obj_tag(x_511) == 0) { -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_554; -x_536 = lean_ctor_get(x_535, 0); -lean_inc(x_536); -lean_dec(x_535); -lean_inc(x_529); -x_537 = l_Lean_Compiler_LCNF_normArgs___at_Lean_Compiler_LCNF_Simp_simp___spec__2(x_534, x_529, x_2, x_3, x_4, x_5, x_6, x_387, x_8, x_532); -x_538 = lean_ctor_get(x_537, 0); -lean_inc(x_538); -x_539 = lean_ctor_get(x_537, 1); -lean_inc(x_539); -if (lean_is_exclusive(x_537)) { - lean_ctor_release(x_537, 0); - lean_ctor_release(x_537, 1); - x_540 = x_537; +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_530; +x_512 = lean_ctor_get(x_511, 0); +lean_inc(x_512); +lean_dec(x_511); +lean_inc(x_505); +x_513 = l_Lean_Compiler_LCNF_normArgs___at_Lean_Compiler_LCNF_Simp_simp___spec__2(x_510, x_505, x_2, x_3, x_4, x_5, x_6, x_375, x_8, x_508); +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; } else { - lean_dec_ref(x_537); - x_540 = lean_box(0); + lean_dec_ref(x_513); + x_516 = lean_box(0); } lean_inc(x_8); -lean_inc(x_387); +lean_inc(x_375); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -lean_inc(x_538); -lean_inc(x_536); -x_554 = l_Lean_Compiler_LCNF_Simp_inlineJp_x3f(x_536, x_538, x_2, x_3, x_4, x_5, x_6, x_387, x_8, x_539); -if (lean_obj_tag(x_554) == 0) +lean_inc(x_514); +lean_inc(x_512); +x_530 = l_Lean_Compiler_LCNF_Simp_inlineJp_x3f(x_512, x_514, x_2, x_3, x_4, x_5, x_6, x_375, x_8, x_515); +if (lean_obj_tag(x_530) == 0) { -lean_object* x_555; -x_555 = lean_ctor_get(x_554, 0); -lean_inc(x_555); -if (lean_obj_tag(x_555) == 0) +lean_object* x_531; +x_531 = lean_ctor_get(x_530, 0); +lean_inc(x_531); +if (lean_obj_tag(x_531) == 0) { -lean_object* x_556; lean_object* x_557; lean_object* x_558; lean_object* x_559; lean_object* x_560; uint8_t x_561; -x_556 = lean_ctor_get(x_554, 1); -lean_inc(x_556); -lean_dec(x_554); -lean_inc(x_536); -x_557 = l_Lean_Compiler_LCNF_Simp_markUsedFVar(x_536, x_2, x_3, x_4, x_5, x_6, x_387, x_8, x_556); -x_558 = lean_ctor_get(x_557, 1); -lean_inc(x_558); -lean_dec(x_557); -x_559 = lean_array_get_size(x_538); -x_560 = lean_unsigned_to_nat(0u); -x_561 = lean_nat_dec_lt(x_560, x_559); -if (x_561 == 0) +lean_object* x_532; lean_object* x_533; lean_object* x_534; lean_object* x_535; lean_object* x_536; uint8_t x_537; +x_532 = lean_ctor_get(x_530, 1); +lean_inc(x_532); +lean_dec(x_530); +lean_inc(x_512); +x_533 = l_Lean_Compiler_LCNF_Simp_markUsedFVar(x_512, x_2, x_3, x_4, x_5, x_6, x_375, x_8, x_532); +x_534 = lean_ctor_get(x_533, 1); +lean_inc(x_534); +lean_dec(x_533); +x_535 = lean_array_get_size(x_514); +x_536 = lean_unsigned_to_nat(0u); +x_537 = lean_nat_dec_lt(x_536, x_535); +if (x_537 == 0) { -lean_dec(x_559); -lean_dec(x_387); +lean_dec(x_535); +lean_dec(x_375); 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_541 = x_558; -goto block_553; +x_517 = x_534; +goto block_529; } else { -uint8_t x_562; -x_562 = lean_nat_dec_le(x_559, x_559); -if (x_562 == 0) +uint8_t x_538; +x_538 = lean_nat_dec_le(x_535, x_535); +if (x_538 == 0) { -lean_dec(x_559); -lean_dec(x_387); +lean_dec(x_535); +lean_dec(x_375); 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_541 = x_558; -goto block_553; +x_517 = x_534; +goto block_529; } else { -size_t x_563; size_t x_564; lean_object* x_565; lean_object* x_566; lean_object* x_567; -x_563 = 0; -x_564 = lean_usize_of_nat(x_559); -lean_dec(x_559); -x_565 = lean_box(0); -x_566 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Simp_markUsedLetValue___spec__1(x_538, x_563, x_564, x_565, x_2, x_3, x_4, x_5, x_6, x_387, x_8, x_558); +size_t x_539; size_t x_540; lean_object* x_541; lean_object* x_542; lean_object* x_543; +x_539 = 0; +x_540 = lean_usize_of_nat(x_535); +lean_dec(x_535); +x_541 = lean_box(0); +x_542 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Simp_markUsedLetValue___spec__1(x_514, x_539, x_540, x_541, x_2, x_3, x_4, x_5, x_6, x_375, x_8, x_534); lean_dec(x_8); -lean_dec(x_387); +lean_dec(x_375); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_567 = lean_ctor_get(x_566, 1); -lean_inc(x_567); -lean_dec(x_566); -x_541 = x_567; -goto block_553; +x_543 = lean_ctor_get(x_542, 1); +lean_inc(x_543); +lean_dec(x_542); +x_517 = x_543; +goto block_529; } } } else { -lean_object* x_568; lean_object* x_569; -lean_dec(x_540); -lean_dec(x_538); -lean_dec(x_536); -lean_dec(x_529); -lean_dec(x_528); +lean_object* x_544; lean_object* x_545; +lean_dec(x_516); +lean_dec(x_514); +lean_dec(x_512); +lean_dec(x_505); +lean_dec(x_504); lean_dec(x_1); -x_568 = lean_ctor_get(x_554, 1); -lean_inc(x_568); -lean_dec(x_554); -x_569 = lean_ctor_get(x_555, 0); -lean_inc(x_569); -lean_dec(x_555); -x_1 = x_569; -x_7 = x_387; -x_9 = x_568; +x_544 = lean_ctor_get(x_530, 1); +lean_inc(x_544); +lean_dec(x_530); +x_545 = lean_ctor_get(x_531, 0); +lean_inc(x_545); +lean_dec(x_531); +x_1 = x_545; +x_7 = x_375; +x_9 = x_544; goto _start; } } else { -lean_object* x_571; lean_object* x_572; lean_object* x_573; lean_object* x_574; -lean_dec(x_540); -lean_dec(x_538); -lean_dec(x_536); -lean_dec(x_529); -lean_dec(x_528); -lean_dec(x_387); +lean_object* x_547; lean_object* x_548; lean_object* x_549; lean_object* x_550; +lean_dec(x_516); +lean_dec(x_514); +lean_dec(x_512); +lean_dec(x_505); +lean_dec(x_504); +lean_dec(x_375); lean_dec(x_8); lean_dec(x_6); lean_dec(x_5); @@ -13885,565 +13852,565 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_571 = lean_ctor_get(x_554, 0); -lean_inc(x_571); -x_572 = lean_ctor_get(x_554, 1); -lean_inc(x_572); -if (lean_is_exclusive(x_554)) { - lean_ctor_release(x_554, 0); - lean_ctor_release(x_554, 1); - x_573 = x_554; +x_547 = lean_ctor_get(x_530, 0); +lean_inc(x_547); +x_548 = lean_ctor_get(x_530, 1); +lean_inc(x_548); +if (lean_is_exclusive(x_530)) { + lean_ctor_release(x_530, 0); + lean_ctor_release(x_530, 1); + x_549 = x_530; } else { - lean_dec_ref(x_554); - x_573 = lean_box(0); + lean_dec_ref(x_530); + x_549 = lean_box(0); } -if (lean_is_scalar(x_573)) { - x_574 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_549)) { + x_550 = lean_alloc_ctor(1, 2, 0); } else { - x_574 = x_573; + x_550 = x_549; } -lean_ctor_set(x_574, 0, x_571); -lean_ctor_set(x_574, 1, x_572); -return x_574; +lean_ctor_set(x_550, 0, x_547); +lean_ctor_set(x_550, 1, x_548); +return x_550; } -block_553: +block_529: { -uint8_t x_542; -x_542 = lean_name_eq(x_528, x_536); -lean_dec(x_528); -if (x_542 == 0) +uint8_t x_518; +x_518 = lean_name_eq(x_504, x_512); +lean_dec(x_504); +if (x_518 == 0) { -lean_object* x_543; lean_object* x_544; lean_object* x_545; -lean_dec(x_529); +lean_object* x_519; lean_object* x_520; lean_object* x_521; +lean_dec(x_505); if (lean_is_exclusive(x_1)) { lean_ctor_release(x_1, 0); lean_ctor_release(x_1, 1); - x_543 = x_1; + x_519 = x_1; } else { lean_dec_ref(x_1); - x_543 = lean_box(0); + x_519 = lean_box(0); } -if (lean_is_scalar(x_543)) { - x_544 = lean_alloc_ctor(3, 2, 0); +if (lean_is_scalar(x_519)) { + x_520 = lean_alloc_ctor(3, 2, 0); } else { - x_544 = x_543; + x_520 = x_519; } -lean_ctor_set(x_544, 0, x_536); -lean_ctor_set(x_544, 1, x_538); -if (lean_is_scalar(x_540)) { - x_545 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_520, 0, x_512); +lean_ctor_set(x_520, 1, x_514); +if (lean_is_scalar(x_516)) { + x_521 = lean_alloc_ctor(0, 2, 0); } else { - x_545 = x_540; + x_521 = x_516; } -lean_ctor_set(x_545, 0, x_544); -lean_ctor_set(x_545, 1, x_541); -return x_545; +lean_ctor_set(x_521, 0, x_520); +lean_ctor_set(x_521, 1, x_517); +return x_521; } else { -size_t x_546; size_t x_547; uint8_t x_548; -x_546 = lean_ptr_addr(x_529); -lean_dec(x_529); -x_547 = lean_ptr_addr(x_538); -x_548 = lean_usize_dec_eq(x_546, x_547); -if (x_548 == 0) +size_t x_522; size_t x_523; uint8_t x_524; +x_522 = lean_ptr_addr(x_505); +lean_dec(x_505); +x_523 = lean_ptr_addr(x_514); +x_524 = lean_usize_dec_eq(x_522, x_523); +if (x_524 == 0) { -lean_object* x_549; lean_object* x_550; lean_object* x_551; +lean_object* x_525; lean_object* x_526; lean_object* x_527; if (lean_is_exclusive(x_1)) { lean_ctor_release(x_1, 0); lean_ctor_release(x_1, 1); - x_549 = x_1; + x_525 = x_1; } else { lean_dec_ref(x_1); - x_549 = lean_box(0); + x_525 = lean_box(0); } -if (lean_is_scalar(x_549)) { - x_550 = lean_alloc_ctor(3, 2, 0); +if (lean_is_scalar(x_525)) { + x_526 = lean_alloc_ctor(3, 2, 0); } else { - x_550 = x_549; + x_526 = x_525; } -lean_ctor_set(x_550, 0, x_536); -lean_ctor_set(x_550, 1, x_538); -if (lean_is_scalar(x_540)) { - x_551 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_526, 0, x_512); +lean_ctor_set(x_526, 1, x_514); +if (lean_is_scalar(x_516)) { + x_527 = lean_alloc_ctor(0, 2, 0); } else { - x_551 = x_540; + x_527 = x_516; } -lean_ctor_set(x_551, 0, x_550); -lean_ctor_set(x_551, 1, x_541); -return x_551; +lean_ctor_set(x_527, 0, x_526); +lean_ctor_set(x_527, 1, x_517); +return x_527; } else { -lean_object* x_552; -lean_dec(x_538); -lean_dec(x_536); -if (lean_is_scalar(x_540)) { - x_552 = lean_alloc_ctor(0, 2, 0); +lean_object* x_528; +lean_dec(x_514); +lean_dec(x_512); +if (lean_is_scalar(x_516)) { + x_528 = lean_alloc_ctor(0, 2, 0); } else { - x_552 = x_540; + x_528 = x_516; } -lean_ctor_set(x_552, 0, x_1); -lean_ctor_set(x_552, 1, x_541); -return x_552; +lean_ctor_set(x_528, 0, x_1); +lean_ctor_set(x_528, 1, x_517); +return x_528; } } } } else { -lean_object* x_575; -lean_dec(x_529); -lean_dec(x_528); +lean_object* x_551; +lean_dec(x_505); +lean_dec(x_504); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_575 = l_Lean_Compiler_LCNF_mkReturnErased(x_5, x_6, x_387, x_8, x_532); +x_551 = l_Lean_Compiler_LCNF_mkReturnErased(x_5, x_6, x_375, x_8, x_508); lean_dec(x_8); -lean_dec(x_387); +lean_dec(x_375); lean_dec(x_6); lean_dec(x_5); -return x_575; +return x_551; } } case 4: { -lean_object* x_576; lean_object* x_577; lean_object* x_578; -x_576 = lean_ctor_get(x_388, 1); -lean_inc(x_576); -lean_dec(x_388); -x_577 = lean_ctor_get(x_1, 0); -lean_inc(x_577); +lean_object* x_552; lean_object* x_553; lean_object* x_554; +x_552 = lean_ctor_get(x_376, 1); +lean_inc(x_552); +lean_dec(x_376); +x_553 = lean_ctor_get(x_1, 0); +lean_inc(x_553); lean_inc(x_8); -lean_inc(x_387); +lean_inc(x_375); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -lean_inc(x_577); -x_578 = l_Lean_Compiler_LCNF_Simp_simpCasesOnCtor_x3f(x_577, x_2, x_3, x_4, x_5, x_6, x_387, x_8, x_576); -if (lean_obj_tag(x_578) == 0) +lean_inc(x_553); +x_554 = l_Lean_Compiler_LCNF_Simp_simpCasesOnCtor_x3f(x_553, x_2, x_3, x_4, x_5, x_6, x_375, x_8, x_552); +if (lean_obj_tag(x_554) == 0) { -lean_object* x_579; -x_579 = lean_ctor_get(x_578, 0); -lean_inc(x_579); -if (lean_obj_tag(x_579) == 0) +lean_object* x_555; +x_555 = lean_ctor_get(x_554, 0); +lean_inc(x_555); +if (lean_obj_tag(x_555) == 0) { -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; uint8_t x_590; lean_object* x_591; -x_580 = lean_ctor_get(x_578, 1); -lean_inc(x_580); -lean_dec(x_578); -x_581 = lean_ctor_get(x_577, 0); -lean_inc(x_581); -x_582 = lean_ctor_get(x_577, 1); -lean_inc(x_582); -x_583 = lean_ctor_get(x_577, 2); -lean_inc(x_583); -x_584 = lean_ctor_get(x_577, 3); -lean_inc(x_584); -if (lean_is_exclusive(x_577)) { - lean_ctor_release(x_577, 0); - lean_ctor_release(x_577, 1); - lean_ctor_release(x_577, 2); - lean_ctor_release(x_577, 3); - x_585 = x_577; +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; uint8_t x_566; lean_object* x_567; +x_556 = lean_ctor_get(x_554, 1); +lean_inc(x_556); +lean_dec(x_554); +x_557 = lean_ctor_get(x_553, 0); +lean_inc(x_557); +x_558 = lean_ctor_get(x_553, 1); +lean_inc(x_558); +x_559 = lean_ctor_get(x_553, 2); +lean_inc(x_559); +x_560 = lean_ctor_get(x_553, 3); +lean_inc(x_560); +if (lean_is_exclusive(x_553)) { + lean_ctor_release(x_553, 0); + lean_ctor_release(x_553, 1); + lean_ctor_release(x_553, 2); + lean_ctor_release(x_553, 3); + x_561 = x_553; } else { - lean_dec_ref(x_577); - x_585 = lean_box(0); -} -x_586 = lean_st_ref_get(x_3, x_580); -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 = lean_ctor_get(x_587, 0); -lean_inc(x_589); -lean_dec(x_587); -x_590 = 0; -lean_inc(x_583); -x_591 = l___private_Lean_Compiler_LCNF_CompilerM_0__Lean_Compiler_LCNF_normFVarImp(x_589, x_583, x_590); -if (lean_obj_tag(x_591) == 0) -{ -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; -x_592 = lean_ctor_get(x_591, 0); -lean_inc(x_592); -lean_dec(x_591); -x_593 = lean_st_ref_get(x_3, x_588); -x_594 = lean_ctor_get(x_593, 0); -lean_inc(x_594); -x_595 = lean_ctor_get(x_593, 1); -lean_inc(x_595); -lean_dec(x_593); -x_596 = lean_ctor_get(x_594, 0); -lean_inc(x_596); -lean_dec(x_594); -lean_inc(x_582); -x_597 = l___private_Lean_Compiler_LCNF_CompilerM_0__Lean_Compiler_LCNF_normExprImp_go(x_596, x_590, x_582); -lean_inc(x_592); -x_598 = l_Lean_Compiler_LCNF_Simp_markUsedFVar(x_592, x_2, x_3, x_4, x_5, x_6, x_387, x_8, x_595); -x_599 = lean_ctor_get(x_598, 1); -lean_inc(x_599); -lean_dec(x_598); -lean_inc(x_592); -x_600 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_Simp_simp___lambda__7), 10, 1); -lean_closure_set(x_600, 0, x_592); + lean_dec_ref(x_553); + x_561 = lean_box(0); +} +x_562 = lean_st_ref_get(x_3, x_556); +x_563 = lean_ctor_get(x_562, 0); +lean_inc(x_563); +x_564 = lean_ctor_get(x_562, 1); +lean_inc(x_564); +lean_dec(x_562); +x_565 = lean_ctor_get(x_563, 0); +lean_inc(x_565); +lean_dec(x_563); +x_566 = 0; +lean_inc(x_559); +x_567 = l___private_Lean_Compiler_LCNF_CompilerM_0__Lean_Compiler_LCNF_normFVarImp(x_565, x_559, x_566); +if (lean_obj_tag(x_567) == 0) +{ +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; +x_568 = lean_ctor_get(x_567, 0); +lean_inc(x_568); +lean_dec(x_567); +x_569 = lean_st_ref_get(x_3, x_564); +x_570 = lean_ctor_get(x_569, 0); +lean_inc(x_570); +x_571 = lean_ctor_get(x_569, 1); +lean_inc(x_571); +lean_dec(x_569); +x_572 = lean_ctor_get(x_570, 0); +lean_inc(x_572); +lean_dec(x_570); +lean_inc(x_558); +x_573 = l___private_Lean_Compiler_LCNF_CompilerM_0__Lean_Compiler_LCNF_normExprImp_go(x_572, x_566, x_558); +lean_inc(x_568); +x_574 = l_Lean_Compiler_LCNF_Simp_markUsedFVar(x_568, x_2, x_3, x_4, x_5, x_6, x_375, x_8, x_571); +x_575 = lean_ctor_get(x_574, 1); +lean_inc(x_575); +lean_dec(x_574); +lean_inc(x_568); +x_576 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_Simp_simp___lambda__8), 10, 1); +lean_closure_set(x_576, 0, x_568); lean_inc(x_8); -lean_inc(x_387); +lean_inc(x_375); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -lean_inc(x_584); -x_601 = l___private_Init_Data_Array_BasicAux_0__mapMonoMImp___at_Lean_Compiler_LCNF_Simp_simp___spec__4(x_584, x_600, x_2, x_3, x_4, x_5, x_6, x_387, x_8, x_599); -if (lean_obj_tag(x_601) == 0) +lean_inc(x_560); +x_577 = l___private_Init_Data_Array_BasicAux_0__mapMonoMImp___at_Lean_Compiler_LCNF_Simp_simp___spec__4(x_560, x_576, x_2, x_3, x_4, x_5, x_6, x_375, x_8, x_575); +if (lean_obj_tag(x_577) == 0) { -lean_object* x_602; lean_object* x_603; lean_object* x_604; lean_object* x_605; -x_602 = lean_ctor_get(x_601, 0); -lean_inc(x_602); -x_603 = lean_ctor_get(x_601, 1); -lean_inc(x_603); -if (lean_is_exclusive(x_601)) { - lean_ctor_release(x_601, 0); - lean_ctor_release(x_601, 1); - x_604 = x_601; +lean_object* x_578; lean_object* x_579; lean_object* x_580; lean_object* x_581; +x_578 = lean_ctor_get(x_577, 0); +lean_inc(x_578); +x_579 = lean_ctor_get(x_577, 1); +lean_inc(x_579); +if (lean_is_exclusive(x_577)) { + lean_ctor_release(x_577, 0); + lean_ctor_release(x_577, 1); + x_580 = x_577; } else { - lean_dec_ref(x_601); - x_604 = lean_box(0); -} -x_605 = l_Lean_Compiler_LCNF_Simp_addDefaultAlt(x_602, x_2, x_3, x_4, x_5, x_6, x_387, x_8, x_603); -if (lean_obj_tag(x_605) == 0) -{ -lean_object* x_606; lean_object* x_607; lean_object* x_608; lean_object* x_609; lean_object* x_631; lean_object* x_632; uint8_t x_643; -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; + lean_dec_ref(x_577); + x_580 = lean_box(0); +} +x_581 = l_Lean_Compiler_LCNF_Simp_addDefaultAlt(x_578, x_2, x_3, x_4, x_5, x_6, x_375, x_8, x_579); +if (lean_obj_tag(x_581) == 0) +{ +lean_object* x_582; lean_object* x_583; lean_object* x_584; lean_object* x_585; lean_object* x_607; lean_object* x_608; uint8_t x_619; +x_582 = lean_ctor_get(x_581, 0); +lean_inc(x_582); +x_583 = lean_ctor_get(x_581, 1); +lean_inc(x_583); +if (lean_is_exclusive(x_581)) { + lean_ctor_release(x_581, 0); + lean_ctor_release(x_581, 1); + x_584 = x_581; } else { - lean_dec_ref(x_605); - x_608 = lean_box(0); + lean_dec_ref(x_581); + x_584 = lean_box(0); } -x_631 = lean_array_get_size(x_606); -x_643 = lean_nat_dec_eq(x_631, x_385); -if (x_643 == 0) +x_607 = lean_array_get_size(x_582); +x_619 = lean_nat_dec_eq(x_607, x_373); +if (x_619 == 0) { -lean_object* x_644; -lean_dec(x_631); -lean_dec(x_604); -x_644 = lean_box(0); -x_609 = x_644; -goto block_630; +lean_object* x_620; +lean_dec(x_607); +lean_dec(x_580); +x_620 = lean_box(0); +x_585 = x_620; +goto block_606; } else { -lean_object* x_645; uint8_t x_646; -x_645 = lean_unsigned_to_nat(0u); -x_646 = lean_nat_dec_lt(x_645, x_631); -if (x_646 == 0) +lean_object* x_621; uint8_t x_622; +x_621 = lean_unsigned_to_nat(0u); +x_622 = lean_nat_dec_lt(x_621, x_607); +if (x_622 == 0) { -lean_object* x_647; lean_object* x_648; -x_647 = l___private_Lean_Compiler_LCNF_Simp_Main_0__Lean_Compiler_LCNF_Simp_oneExitPointQuick_go___closed__1; -x_648 = l___private_Init_Util_0__outOfBounds___rarg(x_647); -if (lean_obj_tag(x_648) == 0) +lean_object* x_623; lean_object* x_624; +x_623 = l___private_Lean_Compiler_LCNF_Simp_Main_0__Lean_Compiler_LCNF_Simp_oneExitPointQuick_go___closed__1; +x_624 = l___private_Init_Util_0__outOfBounds___rarg(x_623); +if (lean_obj_tag(x_624) == 0) { -lean_object* x_649; -lean_dec(x_648); -lean_dec(x_631); -lean_dec(x_604); -x_649 = lean_box(0); -x_609 = x_649; -goto block_630; +lean_object* x_625; +lean_dec(x_624); +lean_dec(x_607); +lean_dec(x_580); +x_625 = lean_box(0); +x_585 = x_625; +goto block_606; } else { -lean_object* x_650; -lean_dec(x_648); -lean_dec(x_608); -lean_dec(x_597); -lean_dec(x_592); -lean_dec(x_585); +lean_object* x_626; +lean_dec(x_624); lean_dec(x_584); -lean_dec(x_583); -lean_dec(x_582); -lean_dec(x_581); +lean_dec(x_573); +lean_dec(x_568); +lean_dec(x_561); +lean_dec(x_560); +lean_dec(x_559); +lean_dec(x_558); +lean_dec(x_557); lean_dec(x_1); -x_650 = lean_box(0); -x_632 = x_650; -goto block_642; +x_626 = lean_box(0); +x_608 = x_626; +goto block_618; } } else { -lean_object* x_651; -x_651 = lean_array_fget(x_606, x_645); -if (lean_obj_tag(x_651) == 0) +lean_object* x_627; +x_627 = lean_array_fget(x_582, x_621); +if (lean_obj_tag(x_627) == 0) { -lean_object* x_652; -lean_dec(x_651); -lean_dec(x_631); -lean_dec(x_604); -x_652 = lean_box(0); -x_609 = x_652; -goto block_630; +lean_object* x_628; +lean_dec(x_627); +lean_dec(x_607); +lean_dec(x_580); +x_628 = lean_box(0); +x_585 = x_628; +goto block_606; } else { -lean_object* x_653; -lean_dec(x_651); -lean_dec(x_608); -lean_dec(x_597); -lean_dec(x_592); -lean_dec(x_585); +lean_object* x_629; +lean_dec(x_627); lean_dec(x_584); -lean_dec(x_583); -lean_dec(x_582); -lean_dec(x_581); +lean_dec(x_573); +lean_dec(x_568); +lean_dec(x_561); +lean_dec(x_560); +lean_dec(x_559); +lean_dec(x_558); +lean_dec(x_557); lean_dec(x_1); -x_653 = lean_box(0); -x_632 = x_653; -goto block_642; +x_629 = lean_box(0); +x_608 = x_629; +goto block_618; } } } -block_630: +block_606: { -size_t x_610; size_t x_611; uint8_t x_612; -lean_dec(x_609); -x_610 = lean_ptr_addr(x_584); -lean_dec(x_584); -x_611 = lean_ptr_addr(x_606); -x_612 = lean_usize_dec_eq(x_610, x_611); -if (x_612 == 0) +size_t x_586; size_t x_587; uint8_t x_588; +lean_dec(x_585); +x_586 = lean_ptr_addr(x_560); +lean_dec(x_560); +x_587 = lean_ptr_addr(x_582); +x_588 = lean_usize_dec_eq(x_586, x_587); +if (x_588 == 0) { -lean_object* x_613; lean_object* x_614; lean_object* x_615; lean_object* x_616; -lean_dec(x_583); -lean_dec(x_582); +lean_object* x_589; lean_object* x_590; lean_object* x_591; lean_object* x_592; +lean_dec(x_559); +lean_dec(x_558); if (lean_is_exclusive(x_1)) { lean_ctor_release(x_1, 0); - x_613 = x_1; + x_589 = x_1; } else { lean_dec_ref(x_1); - x_613 = lean_box(0); + x_589 = lean_box(0); } -if (lean_is_scalar(x_585)) { - x_614 = lean_alloc_ctor(0, 4, 0); +if (lean_is_scalar(x_561)) { + x_590 = lean_alloc_ctor(0, 4, 0); } else { - x_614 = x_585; -} -lean_ctor_set(x_614, 0, x_581); -lean_ctor_set(x_614, 1, x_597); -lean_ctor_set(x_614, 2, x_592); -lean_ctor_set(x_614, 3, x_606); -if (lean_is_scalar(x_613)) { - x_615 = lean_alloc_ctor(4, 1, 0); + x_590 = x_561; +} +lean_ctor_set(x_590, 0, x_557); +lean_ctor_set(x_590, 1, x_573); +lean_ctor_set(x_590, 2, x_568); +lean_ctor_set(x_590, 3, x_582); +if (lean_is_scalar(x_589)) { + x_591 = lean_alloc_ctor(4, 1, 0); } else { - x_615 = x_613; + x_591 = x_589; } -lean_ctor_set(x_615, 0, x_614); -if (lean_is_scalar(x_608)) { - x_616 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_591, 0, x_590); +if (lean_is_scalar(x_584)) { + x_592 = lean_alloc_ctor(0, 2, 0); } else { - x_616 = x_608; + x_592 = x_584; } -lean_ctor_set(x_616, 0, x_615); -lean_ctor_set(x_616, 1, x_607); -return x_616; +lean_ctor_set(x_592, 0, x_591); +lean_ctor_set(x_592, 1, x_583); +return x_592; } else { -size_t x_617; size_t x_618; uint8_t x_619; -x_617 = lean_ptr_addr(x_582); -lean_dec(x_582); -x_618 = lean_ptr_addr(x_597); -x_619 = lean_usize_dec_eq(x_617, x_618); -if (x_619 == 0) +size_t x_593; size_t x_594; uint8_t x_595; +x_593 = lean_ptr_addr(x_558); +lean_dec(x_558); +x_594 = lean_ptr_addr(x_573); +x_595 = lean_usize_dec_eq(x_593, x_594); +if (x_595 == 0) { -lean_object* x_620; lean_object* x_621; lean_object* x_622; lean_object* x_623; -lean_dec(x_583); +lean_object* x_596; lean_object* x_597; lean_object* x_598; lean_object* x_599; +lean_dec(x_559); if (lean_is_exclusive(x_1)) { lean_ctor_release(x_1, 0); - x_620 = x_1; + x_596 = x_1; } else { lean_dec_ref(x_1); - x_620 = lean_box(0); + x_596 = lean_box(0); } -if (lean_is_scalar(x_585)) { - x_621 = lean_alloc_ctor(0, 4, 0); +if (lean_is_scalar(x_561)) { + x_597 = lean_alloc_ctor(0, 4, 0); } else { - x_621 = x_585; -} -lean_ctor_set(x_621, 0, x_581); -lean_ctor_set(x_621, 1, x_597); -lean_ctor_set(x_621, 2, x_592); -lean_ctor_set(x_621, 3, x_606); -if (lean_is_scalar(x_620)) { - x_622 = lean_alloc_ctor(4, 1, 0); + x_597 = x_561; +} +lean_ctor_set(x_597, 0, x_557); +lean_ctor_set(x_597, 1, x_573); +lean_ctor_set(x_597, 2, x_568); +lean_ctor_set(x_597, 3, x_582); +if (lean_is_scalar(x_596)) { + x_598 = lean_alloc_ctor(4, 1, 0); } else { - x_622 = x_620; + x_598 = x_596; } -lean_ctor_set(x_622, 0, x_621); -if (lean_is_scalar(x_608)) { - x_623 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_598, 0, x_597); +if (lean_is_scalar(x_584)) { + x_599 = lean_alloc_ctor(0, 2, 0); } else { - x_623 = x_608; + x_599 = x_584; } -lean_ctor_set(x_623, 0, x_622); -lean_ctor_set(x_623, 1, x_607); -return x_623; +lean_ctor_set(x_599, 0, x_598); +lean_ctor_set(x_599, 1, x_583); +return x_599; } else { -uint8_t x_624; -x_624 = lean_name_eq(x_583, x_592); -lean_dec(x_583); -if (x_624 == 0) +uint8_t x_600; +x_600 = lean_name_eq(x_559, x_568); +lean_dec(x_559); +if (x_600 == 0) { -lean_object* x_625; lean_object* x_626; lean_object* x_627; lean_object* x_628; +lean_object* x_601; lean_object* x_602; lean_object* x_603; lean_object* x_604; if (lean_is_exclusive(x_1)) { lean_ctor_release(x_1, 0); - x_625 = x_1; + x_601 = x_1; } else { lean_dec_ref(x_1); - x_625 = lean_box(0); + x_601 = lean_box(0); } -if (lean_is_scalar(x_585)) { - x_626 = lean_alloc_ctor(0, 4, 0); +if (lean_is_scalar(x_561)) { + x_602 = lean_alloc_ctor(0, 4, 0); } else { - x_626 = x_585; -} -lean_ctor_set(x_626, 0, x_581); -lean_ctor_set(x_626, 1, x_597); -lean_ctor_set(x_626, 2, x_592); -lean_ctor_set(x_626, 3, x_606); -if (lean_is_scalar(x_625)) { - x_627 = lean_alloc_ctor(4, 1, 0); + x_602 = x_561; +} +lean_ctor_set(x_602, 0, x_557); +lean_ctor_set(x_602, 1, x_573); +lean_ctor_set(x_602, 2, x_568); +lean_ctor_set(x_602, 3, x_582); +if (lean_is_scalar(x_601)) { + x_603 = lean_alloc_ctor(4, 1, 0); } else { - x_627 = x_625; + x_603 = x_601; } -lean_ctor_set(x_627, 0, x_626); -if (lean_is_scalar(x_608)) { - x_628 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_603, 0, x_602); +if (lean_is_scalar(x_584)) { + x_604 = lean_alloc_ctor(0, 2, 0); } else { - x_628 = x_608; + x_604 = x_584; } -lean_ctor_set(x_628, 0, x_627); -lean_ctor_set(x_628, 1, x_607); -return x_628; +lean_ctor_set(x_604, 0, x_603); +lean_ctor_set(x_604, 1, x_583); +return x_604; } else { -lean_object* x_629; -lean_dec(x_606); -lean_dec(x_597); -lean_dec(x_592); -lean_dec(x_585); -lean_dec(x_581); -if (lean_is_scalar(x_608)) { - x_629 = lean_alloc_ctor(0, 2, 0); +lean_object* x_605; +lean_dec(x_582); +lean_dec(x_573); +lean_dec(x_568); +lean_dec(x_561); +lean_dec(x_557); +if (lean_is_scalar(x_584)) { + x_605 = lean_alloc_ctor(0, 2, 0); } else { - x_629 = x_608; + x_605 = x_584; } -lean_ctor_set(x_629, 0, x_1); -lean_ctor_set(x_629, 1, x_607); -return x_629; +lean_ctor_set(x_605, 0, x_1); +lean_ctor_set(x_605, 1, x_583); +return x_605; } } } } -block_642: +block_618: { -lean_object* x_633; uint8_t x_634; -lean_dec(x_632); -x_633 = lean_unsigned_to_nat(0u); -x_634 = lean_nat_dec_lt(x_633, x_631); -lean_dec(x_631); -if (x_634 == 0) +lean_object* x_609; uint8_t x_610; +lean_dec(x_608); +x_609 = lean_unsigned_to_nat(0u); +x_610 = lean_nat_dec_lt(x_609, x_607); +lean_dec(x_607); +if (x_610 == 0) { -lean_object* x_635; lean_object* x_636; lean_object* x_637; lean_object* x_638; -lean_dec(x_606); -x_635 = l___private_Lean_Compiler_LCNF_Simp_Main_0__Lean_Compiler_LCNF_Simp_oneExitPointQuick_go___closed__1; -x_636 = l___private_Init_Util_0__outOfBounds___rarg(x_635); -x_637 = l_Lean_Compiler_LCNF_AltCore_getCode(x_636); -lean_dec(x_636); -if (lean_is_scalar(x_604)) { - x_638 = lean_alloc_ctor(0, 2, 0); +lean_object* x_611; lean_object* x_612; lean_object* x_613; lean_object* x_614; +lean_dec(x_582); +x_611 = l___private_Lean_Compiler_LCNF_Simp_Main_0__Lean_Compiler_LCNF_Simp_oneExitPointQuick_go___closed__1; +x_612 = l___private_Init_Util_0__outOfBounds___rarg(x_611); +x_613 = l_Lean_Compiler_LCNF_AltCore_getCode(x_612); +lean_dec(x_612); +if (lean_is_scalar(x_580)) { + x_614 = lean_alloc_ctor(0, 2, 0); } else { - x_638 = x_604; + x_614 = x_580; } -lean_ctor_set(x_638, 0, x_637); -lean_ctor_set(x_638, 1, x_607); -return x_638; +lean_ctor_set(x_614, 0, x_613); +lean_ctor_set(x_614, 1, x_583); +return x_614; } else { -lean_object* x_639; lean_object* x_640; lean_object* x_641; -x_639 = lean_array_fget(x_606, x_633); -lean_dec(x_606); -x_640 = l_Lean_Compiler_LCNF_AltCore_getCode(x_639); -lean_dec(x_639); -if (lean_is_scalar(x_604)) { - x_641 = lean_alloc_ctor(0, 2, 0); +lean_object* x_615; lean_object* x_616; lean_object* x_617; +x_615 = lean_array_fget(x_582, x_609); +lean_dec(x_582); +x_616 = l_Lean_Compiler_LCNF_AltCore_getCode(x_615); +lean_dec(x_615); +if (lean_is_scalar(x_580)) { + x_617 = lean_alloc_ctor(0, 2, 0); } else { - x_641 = x_604; + x_617 = x_580; } -lean_ctor_set(x_641, 0, x_640); -lean_ctor_set(x_641, 1, x_607); -return x_641; +lean_ctor_set(x_617, 0, x_616); +lean_ctor_set(x_617, 1, x_583); +return x_617; } } } else { -lean_object* x_654; lean_object* x_655; lean_object* x_656; lean_object* x_657; -lean_dec(x_604); -lean_dec(x_597); -lean_dec(x_592); -lean_dec(x_585); -lean_dec(x_584); -lean_dec(x_583); -lean_dec(x_582); -lean_dec(x_581); +lean_object* x_630; lean_object* x_631; lean_object* x_632; lean_object* x_633; +lean_dec(x_580); +lean_dec(x_573); +lean_dec(x_568); +lean_dec(x_561); +lean_dec(x_560); +lean_dec(x_559); +lean_dec(x_558); +lean_dec(x_557); lean_dec(x_1); -x_654 = lean_ctor_get(x_605, 0); -lean_inc(x_654); -x_655 = lean_ctor_get(x_605, 1); -lean_inc(x_655); -if (lean_is_exclusive(x_605)) { - lean_ctor_release(x_605, 0); - lean_ctor_release(x_605, 1); - x_656 = x_605; +x_630 = lean_ctor_get(x_581, 0); +lean_inc(x_630); +x_631 = lean_ctor_get(x_581, 1); +lean_inc(x_631); +if (lean_is_exclusive(x_581)) { + lean_ctor_release(x_581, 0); + lean_ctor_release(x_581, 1); + x_632 = x_581; } else { - lean_dec_ref(x_605); - x_656 = lean_box(0); + lean_dec_ref(x_581); + x_632 = lean_box(0); } -if (lean_is_scalar(x_656)) { - x_657 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_632)) { + x_633 = lean_alloc_ctor(1, 2, 0); } else { - x_657 = x_656; + x_633 = x_632; } -lean_ctor_set(x_657, 0, x_654); -lean_ctor_set(x_657, 1, x_655); -return x_657; +lean_ctor_set(x_633, 0, x_630); +lean_ctor_set(x_633, 1, x_631); +return x_633; } } else { -lean_object* x_658; lean_object* x_659; lean_object* x_660; lean_object* x_661; -lean_dec(x_597); -lean_dec(x_592); -lean_dec(x_585); -lean_dec(x_584); -lean_dec(x_583); -lean_dec(x_582); -lean_dec(x_581); -lean_dec(x_387); +lean_object* x_634; lean_object* x_635; lean_object* x_636; lean_object* x_637; +lean_dec(x_573); +lean_dec(x_568); +lean_dec(x_561); +lean_dec(x_560); +lean_dec(x_559); +lean_dec(x_558); +lean_dec(x_557); +lean_dec(x_375); lean_dec(x_8); lean_dec(x_6); lean_dec(x_5); @@ -14451,53 +14418,53 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_658 = lean_ctor_get(x_601, 0); -lean_inc(x_658); -x_659 = lean_ctor_get(x_601, 1); -lean_inc(x_659); -if (lean_is_exclusive(x_601)) { - lean_ctor_release(x_601, 0); - lean_ctor_release(x_601, 1); - x_660 = x_601; +x_634 = lean_ctor_get(x_577, 0); +lean_inc(x_634); +x_635 = lean_ctor_get(x_577, 1); +lean_inc(x_635); +if (lean_is_exclusive(x_577)) { + lean_ctor_release(x_577, 0); + lean_ctor_release(x_577, 1); + x_636 = x_577; } else { - lean_dec_ref(x_601); - x_660 = lean_box(0); + lean_dec_ref(x_577); + x_636 = lean_box(0); } -if (lean_is_scalar(x_660)) { - x_661 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_636)) { + x_637 = lean_alloc_ctor(1, 2, 0); } else { - x_661 = x_660; + x_637 = x_636; } -lean_ctor_set(x_661, 0, x_658); -lean_ctor_set(x_661, 1, x_659); -return x_661; +lean_ctor_set(x_637, 0, x_634); +lean_ctor_set(x_637, 1, x_635); +return x_637; } } else { -lean_object* x_662; -lean_dec(x_585); -lean_dec(x_584); -lean_dec(x_583); -lean_dec(x_582); -lean_dec(x_581); +lean_object* x_638; +lean_dec(x_561); +lean_dec(x_560); +lean_dec(x_559); +lean_dec(x_558); +lean_dec(x_557); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_662 = l_Lean_Compiler_LCNF_mkReturnErased(x_5, x_6, x_387, x_8, x_588); +x_638 = l_Lean_Compiler_LCNF_mkReturnErased(x_5, x_6, x_375, x_8, x_564); lean_dec(x_8); -lean_dec(x_387); +lean_dec(x_375); lean_dec(x_6); lean_dec(x_5); -return x_662; +return x_638; } } else { -lean_object* x_663; lean_object* x_664; lean_object* x_665; lean_object* x_666; -lean_dec(x_577); -lean_dec(x_387); +lean_object* x_639; lean_object* x_640; lean_object* x_641; lean_object* x_642; +lean_dec(x_553); +lean_dec(x_375); lean_dec(x_8); lean_dec(x_6); lean_dec(x_5); @@ -14505,34 +14472,34 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_663 = lean_ctor_get(x_578, 1); -lean_inc(x_663); -if (lean_is_exclusive(x_578)) { - lean_ctor_release(x_578, 0); - lean_ctor_release(x_578, 1); - x_664 = x_578; +x_639 = lean_ctor_get(x_554, 1); +lean_inc(x_639); +if (lean_is_exclusive(x_554)) { + lean_ctor_release(x_554, 0); + lean_ctor_release(x_554, 1); + x_640 = x_554; } else { - lean_dec_ref(x_578); - x_664 = lean_box(0); + lean_dec_ref(x_554); + x_640 = lean_box(0); } -x_665 = lean_ctor_get(x_579, 0); -lean_inc(x_665); -lean_dec(x_579); -if (lean_is_scalar(x_664)) { - x_666 = lean_alloc_ctor(0, 2, 0); +x_641 = lean_ctor_get(x_555, 0); +lean_inc(x_641); +lean_dec(x_555); +if (lean_is_scalar(x_640)) { + x_642 = lean_alloc_ctor(0, 2, 0); } else { - x_666 = x_664; + x_642 = x_640; } -lean_ctor_set(x_666, 0, x_665); -lean_ctor_set(x_666, 1, x_663); -return x_666; +lean_ctor_set(x_642, 0, x_641); +lean_ctor_set(x_642, 1, x_639); +return x_642; } } else { -lean_object* x_667; lean_object* x_668; lean_object* x_669; lean_object* x_670; -lean_dec(x_577); -lean_dec(x_387); +lean_object* x_643; lean_object* x_644; lean_object* x_645; lean_object* x_646; +lean_dec(x_553); +lean_dec(x_375); lean_dec(x_8); lean_dec(x_6); lean_dec(x_5); @@ -14540,205 +14507,205 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_667 = lean_ctor_get(x_578, 0); -lean_inc(x_667); -x_668 = lean_ctor_get(x_578, 1); -lean_inc(x_668); -if (lean_is_exclusive(x_578)) { - lean_ctor_release(x_578, 0); - lean_ctor_release(x_578, 1); - x_669 = x_578; +x_643 = lean_ctor_get(x_554, 0); +lean_inc(x_643); +x_644 = lean_ctor_get(x_554, 1); +lean_inc(x_644); +if (lean_is_exclusive(x_554)) { + lean_ctor_release(x_554, 0); + lean_ctor_release(x_554, 1); + x_645 = x_554; } else { - lean_dec_ref(x_578); - x_669 = lean_box(0); + lean_dec_ref(x_554); + x_645 = lean_box(0); } -if (lean_is_scalar(x_669)) { - x_670 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_645)) { + x_646 = lean_alloc_ctor(1, 2, 0); } else { - x_670 = x_669; + x_646 = x_645; } -lean_ctor_set(x_670, 0, x_667); -lean_ctor_set(x_670, 1, x_668); -return x_670; +lean_ctor_set(x_646, 0, x_643); +lean_ctor_set(x_646, 1, x_644); +return x_646; } } case 5: { -lean_object* x_671; lean_object* x_672; lean_object* x_673; lean_object* x_674; lean_object* x_675; lean_object* x_676; uint8_t x_677; lean_object* x_678; -x_671 = lean_ctor_get(x_388, 1); -lean_inc(x_671); -lean_dec(x_388); -x_672 = lean_ctor_get(x_1, 0); -lean_inc(x_672); -x_673 = lean_st_ref_get(x_3, x_671); -x_674 = lean_ctor_get(x_673, 0); -lean_inc(x_674); -x_675 = lean_ctor_get(x_673, 1); -lean_inc(x_675); -lean_dec(x_673); -x_676 = lean_ctor_get(x_674, 0); -lean_inc(x_676); -lean_dec(x_674); -x_677 = 0; -lean_inc(x_672); -x_678 = l___private_Lean_Compiler_LCNF_CompilerM_0__Lean_Compiler_LCNF_normFVarImp(x_676, x_672, x_677); -if (lean_obj_tag(x_678) == 0) -{ -lean_object* x_679; lean_object* x_680; lean_object* x_681; lean_object* x_682; uint8_t x_683; -x_679 = lean_ctor_get(x_678, 0); -lean_inc(x_679); -lean_dec(x_678); -lean_inc(x_679); -x_680 = l_Lean_Compiler_LCNF_Simp_markUsedFVar(x_679, x_2, x_3, x_4, x_5, x_6, x_387, x_8, x_675); +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; lean_object* x_654; +x_647 = lean_ctor_get(x_376, 1); +lean_inc(x_647); +lean_dec(x_376); +x_648 = lean_ctor_get(x_1, 0); +lean_inc(x_648); +x_649 = lean_st_ref_get(x_3, x_647); +x_650 = lean_ctor_get(x_649, 0); +lean_inc(x_650); +x_651 = lean_ctor_get(x_649, 1); +lean_inc(x_651); +lean_dec(x_649); +x_652 = lean_ctor_get(x_650, 0); +lean_inc(x_652); +lean_dec(x_650); +x_653 = 0; +lean_inc(x_648); +x_654 = l___private_Lean_Compiler_LCNF_CompilerM_0__Lean_Compiler_LCNF_normFVarImp(x_652, x_648, x_653); +if (lean_obj_tag(x_654) == 0) +{ +lean_object* x_655; lean_object* x_656; lean_object* x_657; lean_object* x_658; uint8_t x_659; +x_655 = lean_ctor_get(x_654, 0); +lean_inc(x_655); +lean_dec(x_654); +lean_inc(x_655); +x_656 = l_Lean_Compiler_LCNF_Simp_markUsedFVar(x_655, x_2, x_3, x_4, x_5, x_6, x_375, x_8, x_651); lean_dec(x_8); -lean_dec(x_387); +lean_dec(x_375); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_681 = lean_ctor_get(x_680, 1); -lean_inc(x_681); -if (lean_is_exclusive(x_680)) { - lean_ctor_release(x_680, 0); - lean_ctor_release(x_680, 1); - x_682 = x_680; +x_657 = lean_ctor_get(x_656, 1); +lean_inc(x_657); +if (lean_is_exclusive(x_656)) { + lean_ctor_release(x_656, 0); + lean_ctor_release(x_656, 1); + x_658 = x_656; } else { - lean_dec_ref(x_680); - x_682 = lean_box(0); + lean_dec_ref(x_656); + x_658 = lean_box(0); } -x_683 = lean_name_eq(x_672, x_679); -lean_dec(x_672); -if (x_683 == 0) +x_659 = lean_name_eq(x_648, x_655); +lean_dec(x_648); +if (x_659 == 0) { -lean_object* x_684; lean_object* x_685; lean_object* x_686; +lean_object* x_660; lean_object* x_661; lean_object* x_662; if (lean_is_exclusive(x_1)) { lean_ctor_release(x_1, 0); - x_684 = x_1; + x_660 = x_1; } else { lean_dec_ref(x_1); - x_684 = lean_box(0); + x_660 = lean_box(0); } -if (lean_is_scalar(x_684)) { - x_685 = lean_alloc_ctor(5, 1, 0); +if (lean_is_scalar(x_660)) { + x_661 = lean_alloc_ctor(5, 1, 0); } else { - x_685 = x_684; + x_661 = x_660; } -lean_ctor_set(x_685, 0, x_679); -if (lean_is_scalar(x_682)) { - x_686 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_661, 0, x_655); +if (lean_is_scalar(x_658)) { + x_662 = lean_alloc_ctor(0, 2, 0); } else { - x_686 = x_682; + x_662 = x_658; } -lean_ctor_set(x_686, 0, x_685); -lean_ctor_set(x_686, 1, x_681); -return x_686; +lean_ctor_set(x_662, 0, x_661); +lean_ctor_set(x_662, 1, x_657); +return x_662; } else { -lean_object* x_687; -lean_dec(x_679); -if (lean_is_scalar(x_682)) { - x_687 = lean_alloc_ctor(0, 2, 0); +lean_object* x_663; +lean_dec(x_655); +if (lean_is_scalar(x_658)) { + x_663 = lean_alloc_ctor(0, 2, 0); } else { - x_687 = x_682; + x_663 = x_658; } -lean_ctor_set(x_687, 0, x_1); -lean_ctor_set(x_687, 1, x_681); -return x_687; +lean_ctor_set(x_663, 0, x_1); +lean_ctor_set(x_663, 1, x_657); +return x_663; } } else { -lean_object* x_688; -lean_dec(x_672); +lean_object* x_664; +lean_dec(x_648); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_688 = l_Lean_Compiler_LCNF_mkReturnErased(x_5, x_6, x_387, x_8, x_675); +x_664 = l_Lean_Compiler_LCNF_mkReturnErased(x_5, x_6, x_375, x_8, x_651); lean_dec(x_8); -lean_dec(x_387); +lean_dec(x_375); lean_dec(x_6); lean_dec(x_5); -return x_688; +return x_664; } } default: { -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; uint8_t x_696; lean_object* x_697; size_t x_698; size_t x_699; uint8_t x_700; -lean_dec(x_387); +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_object* x_673; size_t x_674; size_t x_675; uint8_t x_676; +lean_dec(x_375); lean_dec(x_8); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); -x_689 = lean_ctor_get(x_388, 1); -lean_inc(x_689); -lean_dec(x_388); -x_690 = lean_ctor_get(x_1, 0); -lean_inc(x_690); -x_691 = lean_st_ref_get(x_3, x_689); +x_665 = lean_ctor_get(x_376, 1); +lean_inc(x_665); +lean_dec(x_376); +x_666 = lean_ctor_get(x_1, 0); +lean_inc(x_666); +x_667 = lean_st_ref_get(x_3, x_665); lean_dec(x_3); -x_692 = lean_ctor_get(x_691, 0); -lean_inc(x_692); -x_693 = lean_ctor_get(x_691, 1); -lean_inc(x_693); -if (lean_is_exclusive(x_691)) { - lean_ctor_release(x_691, 0); - lean_ctor_release(x_691, 1); - x_694 = x_691; +x_668 = lean_ctor_get(x_667, 0); +lean_inc(x_668); +x_669 = lean_ctor_get(x_667, 1); +lean_inc(x_669); +if (lean_is_exclusive(x_667)) { + lean_ctor_release(x_667, 0); + lean_ctor_release(x_667, 1); + x_670 = x_667; } else { - lean_dec_ref(x_691); - x_694 = lean_box(0); -} -x_695 = lean_ctor_get(x_692, 0); -lean_inc(x_695); -lean_dec(x_692); -x_696 = 0; -lean_inc(x_690); -x_697 = l___private_Lean_Compiler_LCNF_CompilerM_0__Lean_Compiler_LCNF_normExprImp_go(x_695, x_696, x_690); -x_698 = lean_ptr_addr(x_690); -lean_dec(x_690); -x_699 = lean_ptr_addr(x_697); -x_700 = lean_usize_dec_eq(x_698, x_699); -if (x_700 == 0) + lean_dec_ref(x_667); + x_670 = lean_box(0); +} +x_671 = lean_ctor_get(x_668, 0); +lean_inc(x_671); +lean_dec(x_668); +x_672 = 0; +lean_inc(x_666); +x_673 = l___private_Lean_Compiler_LCNF_CompilerM_0__Lean_Compiler_LCNF_normExprImp_go(x_671, x_672, x_666); +x_674 = lean_ptr_addr(x_666); +lean_dec(x_666); +x_675 = lean_ptr_addr(x_673); +x_676 = lean_usize_dec_eq(x_674, x_675); +if (x_676 == 0) { -lean_object* x_701; lean_object* x_702; lean_object* x_703; +lean_object* x_677; lean_object* x_678; lean_object* x_679; if (lean_is_exclusive(x_1)) { lean_ctor_release(x_1, 0); - x_701 = x_1; + x_677 = x_1; } else { lean_dec_ref(x_1); - x_701 = lean_box(0); + x_677 = lean_box(0); } -if (lean_is_scalar(x_701)) { - x_702 = lean_alloc_ctor(6, 1, 0); +if (lean_is_scalar(x_677)) { + x_678 = lean_alloc_ctor(6, 1, 0); } else { - x_702 = x_701; + x_678 = x_677; } -lean_ctor_set(x_702, 0, x_697); -if (lean_is_scalar(x_694)) { - x_703 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_678, 0, x_673); +if (lean_is_scalar(x_670)) { + x_679 = lean_alloc_ctor(0, 2, 0); } else { - x_703 = x_694; + x_679 = x_670; } -lean_ctor_set(x_703, 0, x_702); -lean_ctor_set(x_703, 1, x_693); -return x_703; +lean_ctor_set(x_679, 0, x_678); +lean_ctor_set(x_679, 1, x_669); +return x_679; } else { -lean_object* x_704; -lean_dec(x_697); -if (lean_is_scalar(x_694)) { - x_704 = lean_alloc_ctor(0, 2, 0); +lean_object* x_680; +lean_dec(x_673); +if (lean_is_scalar(x_670)) { + x_680 = lean_alloc_ctor(0, 2, 0); } else { - x_704 = x_694; + x_680 = x_670; } -lean_ctor_set(x_704, 0, x_1); -lean_ctor_set(x_704, 1, x_693); -return x_704; +lean_ctor_set(x_680, 0, x_1); +lean_ctor_set(x_680, 1, x_669); +return x_680; } } } @@ -14746,7 +14713,7 @@ return x_704; } else { -lean_object* x_705; +lean_object* x_681; lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); @@ -14759,8 +14726,8 @@ lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_1); -x_705 = l_Lean_Compiler_LCNF_Simp_withIncRecDepth_throwMaxRecDepth___rarg(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_705; +x_681 = l_Lean_Compiler_LCNF_Simp_withIncRecDepth_throwMaxRecDepth___rarg(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_681; } } } @@ -16558,11 +16525,11 @@ lean_dec(x_1); return x_14; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_simp___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_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_simp___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) { _start: { lean_object* x_15; -x_15 = l_Lean_Compiler_LCNF_Simp_simp___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 = l_Lean_Compiler_LCNF_Simp_simp___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, x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -16576,30 +16543,30 @@ lean_dec(x_1); return x_15; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_simp___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_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_simp___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: { uint8_t x_15; lean_object* x_16; x_15 = lean_unbox(x_4); lean_dec(x_4); -x_16 = l_Lean_Compiler_LCNF_Simp_simp___lambda__3(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); +x_16 = l_Lean_Compiler_LCNF_Simp_simp___lambda__4(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; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_simp___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_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_simp___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) { _start: { lean_object* x_12; -x_12 = l_Lean_Compiler_LCNF_Simp_simp___lambda__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 = l_Lean_Compiler_LCNF_Simp_simp___lambda__5(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_3); return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_simp___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_Lean_Compiler_LCNF_Simp_simp___lambda__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) { _start: { lean_object* x_15; -x_15 = l_Lean_Compiler_LCNF_Simp_simp___lambda__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 = l_Lean_Compiler_LCNF_Simp_simp___lambda__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_12); lean_dec(x_11); @@ -16613,13 +16580,13 @@ lean_dec(x_1); return x_15; } } -LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_Simp_simp___lambda__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_Lean_Compiler_LCNF_Simp_simp___lambda__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) { _start: { uint8_t x_15; lean_object* x_16; x_15 = lean_unbox(x_4); lean_dec(x_4); -x_16 = l_Lean_Compiler_LCNF_Simp_simp___lambda__6(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); +x_16 = l_Lean_Compiler_LCNF_Simp_simp___lambda__7(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; } } diff --git a/stage0/stdlib/Lean/Compiler/Main.c b/stage0/stdlib/Lean/Compiler/Main.c index dd8915de7d9b..a9cb25a93a0b 100644 --- a/stage0/stdlib/Lean/Compiler/Main.c +++ b/stage0/stdlib/Lean/Compiler/Main.c @@ -13,7 +13,6 @@ #ifdef __cplusplus extern "C" { #endif -lean_object* l_Lean_profileitM___at_Lean_addDecl___spec__7___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_64____closed__10; static lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_64____closed__7; static lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_64____closed__11; @@ -35,6 +34,7 @@ static lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_64___ lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_64____closed__12; LEAN_EXPORT lean_object* l_Lean_Compiler_compile___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_profileitM___at_Lean_addDecl___spec__14___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_64____closed__8; static lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_64____closed__13; static lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_64____closed__17; @@ -114,7 +114,7 @@ x_6 = lean_alloc_closure((void*)(l_Lean_Compiler_compile___lambda__1), 4, 1); lean_closure_set(x_6, 0, x_1); x_7 = l_Lean_Compiler_compile___closed__1; x_8 = lean_box(0); -x_9 = l_Lean_profileitM___at_Lean_addDecl___spec__7___rarg(x_7, x_5, x_6, x_8, x_2, x_3, x_4); +x_9 = l_Lean_profileitM___at_Lean_addDecl___spec__14___rarg(x_7, x_5, x_6, x_8, x_2, x_3, x_4); lean_dec(x_5); return x_9; } diff --git a/stage0/stdlib/Lean/Compiler/Options.c b/stage0/stdlib/Lean/Compiler/Options.c index 63f4758adbb9..2a9d297e6257 100644 --- a/stage0/stdlib/Lean/Compiler/Options.c +++ b/stage0/stdlib/Lean/Compiler/Options.c @@ -19,10 +19,10 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_Options___h static lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_Options___hyg_6____closed__8; static lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_Options___hyg_6____closed__5; static lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_Options___hyg_6____closed__6; -lean_object* l_Lean_Option_register___at_Std_Format_initFn____x40_Lean_Data_Format___hyg_95____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_compiler_check; static lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_Options___hyg_6____closed__4; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); +lean_object* l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_Options___hyg_6____closed__1; lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_Options___hyg_6____closed__2; @@ -110,7 +110,7 @@ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_2 = l_Lean_Compiler_initFn____x40_Lean_Compiler_Options___hyg_6____closed__3; x_3 = l_Lean_Compiler_initFn____x40_Lean_Compiler_Options___hyg_6____closed__5; x_4 = l_Lean_Compiler_initFn____x40_Lean_Compiler_Options___hyg_6____closed__8; -x_5 = l_Lean_Option_register___at_Std_Format_initFn____x40_Lean_Data_Format___hyg_95____spec__1(x_2, x_3, x_4, x_1); +x_5 = l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__1(x_2, x_3, x_4, x_1); return x_5; } } diff --git a/stage0/stdlib/Lean/CoreM.c b/stage0/stdlib/Lean/CoreM.c index d58b88af7452..ced1fb5741d9 100644 --- a/stage0/stdlib/Lean/CoreM.c +++ b/stage0/stdlib/Lean/CoreM.c @@ -13,15 +13,15 @@ #ifdef __cplusplus extern "C" { #endif +lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___spec__1(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_getMessageLog___boxed(lean_object*); lean_object* l_Lean_ConstantInfo_instantiateTypeLevelParams(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadCoreM___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadTraceCoreM; -uint8_t l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Core_instantiateTypeLevelParams___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__3; +static lean_object* l_Lean_addDecl___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Core_instMonadLogCoreM___lambda__1(lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Option_register___at_Std_Format_initFn____x40_Lean_Data_Format___hyg_30____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_checkMaxHeartbeats___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_instMonadCoreM___closed__1; static lean_object* l_Lean_Core_instMonadCoreM___closed__5; @@ -29,54 +29,65 @@ LEAN_EXPORT lean_object* l___private_Lean_CoreM_0__Lean_Core_withCurrHeartbeatsI LEAN_EXPORT uint8_t l_Lean_Exception_isMaxHeartbeat(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Core_instantiateTypeLevelParams___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_Context_ref___default; -LEAN_EXPORT lean_object* l_Lean_addDecl___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_addDecl___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Core_withIncRecDepth___spec__1___rarg___closed__2; LEAN_EXPORT lean_object* l_Lean_Core_instInhabitedCoreM(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_initFn____x40_Lean_CoreM___hyg_7____closed__6; LEAN_EXPORT lean_object* l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__4(lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_Context_currMacroScope___default; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Core_instantiateTypeLevelParams___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_profiler; LEAN_EXPORT lean_object* l_Lean_Core_getMaxHeartbeats(lean_object*); LEAN_EXPORT lean_object* l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadTraceCoreM___lambda__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_Context_options___default; static lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Core_instantiateTypeLevelParams___spec__2___closed__3; lean_object* lean_mk_empty_array_with_capacity(lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_CoreM___hyg_5108____closed__3; size_t lean_usize_shift_right(size_t, size_t); +static lean_object* l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__13; LEAN_EXPORT lean_object* l_List_elem___at_Lean_catchInternalIds___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Core_instMetaEvalCoreM___spec__5(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkArrow___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadTraceCoreM___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__4; static lean_object* l_Lean_Core_initFn____x40_Lean_CoreM___hyg_7____closed__5; +lean_object* l_Lean_PersistentArray_toArray___rarg(lean_object*); lean_object* l_Lean_profileitIOUnsafe___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMetaEvalCoreM(lean_object*); -LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_addDecl___spec__7___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_catchInternalIds___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadResolveNameCoreM___lambda__1(lean_object*, lean_object*, lean_object*); static lean_object* l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Core_instMonadQuotationCoreM___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_compileDecls___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_maxRecDepthErrorMessage; LEAN_EXPORT lean_object* l_Lean_Core_instantiateTypeLevelParams(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_le(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Declaration_foldExprM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_toString(lean_object*, lean_object*); +double lean_float_div(double, double); LEAN_EXPORT lean_object* l_Lean_log___at_Lean_addDecl___spec__5(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_withCurrHeartbeats(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Core_withIncRecDepth___spec__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_logAt___at_Lean_addDecl___spec__6___closed__1; +static lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__4; lean_object* l_Lean_FileMap_toPosition(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_CoreM_0__Lean_checkUnsupported___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__12; size_t lean_uint64_to_usize(uint64_t); LEAN_EXPORT lean_object* l_Lean_Core_CoreM_run(lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instInhabitedCoreM___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_instMonadNameGeneratorCoreM___closed__3; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadTraceCoreM___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_Cache_instLevelType___default; lean_object* l_Lean_Name_toString(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Core_instMonadRefCoreM; +static lean_object* l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__5; +double l_Lean_trace_profiler_threshold_getSecs(lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_withIncRecDepth(lean_object*, lean_object*); static lean_object* l_Lean_Core_instMonadInfoTreeCoreM___closed__1; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__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_Core_instMonadResolveNameCoreM; lean_object* l_Lean_PersistentArray_push___rarg(lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); @@ -87,7 +98,6 @@ size_t lean_usize_mul(size_t, size_t); static lean_object* l_Lean_Core_instMonadTraceCoreM___closed__3; LEAN_EXPORT lean_object* l_Lean_Core_instMonadLiftIOCoreM(lean_object*); lean_object* l_ReaderT_instFunctorReaderT___rarg(lean_object*); -static lean_object* l_Lean_addDecl___lambda__2___closed__3; LEAN_EXPORT lean_object* l_Lean_Core_instantiateTypeLevelParams___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_throwMaxHeartbeat___closed__4; LEAN_EXPORT lean_object* l_Lean_Core_instMonadInfoTreeCoreM___lambda__1___boxed(lean_object*, lean_object*, lean_object*); @@ -97,24 +107,28 @@ LEAN_EXPORT lean_object* l_Lean_Core_throwMaxHeartbeat___boxed(lean_object*, lea static lean_object* l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__26; LEAN_EXPORT lean_object* l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__4___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_compileDecl(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*); lean_object* l_Lean_replaceRef(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_addDecl___spec__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getPos_x3f(lean_object*, uint8_t); lean_object* lean_array_fget(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_liftIOCore___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getTailPos_x3f(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Core_liftIOCore___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_instMetaEvalCoreM___rarg___lambda__1___closed__5; LEAN_EXPORT lean_object* l_Lean_Core_instMonadEnvCoreM; lean_object* l_Lean_ConstantInfo_name(lean_object*); static lean_object* l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__16; LEAN_EXPORT lean_object* l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_ReaderT_instMonadReaderT___rarg(lean_object*); +static lean_object* l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__7; static lean_object* l_Lean_Core_instMonadEnvCoreM___closed__2; +uint8_t lean_float_decLt(double, double); LEAN_EXPORT lean_object* l_Lean_Core_instInhabitedCoreM___rarg(lean_object*); static lean_object* l_Lean_Core_instMonadCoreM___closed__4; LEAN_EXPORT lean_object* l_Lean_logWarning___at_Lean_addDecl___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_compileDecls___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_hasSyntheticSorry(lean_object*); lean_object* l_EStateM_instMonadEStateM(lean_object*, lean_object*); @@ -126,11 +140,14 @@ static lean_object* l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessag LEAN_EXPORT lean_object* l___private_Lean_CoreM_0__Lean_Core_withCurrHeartbeatsImp___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_getDeclNamesForCodeGen(lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadRecDepthCoreM___lambda__3(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_addDecl___closed__2; lean_object* lean_io_get_num_heartbeats(lean_object*); static lean_object* l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__1; lean_object* lean_lcnf_compile_decls(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadLogCoreM___lambda__3(lean_object*, lean_object*, lean_object*); +uint8_t l___private_Lean_Util_Trace_0__Lean_checkTraceOption(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_addDecl___lambda__3___closed__2; LEAN_EXPORT lean_object* l_Lean_Core_mkFreshUserName___boxed(lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_maxRecDepth; static size_t l_Lean_PersistentHashMap_insertAux___at_Lean_Core_instantiateTypeLevelParams___spec__2___closed__1; @@ -140,6 +157,7 @@ static lean_object* l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__ LEAN_EXPORT lean_object* l_Lean_Core_instMonadRecDepthCoreM___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Core_instantiateTypeLevelParams___spec__3(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Core_instantiateTypeLevelParams___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_addDecl___spec__9___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Core_instantiateTypeLevelParams___spec__6___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_compileDecl___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__14; @@ -149,34 +167,42 @@ LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_Core_instMetaEva LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_compileDecl___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Core_instantiateTypeLevelParams___spec__2(lean_object*, size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__7___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_Core_instMonadRefCoreM___closed__1; static lean_object* l_Lean_Core_initFn____x40_Lean_CoreM___hyg_7____closed__8; lean_object* l_Lean_Environment_compileDecl(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadRecDepthCoreM___lambda__3___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__15; LEAN_EXPORT lean_object* l_Lean_Core_CoreM_run_x27___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__11; LEAN_EXPORT lean_object* l_Lean_Core_instMonadOptionsCoreM(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Declaration_foldExprM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_State_messages___default; static lean_object* l_Lean_Core_instMetaEvalCoreM___rarg___lambda__1___closed__4; lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_instInhabitedCache___closed__1; +LEAN_EXPORT lean_object* l_Lean_withSeconds___at_Lean_addDecl___spec__10(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_instMonadEnvCoreM___closed__1; lean_object* l_ReaderT_instApplicativeReaderT___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_checkMaxHeartbeatsCore___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_State_traceState___default; size_t lean_usize_of_nat(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_addDecl___spec__12(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_compileDecl___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_List_elem___at_Lean_catchInternalIds___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_modifyInstLevelTypeCache___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMetaEvalCoreM___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Lean_addDecl___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_modifyInstLevelValueCache___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadQuotationCoreM___lambda__2___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_addDecl___spec__14___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_beq___at_Lean_Core_instantiateTypeLevelParams___spec__8___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_logAt___at_Lean_addDecl___spec__6___boxed(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_Core_instMonadLogCoreM___lambda__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__6; static lean_object* l_Lean_Core_instMonadQuotationCoreM___closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_addDecl___spec__9___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_List_beq___at_Lean_Core_instantiateTypeLevelParams___spec__8(lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_catchInternalId___boxed(lean_object*, lean_object*, lean_object*); @@ -195,8 +221,10 @@ LEAN_EXPORT lean_object* l_Lean_Core_instMonadRecDepthCoreM; LEAN_EXPORT lean_object* l_Lean_setEnv___at_Lean_addDecl___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_div(lean_object*, lean_object*); static lean_object* l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg___lambda__2___closed__2; +lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__29; static lean_object* l_Lean_addDecl___closed__1; +lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_Context_openDecls___default; LEAN_EXPORT lean_object* l_Lean_Core_instMonadRefCoreM___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ImportM_runCoreM___rarg(lean_object*, lean_object*, lean_object*); @@ -207,10 +235,14 @@ lean_object* lean_compile_decls(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_Context_currRecDepth___default; static lean_object* l_Lean_Core_instInhabitedCoreM___rarg___closed__2; LEAN_EXPORT lean_object* l_Lean_Core_withIncRecDepth___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_addDecl___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Core_instantiateTypeLevelParams___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_PersistentArray_append___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Core_State_ngen___default___closed__1; LEAN_EXPORT lean_object* l___private_Lean_CoreM_0__Lean_checkUnsupported(lean_object*); static lean_object* l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__20; +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_CoreM___hyg_5108_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_addDecl___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_forallE___override(lean_object*, lean_object*, lean_object*, uint8_t); static lean_object* l_Lean_Core_instMonadNameGeneratorCoreM___closed__2; static lean_object* l_Lean_Core_instMonadRefCoreM___closed__3; @@ -223,18 +255,24 @@ lean_object* lean_st_mk_ref(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_catchInternalId(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__21; LEAN_EXPORT lean_object* l_Lean_Core_instantiateValueLevelParams___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, double, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_State_cache___default; +static lean_object* l_Lean_initFn____x40_Lean_CoreM___hyg_5108____closed__5; LEAN_EXPORT lean_object* l_Lean_Core_Cache_instLevelValue___default; +lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); static lean_object* l_Lean_Core_instMonadTraceCoreM___closed__1; LEAN_EXPORT lean_object* l_Lean_catchInternalIds(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_printTraces___at_Lean_Core_instMetaEvalCoreM___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instantiateTypeLevelParams___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_instMonadNameGeneratorCoreM___closed__1; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__7___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_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Core_instantiateTypeLevelParams___spec__6(lean_object*, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_addDecl___spec__7(lean_object*); +lean_object* lean_io_mono_nanos_now(lean_object*); static lean_object* l_Lean_Core_Cache_instLevelType___default___closed__3; static lean_object* l_Lean_Core_Cache_instLevelType___default___closed__1; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__7(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_addDecl___spec__11(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_Cache_instLevelType___default___closed__2; static lean_object* l_Lean_ImportM_runCoreM___rarg___closed__1; static lean_object* l_Lean_Core_instMonadLogCoreM___closed__3; @@ -245,6 +283,7 @@ static lean_object* l_Lean_Core_instMonadInfoTreeCoreM___closed__2; LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_Core_instMetaEvalCoreM___spec__3___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_maxHeartbeats; lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_addDecl___spec__8(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_initFn____x40_Lean_CoreM___hyg_7____closed__3; static lean_object* l_Lean_Core_instMonadRecDepthCoreM___closed__2; static lean_object* l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__6; @@ -256,10 +295,13 @@ static lean_object* l_Lean_Core_State_ngen___default___closed__2; static lean_object* l_Lean_Core_throwMaxHeartbeat___closed__3; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); extern lean_object* l_Lean_warningAsError; +static lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__1; LEAN_EXPORT lean_object* l_Lean_Core_instMonadLogCoreM; LEAN_EXPORT lean_object* l_Lean_Core_instMonadInfoTreeCoreM___lambda__1(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__9; LEAN_EXPORT lean_object* l_Lean_Core_Context_maxHeartbeats___default___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadInfoTreeCoreM___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__17; lean_object* l_Lean_ConstantInfo_instantiateValueLevelParams_x21(lean_object*, lean_object*); static lean_object* l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__22; @@ -268,7 +310,6 @@ static lean_object* l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Core_instMetaEvalCoreM___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_throwMaxHeartbeat___closed__2; static lean_object* l_Lean_Core_State_nextMacroScope___default___closed__1; -static lean_object* l_Lean_addDecl___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_catchInternalId___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_mkArrow___closed__1; LEAN_EXPORT lean_object* l_Lean_Core_setMessageLog(lean_object*, lean_object*, lean_object*, lean_object*); @@ -276,8 +317,10 @@ LEAN_EXPORT lean_object* l_Lean_catchInternalId___rarg(lean_object*, lean_object static lean_object* l_Lean_Exception_isMaxHeartbeat___closed__1; LEAN_EXPORT lean_object* l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_Core_instMetaEvalCoreM___spec__3___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_addDecl___spec__13(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instantiateValueLevelParams___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_compileDecls(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_addDecl___lambda__1___closed__1; static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Core_withIncRecDepth___spec__1___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_Core_CoreM_toIO(lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_Context_maxHeartbeats___default(lean_object*); @@ -291,6 +334,7 @@ LEAN_EXPORT lean_object* l_Lean_Core_instantiateValueLevelParams___lambda__1___b lean_object* lean_usize_to_nat(size_t); LEAN_EXPORT lean_object* l_Lean_Core_mkFreshUserName(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_withCurrHeartbeats___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_inheritedTraceOptions; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Core_instantiateTypeLevelParams___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadWithOptionsCoreM___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Core_instMetaEvalCoreM___spec__4(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); @@ -301,15 +345,18 @@ LEAN_EXPORT lean_object* l_Lean_Core_instantiateTypeLevelParams___lambda__1___bo LEAN_EXPORT lean_object* l_Lean_Core_instMonadLogCoreM___lambda__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Declaration_foldExprM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__1___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_instMonadRecDepthCoreM___closed__1; -static lean_object* l_Lean_addDecl___lambda__2___closed__2; static lean_object* l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__28; LEAN_EXPORT lean_object* l_Lean_compileDeclsNew___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_PersistentArray_anyM___at_Lean_MessageLog_hasErrors___spec__1(lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_CoreM___hyg_5108____closed__1; +double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); static lean_object* l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__19; LEAN_EXPORT lean_object* l_Lean_catchInternalIds___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_getMessageLog(lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadEnvCoreM___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_mkArrow___closed__2; +static lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__5; LEAN_EXPORT lean_object* l_Lean_Core_instMonadQuotationCoreM___lambda__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_State_infoState___default___closed__2; LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_compileDecl___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -317,11 +364,18 @@ lean_object* l_Lean_Expr_FindImpl_findUnsafe_x3f(lean_object*, lean_object*); static lean_object* l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg___lambda__2___closed__4; LEAN_EXPORT lean_object* l_Lean_Core_instMonadLogCoreM___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_instantiateValueLevelParams___lambda__2___closed__3; +static lean_object* l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__4; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__2(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_catchInternalIds___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__2; +static double l_Lean_withSeconds___at_Lean_addDecl___spec__10___closed__1; static lean_object* l_Lean_Core_instMetaEvalCoreM___rarg___lambda__1___closed__1; lean_object* l_Lean_throwError___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_CoreM_run___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_addDecl___spec__9___rarg(lean_object*, lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_CoreM___hyg_5108____closed__4; static lean_object* l_Lean_Core_instMonadLogCoreM___closed__4; +static lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__3; LEAN_EXPORT lean_object* l_Lean_Core_withIncRecDepth___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withAtLeastMaxRecDepth___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_instMonadRefCoreM___closed__2; @@ -331,17 +385,21 @@ LEAN_EXPORT lean_object* l_Lean_Core_modifyInstLevelValueCache(lean_object*, lea static lean_object* l_Lean_Core_initFn____x40_Lean_CoreM___hyg_7____closed__2; static lean_object* l_Lean_Core_initFn____x40_Lean_CoreM___hyg_7____closed__4; LEAN_EXPORT lean_object* l_Lean_Core_modifyCache___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_CoreM___hyg_5108____closed__2; LEAN_EXPORT lean_object* l_Lean_Core_instMonadNameGeneratorCoreM; static lean_object* l_Lean_Core_State_ngen___default___closed__3; +static lean_object* l_Lean_compileDecl___closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Core_instMetaEvalCoreM___spec__6(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); lean_object* lean_environment_main_module(lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadQuotationCoreM___lambda__2(lean_object*, lean_object*, lean_object*); +lean_object* lean_float_to_string(double); LEAN_EXPORT lean_object* l_Lean_Core_instMonadNameGeneratorCoreM___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_State_traceState___default___closed__3; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadLiftIOCoreM___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Declaration_foldExprM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__1___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_addDecl___lambda__3___closed__3; LEAN_EXPORT lean_object* l_Lean_catchInternalIds___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadQuotationCoreM___lambda__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg___lambda__2___closed__3; @@ -350,6 +408,7 @@ LEAN_EXPORT lean_object* l_Lean_Core_checkMaxHeartbeats(lean_object*, lean_objec static lean_object* l_Lean_Core_instMonadLogCoreM___closed__5; LEAN_EXPORT lean_object* l_Lean_Core_checkMaxHeartbeatsCore(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_String_isPrefixOf(lean_object*, lean_object*); +static lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__2; LEAN_EXPORT lean_object* l_Lean_Core_instMonadRefCoreM___lambda__1___boxed(lean_object*, lean_object*, lean_object*); uint64_t l_Lean_Name_hash___override(lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadNameGeneratorCoreM___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); @@ -357,14 +416,18 @@ LEAN_EXPORT lean_object* l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUn uint8_t l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadCoreM___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__11; +static lean_object* l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__14; static lean_object* l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__18; LEAN_EXPORT lean_object* l_Lean_Core_State_nextMacroScope___default; LEAN_EXPORT lean_object* l_Lean_Core_instMonadInfoTreeCoreM___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_compileDecl___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_modifyCache(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); +lean_object* l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_modifyInstLevelTypeCache(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__10; LEAN_EXPORT uint8_t l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg___lambda__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_addDecl___spec__14___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_mul(lean_object*, lean_object*); static lean_object* l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__5; static size_t l_Lean_PersistentHashMap_insertAux___at_Lean_Core_instantiateTypeLevelParams___spec__2___closed__2; @@ -377,15 +440,17 @@ LEAN_EXPORT lean_object* l_Lean_Core_restore___boxed(lean_object*, lean_object*, static lean_object* l_Lean_Core_instMonadCoreM___closed__3; static lean_object* l_Lean_Core_throwMaxHeartbeat___closed__1; LEAN_EXPORT lean_object* l_Lean_mkArrow(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_addDecl___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkCollisionNode___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Core_instMetaEvalCoreM___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_withTraceNode___at_Lean_addDecl___spec__7___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadCoreM; -LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_addDecl___spec__7___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_instMonadCoreM___closed__6; uint8_t l_Lean_Declaration_foldExprM___at_Lean_Declaration_hasSorry___spec__1(lean_object*, uint8_t); static lean_object* l_Lean_Core_State_traceState___default___closed__2; uint8_t l_Lean_isRecCore(lean_object*, lean_object*); +static lean_object* l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__8; LEAN_EXPORT lean_object* l_Lean_Core_throwMaxHeartbeat(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__3___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_firstFrontendMacroScope; @@ -393,8 +458,10 @@ LEAN_EXPORT lean_object* l_Lean_logAt___at_Lean_addDecl___spec__6(lean_object*, size_t lean_usize_sub(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Core_instMonadInfoTreeCoreM; LEAN_EXPORT lean_object* l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_addDecl___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_addDecl___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadEnvCoreM___lambda__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__1; +static lean_object* l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__3; static lean_object* l_Lean_Core_instMonadLogCoreM___closed__2; LEAN_EXPORT lean_object* l_Lean_Core_CoreM_toIO___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_withFreshMacroScope(lean_object*); @@ -405,12 +472,14 @@ LEAN_EXPORT lean_object* l___private_Lean_CoreM_0__Lean_checkUnsupported___rarg_ LEAN_EXPORT lean_object* l_Lean_setEnv___at_Lean_addDecl___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_Context_initHeartbeats___default; LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_compileDecl___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_compileDecl___lambda__1(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_Lean_Core_instMetaEvalCoreM___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instantiateValueLevelParams___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadTraceCoreM___lambda__2(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_instMonadResolveNameCoreM___closed__2; LEAN_EXPORT lean_object* l_Lean_Core_instMonadNameGeneratorCoreM___lambda__1___boxed(lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_trace_profiler; LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_Core_instMetaEvalCoreM___spec__2___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_io_error_to_string(lean_object*); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); @@ -418,8 +487,10 @@ static lean_object* l_Lean_Core_instMonadLogCoreM___closed__1; LEAN_EXPORT lean_object* l_Lean_Core_instMetaEvalCoreM___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_liftIOCore(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_Core_instMetaEvalCoreM___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_addDecl___spec__9(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Core_instantiateValueLevelParams___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_shift_left(size_t, size_t); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_addDecl___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadResolveNameCoreM___lambda__2(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__9; static lean_object* l_Lean_Core_CoreM_toIO___rarg___closed__1; @@ -428,8 +499,10 @@ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Core_instan LEAN_EXPORT lean_object* l_Lean_Core_instMonadLogCoreM___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_Core_instMetaEvalCoreM___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_CoreM_0__Lean_supportedRecursors; +static lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__1; static lean_object* l_Lean_Core_State_traceState___default___closed__1; lean_object* lean_string_append(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_addDecl___spec__14(lean_object*); LEAN_EXPORT lean_object* l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__3(lean_object*); static lean_object* l_Lean_Core_instMonadEnvCoreM___closed__3; lean_object* lean_array_get_size(lean_object*); @@ -437,6 +510,7 @@ LEAN_EXPORT lean_object* l_Lean_Core_restore(lean_object*, lean_object*, lean_ob LEAN_EXPORT lean_object* l_Lean_Core_instAddMessageContextCoreM; LEAN_EXPORT lean_object* l_Lean_Core_instMonadEnvCoreM___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__8; +static lean_object* l_Lean_addDecl___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Lean_Core_withCurrHeartbeats___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ImportM_runCoreM(lean_object*); uint8_t lean_level_eq(lean_object*, lean_object*); @@ -452,6 +526,7 @@ lean_object* l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(lean_object*); static lean_object* l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__23; LEAN_EXPORT lean_object* l_Lean_Core_instMonadCoreM___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_instMonadQuotationCoreM___closed__4; +lean_object* l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_68____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_initFn____x40_Lean_CoreM___hyg_7_(lean_object*); static lean_object* l_Lean_Core_instInhabitedCoreM___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_withAtLeastMaxRecDepth(lean_object*, lean_object*); @@ -460,6 +535,7 @@ uint8_t l___private_Lean_Message_0__Lean_beqMessageSeverity____x40_Lean_Message_ LEAN_EXPORT lean_object* l_Lean_Core_instMonadRefCoreM___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_KernelException_toMessageData(lean_object*, lean_object*); static lean_object* l_Lean_Core_instMonadResolveNameCoreM___closed__3; +static lean_object* l_Lean_initFn____x40_Lean_CoreM___hyg_5108____closed__6; LEAN_EXPORT lean_object* l_Lean_Declaration_foldExprM___at_Lean_compileDecl___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_addDecl___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1___closed__4; @@ -471,6 +547,7 @@ static lean_object* l_Lean_Core_instMonadRecDepthCoreM___closed__4; LEAN_EXPORT lean_object* l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2(lean_object*); static lean_object* l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__27; LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Core_withIncRecDepth___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_isTracingEnabledFor___at_Lean_addDecl___spec__8___closed__1; LEAN_EXPORT lean_object* l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Nat_repr(lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadResolveNameCoreM___lambda__2___boxed(lean_object*, lean_object*, lean_object*); @@ -479,10 +556,12 @@ lean_object* l_Lean_MessageData_format(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addDecl(lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_land(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_log___at_Lean_addDecl___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_compiler_enableNew; LEAN_EXPORT lean_object* l_Lean_Core_getMaxHeartbeats___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_Context_currNamespace___default; static lean_object* l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1___closed__1; static lean_object* l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__24; +LEAN_EXPORT lean_object* l_Lean_compileDecl___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_initFn____x40_Lean_CoreM___hyg_7____closed__1; LEAN_EXPORT lean_object* l_Lean_Core_instMonadOptionsCoreM___boxed(lean_object*, lean_object*, lean_object*); static lean_object* _init_l_Lean_Core_initFn____x40_Lean_CoreM___hyg_7____closed__1() { @@ -567,7 +646,148 @@ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_2 = l_Lean_Core_initFn____x40_Lean_CoreM___hyg_7____closed__2; x_3 = l_Lean_Core_initFn____x40_Lean_CoreM___hyg_7____closed__5; x_4 = l_Lean_Core_initFn____x40_Lean_CoreM___hyg_7____closed__8; -x_5 = l_Lean_Option_register___at_Std_Format_initFn____x40_Lean_Data_Format___hyg_30____spec__1(x_2, x_3, x_4, x_1); +x_5 = l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_68____spec__1(x_2, x_3, x_4, x_1); +return x_5; +} +} +static lean_object* _init_l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Kernel", 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Core_initFn____x40_Lean_CoreM___hyg_7____closed__6; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__3; +x_2 = l_Lean_Core_initFn____x40_Lean_CoreM___hyg_7____closed__7; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("initFn", 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__4; +x_2 = l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__5; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("_@", 2); +return x_1; +} +} +static lean_object* _init_l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__6; +x_2 = l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__7; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__8; +x_2 = l_Lean_Core_initFn____x40_Lean_CoreM___hyg_7____closed__6; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__10() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("CoreM", 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__9; +x_2 = l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__10; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__12() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("_hyg", 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__11; +x_2 = l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__12; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__13; +x_2 = lean_unsigned_to_nat(44u); +x_3 = l_Lean_Name_num___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44_(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; +x_2 = l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__2; +x_3 = 0; +x_4 = l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__14; +x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } } @@ -584,7 +804,7 @@ LEAN_EXPORT lean_object* l_Lean_Core_getMaxHeartbeats(lean_object* x_1) { { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_2 = l_Lean_Core_getMaxHeartbeats___closed__1; -x_3 = l_Lean_Option_get___at_Std_Format_pretty_x27___spec__1(x_1, x_2); +x_3 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_2); x_4 = lean_unsigned_to_nat(1000u); x_5 = lean_nat_mul(x_3, x_4); lean_dec(x_3); @@ -6703,7 +6923,7 @@ LEAN_EXPORT lean_object* l_Lean_Core_instMetaEvalCoreM___rarg___lambda__1(lean_o { lean_object* x_7; lean_object* x_8; lean_object* x_29; lean_object* x_30; lean_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_29 = l_Lean_Core_instMetaEvalCoreM___rarg___lambda__1___closed__3; -x_30 = l_Lean_Option_get___at_Std_Format_pretty_x27___spec__1(x_3, x_29); +x_30 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_3, x_29); x_31 = lean_box(0); x_32 = l_Lean_Core_getMaxHeartbeats(x_3); x_33 = l_Lean_Core_State_nextMacroScope___default___closed__1; @@ -8694,7 +8914,7 @@ else { lean_object* x_261; uint8_t x_262; x_261 = l_Lean_logAt___at_Lean_addDecl___spec__6___closed__1; -x_262 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_258, x_261); +x_262 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_258, x_261); if (x_262 == 0) { x_8 = x_3; @@ -9543,7 +9763,7 @@ lean_object* x_5; lean_object* x_6; uint8_t x_7; x_5 = lean_ctor_get(x_2, 2); lean_inc(x_5); x_6 = l_Lean_logAt___at_Lean_addDecl___spec__6___closed__1; -x_7 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_5, x_6); +x_7 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_5, x_6); lean_dec(x_5); if (x_7 == 0) { @@ -9561,122 +9781,1231 @@ return x_11; } } } -LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_addDecl___spec__7___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) { +static lean_object* _init_l_Lean_isTracingEnabledFor___at_Lean_addDecl___spec__8___closed__1() { _start: { -lean_object* x_8; lean_object* x_9; -x_8 = lean_apply_2(x_3, x_5, x_6); -x_9 = l_Lean_profileitIOUnsafe___rarg(x_1, x_2, x_8, x_4, x_7); -return x_9; +lean_object* x_1; +x_1 = l_Lean_inheritedTraceOptions; +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_addDecl___spec__7(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_addDecl___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_profileitM___at_Lean_addDecl___spec__7___rarg___boxed), 7, 0); -return x_2; +lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_5 = l_Lean_isTracingEnabledFor___at_Lean_addDecl___spec__8___closed__1; +x_6 = lean_st_ref_get(x_5, x_4); +x_7 = !lean_is_exclusive(x_6); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; uint8_t x_10; lean_object* x_11; +x_8 = lean_ctor_get(x_6, 0); +x_9 = lean_ctor_get(x_2, 2); +x_10 = l___private_Lean_Util_Trace_0__Lean_checkTraceOption(x_8, x_9, x_1); +lean_dec(x_8); +x_11 = lean_box(x_10); +lean_ctor_set(x_6, 0, x_11); +return x_6; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; +x_12 = lean_ctor_get(x_6, 0); +x_13 = lean_ctor_get(x_6, 1); +lean_inc(x_13); +lean_inc(x_12); +lean_dec(x_6); +x_14 = lean_ctor_get(x_2, 2); +x_15 = l___private_Lean_Util_Trace_0__Lean_checkTraceOption(x_12, x_14, x_1); +lean_dec(x_12); +x_16 = lean_box(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_13); +return x_17; +} } } -LEAN_EXPORT lean_object* l_Lean_addDecl___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___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_addDecl___spec__9___rarg(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -lean_dec(x_2); -x_6 = lean_st_ref_get(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_object* x_3; 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_st_ref_get(x_1, x_2); +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 = lean_ctor_get(x_4, 3); +lean_inc(x_6); +lean_dec(x_4); +x_7 = lean_st_ref_take(x_1, x_5); +x_8 = lean_ctor_get(x_7, 0); lean_inc(x_8); -lean_dec(x_6); -x_9 = lean_ctor_get(x_7, 0); +x_9 = lean_ctor_get(x_7, 1); lean_inc(x_9); lean_dec(x_7); -x_10 = lean_add_decl(x_9, x_1); -lean_dec(x_1); -if (lean_obj_tag(x_10) == 0) +x_10 = !lean_is_exclusive(x_8); +if (x_10 == 0) { -lean_object* x_11; lean_object* x_12; -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -lean_dec(x_10); -x_12 = l_Lean_throwKernelException___at_Lean_addDecl___spec__1(x_11, x_3, x_4, x_8); -lean_dec(x_4); -return x_12; +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_11 = lean_ctor_get(x_8, 3); +lean_dec(x_11); +x_12 = l_Lean_Core_State_traceState___default___closed__3; +lean_ctor_set(x_8, 3, x_12); +x_13 = lean_st_ref_set(x_1, x_8, x_9); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +lean_object* x_15; +x_15 = lean_ctor_get(x_13, 0); +lean_dec(x_15); +lean_ctor_set(x_13, 0, x_6); +return x_13; } else { -lean_object* x_13; lean_object* x_14; -x_13 = lean_ctor_get(x_10, 0); -lean_inc(x_13); -lean_dec(x_10); -x_14 = l_Lean_setEnv___at_Lean_addDecl___spec__3(x_13, x_3, x_4, x_8); -lean_dec(x_4); -lean_dec(x_3); -return x_14; -} +lean_object* x_16; lean_object* x_17; +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); +lean_dec(x_13); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_6); +lean_ctor_set(x_17, 1, x_16); +return x_17; } } -static lean_object* _init_l_Lean_addDecl___lambda__2___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("declaration uses 'sorry'", 24); -return x_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; +x_18 = lean_ctor_get(x_8, 0); +x_19 = lean_ctor_get(x_8, 1); +x_20 = lean_ctor_get(x_8, 2); +x_21 = lean_ctor_get(x_8, 4); +x_22 = lean_ctor_get(x_8, 5); +x_23 = lean_ctor_get(x_8, 6); +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_dec(x_8); +x_24 = l_Lean_Core_State_traceState___default___closed__3; +x_25 = lean_alloc_ctor(0, 7, 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_24); +lean_ctor_set(x_25, 4, x_21); +lean_ctor_set(x_25, 5, x_22); +lean_ctor_set(x_25, 6, x_23); +x_26 = lean_st_ref_set(x_1, x_25, x_9); +x_27 = lean_ctor_get(x_26, 1); +lean_inc(x_27); +if (lean_is_exclusive(x_26)) { + lean_ctor_release(x_26, 0); + lean_ctor_release(x_26, 1); + x_28 = x_26; +} else { + lean_dec_ref(x_26); + x_28 = lean_box(0); +} +if (lean_is_scalar(x_28)) { + x_29 = lean_alloc_ctor(0, 2, 0); +} else { + x_29 = x_28; +} +lean_ctor_set(x_29, 0, x_6); +lean_ctor_set(x_29, 1, x_27); +return x_29; } } -static lean_object* _init_l_Lean_addDecl___lambda__2___closed__2() { +} +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_addDecl___spec__9(lean_object* x_1) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_addDecl___lambda__2___closed__1; -x_2 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_2, 0, x_1); +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_addDecl___spec__9___rarg___boxed), 2, 0); return x_2; } } -static lean_object* _init_l_Lean_addDecl___lambda__2___closed__3() { +static double _init_l_Lean_withSeconds___at_Lean_addDecl___spec__10___closed__1() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_addDecl___lambda__2___closed__2; -x_2 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; +lean_object* x_1; uint8_t x_2; lean_object* x_3; double x_4; +x_1 = lean_unsigned_to_nat(1000000000u); +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_Lean_addDecl___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_withSeconds___at_Lean_addDecl___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; uint8_t x_9; -x_5 = lean_st_ref_get(x_3, x_4); +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_5 = lean_io_mono_nanos_now(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, 5); -lean_inc(x_8); -lean_dec(x_6); -x_9 = l_Lean_PersistentArray_anyM___at_Lean_MessageLog_hasErrors___spec__1(x_8); -if (x_9 == 0) -{ -uint8_t x_10; uint8_t x_11; -x_10 = 0; -lean_inc(x_1); -x_11 = l_Lean_Declaration_foldExprM___at_Lean_Declaration_hasSorry___spec__1(x_1, x_10); -if (x_11 == 0) +x_8 = lean_apply_3(x_1, x_2, x_3, x_7); +if (lean_obj_tag(x_8) == 0) { -lean_object* x_12; lean_object* x_13; -x_12 = lean_box(0); -x_13 = l_Lean_addDecl___lambda__1(x_1, x_12, x_2, x_3, x_7); -return x_13; -} -else +lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +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_dec(x_8); +x_11 = lean_io_mono_nanos_now(x_10); +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_14 = l_Lean_addDecl___lambda__2___closed__3; -lean_inc(x_3); +lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; double x_17; double x_18; double x_19; lean_object* x_20; lean_object* x_21; +x_13 = lean_ctor_get(x_11, 0); +x_14 = lean_nat_sub(x_13, x_6); +lean_dec(x_6); +lean_dec(x_13); +x_15 = 0; +x_16 = lean_unsigned_to_nat(0u); +x_17 = l_Float_ofScientific(x_14, x_15, x_16); +lean_dec(x_14); +x_18 = l_Lean_withSeconds___at_Lean_addDecl___spec__10___closed__1; +x_19 = lean_float_div(x_17, x_18); +x_20 = lean_box_float(x_19); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_9); +lean_ctor_set(x_21, 1, x_20); +lean_ctor_set(x_11, 0, x_21); +return x_11; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; lean_object* x_26; double x_27; double x_28; double x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_22 = lean_ctor_get(x_11, 0); +x_23 = lean_ctor_get(x_11, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_11); +x_24 = lean_nat_sub(x_22, x_6); +lean_dec(x_6); +lean_dec(x_22); +x_25 = 0; +x_26 = lean_unsigned_to_nat(0u); +x_27 = l_Float_ofScientific(x_24, x_25, x_26); +lean_dec(x_24); +x_28 = l_Lean_withSeconds___at_Lean_addDecl___spec__10___closed__1; +x_29 = lean_float_div(x_27, x_28); +x_30 = lean_box_float(x_29); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_9); +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_23); +return x_32; +} +} +else +{ +uint8_t x_33; +lean_dec(x_6); +x_33 = !lean_is_exclusive(x_8); +if (x_33 == 0) +{ +return x_8; +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_8, 0); +x_35 = lean_ctor_get(x_8, 1); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_8); +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_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_addDecl___spec__12(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) { +_start: +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_9 = lean_st_ref_take(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_is_exclusive(x_10); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; size_t x_16; size_t 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; +x_13 = lean_ctor_get(x_10, 3); +x_14 = l_Lean_PersistentArray_toArray___rarg(x_13); +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___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___spec__1(x_16, x_17, x_14); +x_19 = lean_alloc_ctor(9, 3, 1); +lean_ctor_set(x_19, 0, x_2); +lean_ctor_set(x_19, 1, x_4); +lean_ctor_set(x_19, 2, x_18); +lean_ctor_set_uint8(x_19, sizeof(void*)*3, x_5); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_3); +lean_ctor_set(x_20, 1, x_19); +x_21 = l_Lean_PersistentArray_push___rarg(x_1, x_20); +lean_ctor_set(x_10, 3, x_21); +x_22 = lean_st_ref_set(x_7, x_10, x_11); +x_23 = !lean_is_exclusive(x_22); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; +x_24 = lean_ctor_get(x_22, 0); +lean_dec(x_24); +x_25 = lean_box(0); +lean_ctor_set(x_22, 0, x_25); +return x_22; +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_22, 1); +lean_inc(x_26); +lean_dec(x_22); +x_27 = lean_box(0); +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; +} +} +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; lean_object* x_37; size_t x_38; size_t x_39; lean_object* x_40; lean_object* x_41; 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_29 = lean_ctor_get(x_10, 0); +x_30 = lean_ctor_get(x_10, 1); +x_31 = lean_ctor_get(x_10, 2); +x_32 = lean_ctor_get(x_10, 3); +x_33 = lean_ctor_get(x_10, 4); +x_34 = lean_ctor_get(x_10, 5); +x_35 = lean_ctor_get(x_10, 6); +lean_inc(x_35); +lean_inc(x_34); +lean_inc(x_33); +lean_inc(x_32); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_29); +lean_dec(x_10); +x_36 = l_Lean_PersistentArray_toArray___rarg(x_32); +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___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___spec__1(x_38, x_39, x_36); +x_41 = lean_alloc_ctor(9, 3, 1); +lean_ctor_set(x_41, 0, x_2); +lean_ctor_set(x_41, 1, x_4); +lean_ctor_set(x_41, 2, x_40); +lean_ctor_set_uint8(x_41, sizeof(void*)*3, x_5); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_3); +lean_ctor_set(x_42, 1, x_41); +x_43 = l_Lean_PersistentArray_push___rarg(x_1, x_42); +x_44 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_44, 0, x_29); +lean_ctor_set(x_44, 1, x_30); +lean_ctor_set(x_44, 2, x_31); +lean_ctor_set(x_44, 3, x_43); +lean_ctor_set(x_44, 4, x_33); +lean_ctor_set(x_44, 5, x_34); +lean_ctor_set(x_44, 6, x_35); +x_45 = lean_st_ref_set(x_7, x_44, x_11); +x_46 = lean_ctor_get(x_45, 1); +lean_inc(x_46); +if (lean_is_exclusive(x_45)) { + lean_ctor_release(x_45, 0); + lean_ctor_release(x_45, 1); + x_47 = x_45; +} else { + lean_dec_ref(x_45); + x_47 = lean_box(0); +} +x_48 = lean_box(0); +if (lean_is_scalar(x_47)) { + x_49 = lean_alloc_ctor(0, 2, 0); +} else { + x_49 = x_47; +} +lean_ctor_set(x_49, 0, x_48); +lean_ctor_set(x_49, 1, x_46); +return x_49; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_addDecl___spec__11(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) { +_start: +{ +uint8_t x_9; +x_9 = !lean_is_exclusive(x_6); +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; +x_10 = lean_ctor_get(x_6, 5); +x_11 = l_Lean_replaceRef(x_3, x_10); +lean_dec(x_10); +lean_ctor_set(x_6, 5, x_11); +x_12 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_4, 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___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_addDecl___spec__12(x_1, x_2, x_3, x_13, x_5, x_6, x_7, x_14); +lean_dec(x_6); +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; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_16 = lean_ctor_get(x_6, 0); +x_17 = lean_ctor_get(x_6, 1); +x_18 = lean_ctor_get(x_6, 2); +x_19 = lean_ctor_get(x_6, 3); +x_20 = lean_ctor_get(x_6, 4); +x_21 = lean_ctor_get(x_6, 5); +x_22 = lean_ctor_get(x_6, 6); +x_23 = lean_ctor_get(x_6, 7); +x_24 = lean_ctor_get(x_6, 8); +x_25 = lean_ctor_get(x_6, 9); +x_26 = lean_ctor_get(x_6, 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_6); +x_27 = l_Lean_replaceRef(x_3, 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_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_4, x_28, x_7, 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); +lean_dec(x_29); +x_32 = l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_addDecl___spec__12(x_1, x_2, x_3, x_30, x_5, x_28, x_7, x_31); +lean_dec(x_28); +return x_32; +} +} +} +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_addDecl___spec__13(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_5; lean_object* x_6; +x_5 = lean_ctor_get(x_1, 0); +lean_inc(x_5); +x_6 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_6, 0, x_5); +lean_ctor_set(x_6, 1, x_4); +return x_6; +} +else +{ +lean_object* x_7; lean_object* x_8; +x_7 = lean_ctor_get(x_1, 0); +lean_inc(x_7); +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; +} +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = lean_apply_3(x_1, x_2, x_3, x_4); +if (lean_obj_tag(x_5) == 0) +{ +uint8_t x_6; +x_6 = !lean_is_exclusive(x_5); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; +x_7 = lean_ctor_get(x_5, 0); +x_8 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_8, 0, x_7); +lean_ctor_set(x_5, 0, x_8); +return x_5; +} +else +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_9 = lean_ctor_get(x_5, 0); +x_10 = lean_ctor_get(x_5, 1); +lean_inc(x_10); +lean_inc(x_9); +lean_dec(x_5); +x_11 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_11, 0, x_9); +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; +} +} +else +{ +uint8_t x_13; +x_13 = !lean_is_exclusive(x_5); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_ctor_get(x_5, 0); +x_15 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set_tag(x_5, 0); +lean_ctor_set(x_5, 0, x_15); +return x_5; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_16 = lean_ctor_get(x_5, 0); +x_17 = lean_ctor_get(x_5, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_5); +x_18 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_18, 0, x_16); +x_19 = lean_alloc_ctor(0, 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_withTraceNode___at_Lean_addDecl___spec__7___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) { +_start: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; +lean_inc(x_8); +x_11 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_addDecl___spec__11(x_1, x_2, x_3, x_6, x_4, x_8, x_9, x_10); +x_12 = lean_ctor_get(x_11, 1); +lean_inc(x_12); +lean_dec(x_11); +x_13 = l_MonadExcept_ofExcept___at_Lean_addDecl___spec__13(x_5, x_8, x_9, x_12); +lean_dec(x_8); +return x_13; +} +} +static lean_object* _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_profiler; +return x_1; +} +} +static lean_object* _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("[", 1); +return x_1; +} +} +static lean_object* _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__2; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("s] ", 3); +return x_1; +} +} +static lean_object* _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__4; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, uint8_t x_7, double 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_dec(x_9); +x_13 = lean_ctor_get(x_10, 5); +lean_inc(x_13); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_2); +x_14 = lean_apply_4(x_1, x_2, x_10, x_11, x_12); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +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_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__1; +x_18 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_6, x_17); +lean_dec(x_6); +if (x_18 == 0) +{ +if (x_7 == 0) +{ +lean_object* x_19; lean_object* x_20; +x_19 = lean_box(0); +x_20 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__2(x_3, x_4, x_13, x_5, x_2, x_15, x_19, x_10, x_11, x_16); +lean_dec(x_11); +lean_dec(x_2); +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; +x_21 = lean_float_to_string(x_8); +x_22 = lean_alloc_ctor(3, 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 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__3; +x_25 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_23); +x_26 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__5; +x_27 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_15); +x_29 = l_Lean_Core_instantiateValueLevelParams___lambda__2___closed__3; +x_30 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +x_31 = lean_box(0); +x_32 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__2(x_3, x_4, x_13, x_5, x_2, x_30, x_31, x_10, x_11, x_16); +lean_dec(x_11); +lean_dec(x_2); +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; 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_float_to_string(x_8); +x_34 = lean_alloc_ctor(3, 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 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__3; +x_37 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set(x_37, 1, x_35); +x_38 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__5; +x_39 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +x_40 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_15); +x_41 = l_Lean_Core_instantiateValueLevelParams___lambda__2___closed__3; +x_42 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_42, 0, x_40); +lean_ctor_set(x_42, 1, x_41); +x_43 = lean_box(0); +x_44 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__2(x_3, x_4, x_13, x_5, x_2, x_42, x_43, x_10, x_11, x_16); +lean_dec(x_11); +lean_dec(x_2); +return x_44; +} +} +else +{ +uint8_t x_45; +lean_dec(x_13); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_45 = !lean_is_exclusive(x_14); +if (x_45 == 0) +{ +return x_14; +} +else +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_46 = lean_ctor_get(x_14, 0); +x_47 = lean_ctor_get(x_14, 1); +lean_inc(x_47); +lean_inc(x_46); +lean_dec(x_14); +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; +} +} +} +} +static lean_object* _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_trace_profiler; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t 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) { +_start: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +lean_dec(x_7); +x_11 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_addDecl___spec__9___rarg(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_alloc_closure((void*)(l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__1), 4, 1); +lean_closure_set(x_14, 0, x_1); +lean_inc(x_9); +lean_inc(x_8); +x_15 = l_Lean_withSeconds___at_Lean_addDecl___spec__10(x_14, x_8, x_9, x_13); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; lean_object* x_65; uint8_t x_66; +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); +x_19 = lean_ctor_get(x_16, 1); +lean_inc(x_19); +lean_dec(x_16); +x_65 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__1; +x_66 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_5, x_65); +if (x_66 == 0) +{ +uint8_t x_67; +x_67 = 0; +x_20 = x_67; +goto block_64; +} +else +{ +double x_68; double x_69; uint8_t x_70; +x_68 = l_Lean_trace_profiler_threshold_getSecs(x_5); +x_69 = lean_unbox_float(x_19); +x_70 = lean_float_decLt(x_68, x_69); +x_20 = x_70; +goto block_64; +} +block_64: +{ +if (x_6 == 0) +{ +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +lean_dec(x_19); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +x_21 = lean_st_ref_take(x_9, x_17); +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_is_exclusive(x_22); +if (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_22, 3); +x_26 = l_Lean_PersistentArray_append___rarg(x_12, x_25); +lean_ctor_set(x_22, 3, x_26); +x_27 = lean_st_ref_set(x_9, x_22, x_23); +x_28 = lean_ctor_get(x_27, 1); +lean_inc(x_28); +lean_dec(x_27); +x_29 = l_MonadExcept_ofExcept___at_Lean_addDecl___spec__13(x_18, x_8, x_9, x_28); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_18); +if (lean_obj_tag(x_29) == 0) +{ +uint8_t x_30; +x_30 = !lean_is_exclusive(x_29); +if (x_30 == 0) +{ +return x_29; +} +else +{ +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 +{ +uint8_t x_34; +x_34 = !lean_is_exclusive(x_29); +if (x_34 == 0) +{ +return x_29; +} +else +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_29, 0); +x_36 = lean_ctor_get(x_29, 1); +lean_inc(x_36); +lean_inc(x_35); +lean_dec(x_29); +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_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; 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_22, 0); +x_39 = lean_ctor_get(x_22, 1); +x_40 = lean_ctor_get(x_22, 2); +x_41 = lean_ctor_get(x_22, 3); +x_42 = lean_ctor_get(x_22, 4); +x_43 = lean_ctor_get(x_22, 5); +x_44 = lean_ctor_get(x_22, 6); +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_dec(x_22); +x_45 = l_Lean_PersistentArray_append___rarg(x_12, x_41); +x_46 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_46, 0, x_38); +lean_ctor_set(x_46, 1, x_39); +lean_ctor_set(x_46, 2, x_40); +lean_ctor_set(x_46, 3, x_45); +lean_ctor_set(x_46, 4, x_42); +lean_ctor_set(x_46, 5, x_43); +lean_ctor_set(x_46, 6, x_44); +x_47 = lean_st_ref_set(x_9, x_46, x_23); +x_48 = lean_ctor_get(x_47, 1); +lean_inc(x_48); +lean_dec(x_47); +x_49 = l_MonadExcept_ofExcept___at_Lean_addDecl___spec__13(x_18, x_8, x_9, x_48); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_18); +if (lean_obj_tag(x_49) == 0) +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_50 = lean_ctor_get(x_49, 0); +lean_inc(x_50); +x_51 = lean_ctor_get(x_49, 1); +lean_inc(x_51); +if (lean_is_exclusive(x_49)) { + lean_ctor_release(x_49, 0); + lean_ctor_release(x_49, 1); + x_52 = x_49; +} else { + lean_dec_ref(x_49); + x_52 = lean_box(0); +} +if (lean_is_scalar(x_52)) { + x_53 = lean_alloc_ctor(0, 2, 0); +} else { + x_53 = x_52; +} +lean_ctor_set(x_53, 0, x_50); +lean_ctor_set(x_53, 1, x_51); +return x_53; +} +else +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_54 = lean_ctor_get(x_49, 0); +lean_inc(x_54); +x_55 = lean_ctor_get(x_49, 1); +lean_inc(x_55); +if (lean_is_exclusive(x_49)) { + lean_ctor_release(x_49, 0); + lean_ctor_release(x_49, 1); + x_56 = x_49; +} else { + lean_dec_ref(x_49); + x_56 = lean_box(0); +} +if (lean_is_scalar(x_56)) { + x_57 = lean_alloc_ctor(1, 2, 0); +} else { + x_57 = x_56; +} +lean_ctor_set(x_57, 0, x_54); +lean_ctor_set(x_57, 1, x_55); +return x_57; +} +} +} +else +{ +lean_object* x_58; double x_59; lean_object* x_60; +x_58 = lean_box(0); +x_59 = lean_unbox_float(x_19); +lean_dec(x_19); +x_60 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3(x_2, x_18, x_12, x_3, x_4, x_5, x_20, x_59, x_58, x_8, x_9, x_17); +return x_60; +} +} +else +{ +lean_object* x_61; double x_62; lean_object* x_63; +x_61 = lean_box(0); +x_62 = lean_unbox_float(x_19); +lean_dec(x_19); +x_63 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3(x_2, x_18, x_12, x_3, x_4, x_5, x_20, x_62, x_61, x_8, x_9, x_17); +return x_63; +} +} +} +else +{ +uint8_t x_71; +lean_dec(x_12); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +x_71 = !lean_is_exclusive(x_15); +if (x_71 == 0) +{ +return x_15; +} +else +{ +lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_72 = lean_ctor_get(x_15, 0); +x_73 = lean_ctor_get(x_15, 1); +lean_inc(x_73); +lean_inc(x_72); +lean_dec(x_15); +x_74 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_74, 0, x_72); +lean_ctor_set(x_74, 1, x_73); +return x_74; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__7(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) { +_start: +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_8 = lean_ctor_get(x_5, 2); +lean_inc(x_8); +lean_inc(x_1); +x_9 = l_Lean_isTracingEnabledFor___at_Lean_addDecl___spec__8(x_1, x_5, x_6, x_7); +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +x_11 = lean_unbox(x_10); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_12 = lean_ctor_get(x_9, 1); +lean_inc(x_12); +lean_dec(x_9); +x_13 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__1; +x_14 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_8, x_13); +if (x_14 == 0) +{ +lean_object* x_15; +lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_2); +lean_dec(x_1); +x_15 = lean_apply_3(x_3, x_5, x_6, 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; +} +} +} +else +{ +lean_object* x_24; uint8_t x_25; lean_object* x_26; +x_24 = lean_box(0); +x_25 = lean_unbox(x_10); +lean_dec(x_10); +x_26 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4(x_3, x_2, x_1, x_4, x_8, x_25, x_24, x_5, x_6, x_12); +return x_26; +} +} +else +{ +lean_object* x_27; lean_object* x_28; uint8_t x_29; lean_object* x_30; +x_27 = lean_ctor_get(x_9, 1); +lean_inc(x_27); +lean_dec(x_9); +x_28 = lean_box(0); +x_29 = lean_unbox(x_10); +lean_dec(x_10); +x_30 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4(x_3, x_2, x_1, x_4, x_8, x_29, x_28, x_5, x_6, x_27); +return x_30; +} +} +} +LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_addDecl___spec__14___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) { +_start: +{ +lean_object* x_8; lean_object* x_9; +x_8 = lean_apply_2(x_3, x_5, x_6); +x_9 = l_Lean_profileitIOUnsafe___rarg(x_1, x_2, x_8, x_4, x_7); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_addDecl___spec__14(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_profileitM___at_Lean_addDecl___spec__14___rarg___boxed), 7, 0); +return x_2; +} +} +static lean_object* _init_l_Lean_addDecl___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("typechecking declaration", 24); +return x_1; +} +} +static lean_object* _init_l_Lean_addDecl___lambda__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_addDecl___lambda__1___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_addDecl___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; +x_5 = l_Lean_addDecl___lambda__1___closed__2; +x_6 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_6, 0, x_5); +lean_ctor_set(x_6, 1, x_4); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_addDecl___lambda__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; lean_object* x_10; +lean_dec(x_2); +x_6 = lean_st_ref_get(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 = lean_ctor_get(x_7, 0); +lean_inc(x_9); +lean_dec(x_7); +x_10 = lean_add_decl(x_9, x_1); +lean_dec(x_1); +if (lean_obj_tag(x_10) == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +lean_dec(x_10); +x_12 = l_Lean_throwKernelException___at_Lean_addDecl___spec__1(x_11, x_3, x_4, x_8); +lean_dec(x_4); +return x_12; +} +else +{ +lean_object* x_13; lean_object* x_14; +x_13 = lean_ctor_get(x_10, 0); +lean_inc(x_13); +lean_dec(x_10); +x_14 = l_Lean_setEnv___at_Lean_addDecl___spec__3(x_13, x_3, x_4, x_8); +lean_dec(x_4); +lean_dec(x_3); +return x_14; +} +} +} +static lean_object* _init_l_Lean_addDecl___lambda__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("declaration uses 'sorry'", 24); +return x_1; +} +} +static lean_object* _init_l_Lean_addDecl___lambda__3___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_addDecl___lambda__3___closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_addDecl___lambda__3___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_addDecl___lambda__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_addDecl___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; uint8_t x_9; +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, 5); +lean_inc(x_8); +lean_dec(x_6); +x_9 = l_Lean_PersistentArray_anyM___at_Lean_MessageLog_hasErrors___spec__1(x_8); +if (x_9 == 0) +{ +uint8_t x_10; uint8_t x_11; +x_10 = 0; +lean_inc(x_1); +x_11 = l_Lean_Declaration_foldExprM___at_Lean_Declaration_hasSorry___spec__1(x_1, x_10); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; +x_12 = lean_box(0); +x_13 = l_Lean_addDecl___lambda__2(x_1, x_12, x_2, x_3, x_7); +return x_13; +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_14 = l_Lean_addDecl___lambda__3___closed__3; +lean_inc(x_3); lean_inc(x_2); x_15 = l_Lean_logWarning___at_Lean_addDecl___spec__4(x_14, x_2, x_3, x_7); x_16 = lean_ctor_get(x_15, 0); @@ -9684,7 +11013,7 @@ lean_inc(x_16); x_17 = lean_ctor_get(x_15, 1); lean_inc(x_17); lean_dec(x_15); -x_18 = l_Lean_addDecl___lambda__1(x_1, x_16, x_2, x_3, x_17); +x_18 = l_Lean_addDecl___lambda__2(x_1, x_16, x_2, x_3, x_17); return x_18; } } @@ -9692,7 +11021,7 @@ else { lean_object* x_19; lean_object* x_20; x_19 = lean_box(0); -x_20 = l_Lean_addDecl___lambda__1(x_1, x_19, x_2, x_3, x_7); +x_20 = l_Lean_addDecl___lambda__2(x_1, x_19, x_2, x_3, x_7); return x_20; } } @@ -9701,6 +11030,14 @@ static lean_object* _init_l_Lean_addDecl___closed__1() { _start: { lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_addDecl___lambda__1___boxed), 4, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_addDecl___closed__2() { +_start: +{ +lean_object* x_1; x_1 = lean_mk_string_from_bytes("type checking", 13); return x_1; } @@ -9708,16 +11045,25 @@ return x_1; LEAN_EXPORT lean_object* l_Lean_addDecl(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; uint8_t 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_2, 2); lean_inc(x_5); -x_6 = lean_alloc_closure((void*)(l_Lean_addDecl___lambda__2), 4, 1); +x_6 = lean_alloc_closure((void*)(l_Lean_addDecl___lambda__3), 4, 1); lean_closure_set(x_6, 0, x_1); -x_7 = l_Lean_addDecl___closed__1; -x_8 = lean_box(0); -x_9 = l_Lean_profileitM___at_Lean_addDecl___spec__7___rarg(x_7, x_5, x_6, x_8, x_2, x_3, x_4); +x_7 = l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__2; +x_8 = l_Lean_addDecl___closed__1; +x_9 = 1; +x_10 = lean_box(x_9); +x_11 = lean_alloc_closure((void*)(l_Lean_withTraceNode___at_Lean_addDecl___spec__7___boxed), 7, 4); +lean_closure_set(x_11, 0, x_7); +lean_closure_set(x_11, 1, x_8); +lean_closure_set(x_11, 2, x_6); +lean_closure_set(x_11, 3, x_10); +x_12 = l_Lean_addDecl___closed__2; +x_13 = lean_box(0); +x_14 = l_Lean_profileitM___at_Lean_addDecl___spec__14___rarg(x_12, x_5, x_11, x_13, x_2, x_3, x_4); lean_dec(x_5); -return x_9; +return x_14; } } LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_addDecl___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { @@ -9772,16 +11118,138 @@ x_7 = l_Lean_log___at_Lean_addDecl___spec__5(x_1, x_6, x_3, x_4, x_5); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_addDecl___spec__7___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_addDecl___spec__8___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_isTracingEnabledFor___at_Lean_addDecl___spec__8(x_1, x_2, x_3, x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_addDecl___spec__9___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_addDecl___spec__9___rarg(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_addDecl___spec__9___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_addDecl___spec__9(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_addDecl___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) { +_start: +{ +uint8_t x_9; lean_object* x_10; +x_9 = lean_unbox(x_5); +lean_dec(x_5); +x_10 = l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_addDecl___spec__12(x_1, x_2, x_3, x_4, x_9, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_6); +return x_10; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_addDecl___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) { +_start: +{ +uint8_t x_9; lean_object* x_10; +x_9 = lean_unbox(x_5); +lean_dec(x_5); +x_10 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_addDecl___spec__11(x_1, x_2, x_3, x_4, x_9, x_6, x_7, x_8); +lean_dec(x_7); +return x_10; +} +} +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_addDecl___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_MonadExcept_ofExcept___at_Lean_addDecl___spec__13(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_withTraceNode___at_Lean_addDecl___spec__7___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: +{ +uint8_t x_11; lean_object* x_12; +x_11 = lean_unbox(x_4); +lean_dec(x_4); +x_12 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__2(x_1, x_2, x_3, x_11, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_7); +lean_dec(x_5); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__7___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) { +_start: +{ +uint8_t x_13; uint8_t x_14; double x_15; lean_object* x_16; +x_13 = lean_unbox(x_5); +lean_dec(x_5); +x_14 = lean_unbox(x_7); +lean_dec(x_7); +x_15 = lean_unbox_float(x_8); +lean_dec(x_8); +x_16 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3(x_1, x_2, x_3, x_4, x_13, x_6, x_14, x_15, x_9, x_10, x_11, x_12); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__7___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: +{ +uint8_t x_11; uint8_t x_12; lean_object* x_13; +x_11 = lean_unbox(x_4); +lean_dec(x_4); +x_12 = lean_unbox(x_6); +lean_dec(x_6); +x_13 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4(x_1, x_2, x_3, x_11, x_5, x_12, x_7, x_8, x_9, x_10); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___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) { +_start: +{ +uint8_t x_8; lean_object* x_9; +x_8 = lean_unbox(x_4); +lean_dec(x_4); +x_9 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7(x_1, x_2, x_3, x_8, x_5, x_6, x_7); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_addDecl___spec__14___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { lean_object* x_8; -x_8 = l_Lean_profileitM___at_Lean_addDecl___spec__7___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +x_8 = l_Lean_profileitM___at_Lean_addDecl___spec__14___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7); lean_dec(x_2); lean_dec(x_1); return x_8; } } +LEAN_EXPORT lean_object* l_Lean_addDecl___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_addDecl___lambda__1(x_1, x_2, x_3, x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_5; +} +} static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__1() { _start: { @@ -10993,6 +12461,77 @@ lean_dec(x_5); return x_6; } } +static lean_object* _init_l_Lean_initFn____x40_Lean_CoreM___hyg_5108____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("compiler", 8); +return x_1; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_CoreM___hyg_5108____closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("enableNew", 9); +return x_1; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_CoreM___hyg_5108____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_initFn____x40_Lean_CoreM___hyg_5108____closed__1; +x_2 = l_Lean_initFn____x40_Lean_CoreM___hyg_5108____closed__2; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_CoreM___hyg_5108____closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("(compiler) enable the new code generator, this should have no significant effect on your code but it does help to test the new code generator; unset to only use the old code generator instead", 191); +return x_1; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_CoreM___hyg_5108____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_initFn____x40_Lean_CoreM___hyg_5108____closed__1; +x_3 = l_Lean_initFn____x40_Lean_CoreM___hyg_5108____closed__4; +x_4 = lean_box(x_1); +x_5 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_5, 0, x_4); +lean_ctor_set(x_5, 1, x_2); +lean_ctor_set(x_5, 2, x_3); +return x_5; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_CoreM___hyg_5108____closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Core_initFn____x40_Lean_CoreM___hyg_7____closed__6; +x_2 = l_Lean_initFn____x40_Lean_CoreM___hyg_5108____closed__1; +x_3 = l_Lean_initFn____x40_Lean_CoreM___hyg_5108____closed__2; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_CoreM___hyg_5108_(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_initFn____x40_Lean_CoreM___hyg_5108____closed__3; +x_3 = l_Lean_initFn____x40_Lean_CoreM___hyg_5108____closed__5; +x_4 = l_Lean_initFn____x40_Lean_CoreM___hyg_5108____closed__6; +x_5 = l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__1(x_2, x_3, x_4, x_1); +return x_5; +} +} LEAN_EXPORT lean_object* l_Lean_compileDeclsNew___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { @@ -11889,136 +13428,170 @@ lean_dec(x_2); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_compileDecl(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_compileDecl___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_5; lean_object* x_6; -lean_inc(x_1); -x_5 = l_Lean_Compiler_getDeclNamesForCodeGen(x_1); -lean_inc(x_3); -lean_inc(x_2); -x_6 = lean_lcnf_compile_decls(x_5, x_2, x_3, x_4); -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; lean_object* x_12; lean_object* x_13; -x_7 = lean_ctor_get(x_6, 1); -lean_inc(x_7); -lean_dec(x_6); -x_8 = lean_st_ref_get(x_3, x_7); -x_9 = lean_ctor_get(x_8, 0); +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_7 = lean_st_ref_get(x_5, x_6); +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +x_9 = lean_ctor_get(x_7, 1); lean_inc(x_9); -x_10 = lean_ctor_get(x_8, 1); +lean_dec(x_7); +x_10 = lean_ctor_get(x_8, 0); lean_inc(x_10); lean_dec(x_8); -x_11 = lean_ctor_get(x_9, 0); -lean_inc(x_11); -lean_dec(x_9); -x_12 = lean_ctor_get(x_2, 2); -lean_inc(x_12); -lean_inc(x_1); -x_13 = l_Lean_Environment_compileDecl(x_11, x_12, x_1); -lean_dec(x_12); -if (lean_obj_tag(x_13) == 0) +lean_inc(x_2); +x_11 = l_Lean_Environment_compileDecl(x_10, x_1, x_2); +if (lean_obj_tag(x_11) == 0) { -lean_object* x_14; -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -lean_dec(x_13); -if (lean_obj_tag(x_14) == 11) +lean_object* x_12; +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +lean_dec(x_11); +if (lean_obj_tag(x_12) == 11) { -lean_object* x_15; lean_object* x_16; -x_15 = lean_ctor_get(x_14, 0); +lean_object* x_13; lean_object* x_14; +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +lean_dec(x_12); +lean_inc(x_5); +lean_inc(x_4); +x_14 = l___private_Lean_CoreM_0__Lean_checkUnsupported___at_Lean_compileDecl___spec__1(x_2, x_4, x_5, x_9); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_15 = lean_ctor_get(x_14, 1); lean_inc(x_15); lean_dec(x_14); -lean_inc(x_3); -lean_inc(x_2); -x_16 = l___private_Lean_CoreM_0__Lean_checkUnsupported___at_Lean_compileDecl___spec__1(x_1, x_2, x_3, x_10); -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 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_18, 0, x_15); -x_19 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_19, 0, x_18); -x_20 = l_Lean_throwError___at_Lean_addDecl___spec__2(x_19, x_2, x_3, x_17); -lean_dec(x_3); -lean_dec(x_2); -return x_20; +x_16 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_16, 0, x_13); +x_17 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_17, 0, x_16); +x_18 = l_Lean_throwError___at_Lean_addDecl___spec__2(x_17, x_4, x_5, x_15); +lean_dec(x_5); +lean_dec(x_4); +return x_18; } else { -uint8_t x_21; -lean_dec(x_15); -lean_dec(x_3); -lean_dec(x_2); -x_21 = !lean_is_exclusive(x_16); -if (x_21 == 0) +uint8_t x_19; +lean_dec(x_13); +lean_dec(x_5); +lean_dec(x_4); +x_19 = !lean_is_exclusive(x_14); +if (x_19 == 0) { -return x_16; +return x_14; } 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_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_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_25; -lean_dec(x_1); -x_25 = l_Lean_throwKernelException___at_Lean_addDecl___spec__1(x_14, x_2, x_3, x_10); -lean_dec(x_3); -return x_25; +lean_object* x_23; +lean_dec(x_2); +x_23 = l_Lean_throwKernelException___at_Lean_addDecl___spec__1(x_12, x_4, x_5, x_9); +lean_dec(x_5); +return x_23; } } else { -lean_object* x_26; lean_object* x_27; -lean_dec(x_1); -x_26 = lean_ctor_get(x_13, 0); -lean_inc(x_26); -lean_dec(x_13); -x_27 = l_Lean_setEnv___at_Lean_addDecl___spec__3(x_26, x_2, x_3, x_10); -lean_dec(x_3); +lean_object* x_24; lean_object* x_25; lean_dec(x_2); -return x_27; +x_24 = lean_ctor_get(x_11, 0); +lean_inc(x_24); +lean_dec(x_11); +x_25 = l_Lean_setEnv___at_Lean_addDecl___spec__3(x_24, x_4, x_5, x_9); +lean_dec(x_5); +lean_dec(x_4); +return x_25; } } +} +static lean_object* _init_l_Lean_compileDecl___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_compiler_enableNew; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_compileDecl(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; +x_5 = lean_ctor_get(x_2, 2); +lean_inc(x_5); +x_6 = l_Lean_compileDecl___closed__1; +x_7 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_5, x_6); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; +x_8 = lean_box(0); +x_9 = l_Lean_compileDecl___lambda__1(x_5, x_1, x_8, x_2, x_3, x_4); +lean_dec(x_5); +return x_9; +} else { -uint8_t x_28; +lean_object* x_10; lean_object* x_11; +lean_inc(x_1); +x_10 = l_Lean_Compiler_getDeclNamesForCodeGen(x_1); +lean_inc(x_3); +lean_inc(x_2); +x_11 = lean_lcnf_compile_decls(x_10, x_2, x_3, x_4); +if (lean_obj_tag(x_11) == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; +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 = l_Lean_compileDecl___lambda__1(x_5, x_1, x_12, x_2, x_3, x_13); +lean_dec(x_12); +lean_dec(x_5); +return x_14; +} +else +{ +uint8_t x_15; +lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_28 = !lean_is_exclusive(x_6); -if (x_28 == 0) +x_15 = !lean_is_exclusive(x_11); +if (x_15 == 0) { -return x_6; +return x_11; } else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = lean_ctor_get(x_6, 0); -x_30 = lean_ctor_get(x_6, 1); -lean_inc(x_30); -lean_inc(x_29); -lean_dec(x_6); -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_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_11, 0); +x_17 = lean_ctor_get(x_11, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_11); +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; +} } } } @@ -12063,21 +13636,21 @@ lean_dec(x_4); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_compileDecls(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_compileDecl___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_5; -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_5 = lean_lcnf_compile_decls(x_1, x_2, x_3, x_4); -if (lean_obj_tag(x_5) == 0) +lean_object* x_7; +x_7 = l_Lean_compileDecl___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_3); +lean_dec(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_compileDecls___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_6; 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_6 = lean_ctor_get(x_5, 1); -lean_inc(x_6); -lean_dec(x_5); -x_7 = lean_st_ref_get(x_3, x_6); +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_7 = lean_st_ref_get(x_5, x_6); x_8 = lean_ctor_get(x_7, 0); lean_inc(x_8); x_9 = lean_ctor_get(x_7, 1); @@ -12086,77 +13659,125 @@ lean_dec(x_7); x_10 = lean_ctor_get(x_8, 0); lean_inc(x_10); lean_dec(x_8); -x_11 = lean_ctor_get(x_2, 2); -lean_inc(x_11); -x_12 = lean_compile_decls(x_10, x_11, x_1); -lean_dec(x_1); +x_11 = lean_compile_decls(x_10, x_1, x_2); +if (lean_obj_tag(x_11) == 0) +{ +lean_object* x_12; +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); lean_dec(x_11); -if (lean_obj_tag(x_12) == 0) +if (lean_obj_tag(x_12) == 11) { -lean_object* x_13; +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; x_13 = lean_ctor_get(x_12, 0); lean_inc(x_13); lean_dec(x_12); -if (lean_obj_tag(x_13) == 11) -{ -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); -lean_dec(x_13); -x_15 = lean_alloc_ctor(3, 1, 0); +x_14 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_14, 0, x_13); +x_15 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_15, 0, x_14); -x_16 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_16, 0, x_15); -x_17 = l_Lean_throwError___at_Lean_addDecl___spec__2(x_16, x_2, x_3, x_9); -lean_dec(x_3); -lean_dec(x_2); -return x_17; +x_16 = l_Lean_throwError___at_Lean_addDecl___spec__2(x_15, x_4, x_5, x_9); +lean_dec(x_4); +return x_16; } else { -lean_object* x_18; -x_18 = l_Lean_throwKernelException___at_Lean_addDecl___spec__1(x_13, x_2, x_3, x_9); -lean_dec(x_3); -return x_18; +lean_object* x_17; +x_17 = l_Lean_throwKernelException___at_Lean_addDecl___spec__1(x_12, x_4, x_5, x_9); +return x_17; } } else { -lean_object* x_19; lean_object* x_20; -x_19 = lean_ctor_get(x_12, 0); -lean_inc(x_19); -lean_dec(x_12); -x_20 = l_Lean_setEnv___at_Lean_addDecl___spec__3(x_19, x_2, x_3, x_9); +lean_object* x_18; lean_object* x_19; +x_18 = lean_ctor_get(x_11, 0); +lean_inc(x_18); +lean_dec(x_11); +x_19 = l_Lean_setEnv___at_Lean_addDecl___spec__3(x_18, x_4, x_5, x_9); +lean_dec(x_4); +return x_19; +} +} +} +LEAN_EXPORT lean_object* l_Lean_compileDecls(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; +x_5 = lean_ctor_get(x_2, 2); +lean_inc(x_5); +x_6 = l_Lean_compileDecl___closed__1; +x_7 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_5, x_6); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; +x_8 = lean_box(0); +x_9 = l_Lean_compileDecls___lambda__1(x_5, x_1, x_8, x_2, x_3, x_4); lean_dec(x_3); -lean_dec(x_2); -return x_20; +lean_dec(x_1); +lean_dec(x_5); +return x_9; } +else +{ +lean_object* x_10; +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_10 = lean_lcnf_compile_decls(x_1, x_2, x_3, x_4); +if (lean_obj_tag(x_10) == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; +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_compileDecls___lambda__1(x_5, x_1, x_11, x_2, x_3, x_12); +lean_dec(x_3); +lean_dec(x_11); +lean_dec(x_1); +lean_dec(x_5); +return x_13; } else { -uint8_t x_21; +uint8_t x_14; +lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_21 = !lean_is_exclusive(x_5); -if (x_21 == 0) +x_14 = !lean_is_exclusive(x_10); +if (x_14 == 0) { -return x_5; +return x_10; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_5, 0); -x_23 = lean_ctor_get(x_5, 1); -lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_5); -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_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_10, 0); +x_16 = lean_ctor_get(x_10, 1); +lean_inc(x_16); +lean_inc(x_15); +lean_dec(x_10); +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_EXPORT lean_object* l_Lean_compileDecls___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_compileDecls___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_7; } } LEAN_EXPORT lean_object* l_Lean_addAndCompile(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { @@ -12487,7 +14108,38 @@ if (lean_io_result_is_error(res)) return res; l_Lean_Core_maxHeartbeats = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Core_maxHeartbeats); lean_dec_ref(res); -}l_Lean_Core_getMaxHeartbeats___closed__1 = _init_l_Lean_Core_getMaxHeartbeats___closed__1(); +}l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__1 = _init_l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__1(); +lean_mark_persistent(l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__1); +l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__2 = _init_l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__2(); +lean_mark_persistent(l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__2); +l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__3 = _init_l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__3(); +lean_mark_persistent(l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__3); +l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__4 = _init_l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__4(); +lean_mark_persistent(l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__4); +l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__5 = _init_l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__5(); +lean_mark_persistent(l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__5); +l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__6 = _init_l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__6(); +lean_mark_persistent(l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__6); +l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__7 = _init_l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__7(); +lean_mark_persistent(l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__7); +l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__8 = _init_l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__8(); +lean_mark_persistent(l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__8); +l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__9 = _init_l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__9(); +lean_mark_persistent(l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__9); +l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__10 = _init_l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__10(); +lean_mark_persistent(l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__10); +l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__11 = _init_l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__11(); +lean_mark_persistent(l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__11); +l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__12 = _init_l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__12(); +lean_mark_persistent(l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__12); +l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__13 = _init_l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__13(); +lean_mark_persistent(l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__13); +l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__14 = _init_l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__14(); +lean_mark_persistent(l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44____closed__14); +res = l_Lean_Core_initFn____x40_Lean_CoreM___hyg_44_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Lean_Core_getMaxHeartbeats___closed__1 = _init_l_Lean_Core_getMaxHeartbeats___closed__1(); lean_mark_persistent(l_Lean_Core_getMaxHeartbeats___closed__1); l_Lean_Core_Cache_instLevelType___default___closed__1 = _init_l_Lean_Core_Cache_instLevelType___default___closed__1(); lean_mark_persistent(l_Lean_Core_Cache_instLevelType___default___closed__1); @@ -12703,14 +14355,35 @@ l_Lean_mkArrow___closed__2 = _init_l_Lean_mkArrow___closed__2(); lean_mark_persistent(l_Lean_mkArrow___closed__2); l_Lean_logAt___at_Lean_addDecl___spec__6___closed__1 = _init_l_Lean_logAt___at_Lean_addDecl___spec__6___closed__1(); lean_mark_persistent(l_Lean_logAt___at_Lean_addDecl___spec__6___closed__1); -l_Lean_addDecl___lambda__2___closed__1 = _init_l_Lean_addDecl___lambda__2___closed__1(); -lean_mark_persistent(l_Lean_addDecl___lambda__2___closed__1); -l_Lean_addDecl___lambda__2___closed__2 = _init_l_Lean_addDecl___lambda__2___closed__2(); -lean_mark_persistent(l_Lean_addDecl___lambda__2___closed__2); -l_Lean_addDecl___lambda__2___closed__3 = _init_l_Lean_addDecl___lambda__2___closed__3(); -lean_mark_persistent(l_Lean_addDecl___lambda__2___closed__3); +l_Lean_isTracingEnabledFor___at_Lean_addDecl___spec__8___closed__1 = _init_l_Lean_isTracingEnabledFor___at_Lean_addDecl___spec__8___closed__1(); +lean_mark_persistent(l_Lean_isTracingEnabledFor___at_Lean_addDecl___spec__8___closed__1); +l_Lean_withSeconds___at_Lean_addDecl___spec__10___closed__1 = _init_l_Lean_withSeconds___at_Lean_addDecl___spec__10___closed__1(); +l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__1 = _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__1); +l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__2 = _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__2); +l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__3 = _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__3(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__3); +l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__4 = _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__4(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__4); +l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__5 = _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__5(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__5); +l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__1 = _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__1); +l_Lean_addDecl___lambda__1___closed__1 = _init_l_Lean_addDecl___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_addDecl___lambda__1___closed__1); +l_Lean_addDecl___lambda__1___closed__2 = _init_l_Lean_addDecl___lambda__1___closed__2(); +lean_mark_persistent(l_Lean_addDecl___lambda__1___closed__2); +l_Lean_addDecl___lambda__3___closed__1 = _init_l_Lean_addDecl___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_addDecl___lambda__3___closed__1); +l_Lean_addDecl___lambda__3___closed__2 = _init_l_Lean_addDecl___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_addDecl___lambda__3___closed__2); +l_Lean_addDecl___lambda__3___closed__3 = _init_l_Lean_addDecl___lambda__3___closed__3(); +lean_mark_persistent(l_Lean_addDecl___lambda__3___closed__3); l_Lean_addDecl___closed__1 = _init_l_Lean_addDecl___closed__1(); lean_mark_persistent(l_Lean_addDecl___closed__1); +l_Lean_addDecl___closed__2 = _init_l_Lean_addDecl___closed__2(); +lean_mark_persistent(l_Lean_addDecl___closed__2); l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__1 = _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__1(); lean_mark_persistent(l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__1); l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__2 = _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__2(); @@ -12779,6 +14452,25 @@ l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rar lean_mark_persistent(l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg___lambda__2___closed__3); l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg___lambda__2___closed__4 = _init_l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg___lambda__2___closed__4(); lean_mark_persistent(l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg___lambda__2___closed__4); +l_Lean_initFn____x40_Lean_CoreM___hyg_5108____closed__1 = _init_l_Lean_initFn____x40_Lean_CoreM___hyg_5108____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_CoreM___hyg_5108____closed__1); +l_Lean_initFn____x40_Lean_CoreM___hyg_5108____closed__2 = _init_l_Lean_initFn____x40_Lean_CoreM___hyg_5108____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_CoreM___hyg_5108____closed__2); +l_Lean_initFn____x40_Lean_CoreM___hyg_5108____closed__3 = _init_l_Lean_initFn____x40_Lean_CoreM___hyg_5108____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_CoreM___hyg_5108____closed__3); +l_Lean_initFn____x40_Lean_CoreM___hyg_5108____closed__4 = _init_l_Lean_initFn____x40_Lean_CoreM___hyg_5108____closed__4(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_CoreM___hyg_5108____closed__4); +l_Lean_initFn____x40_Lean_CoreM___hyg_5108____closed__5 = _init_l_Lean_initFn____x40_Lean_CoreM___hyg_5108____closed__5(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_CoreM___hyg_5108____closed__5); +l_Lean_initFn____x40_Lean_CoreM___hyg_5108____closed__6 = _init_l_Lean_initFn____x40_Lean_CoreM___hyg_5108____closed__6(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_CoreM___hyg_5108____closed__6); +if (builtin) {res = l_Lean_initFn____x40_Lean_CoreM___hyg_5108_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +l_Lean_compiler_enableNew = lean_io_result_get_value(res); +lean_mark_persistent(l_Lean_compiler_enableNew); +lean_dec_ref(res); +}l_Lean_compileDecl___closed__1 = _init_l_Lean_compileDecl___closed__1(); +lean_mark_persistent(l_Lean_compileDecl___closed__1); l_Lean_ImportM_runCoreM___rarg___closed__1 = _init_l_Lean_ImportM_runCoreM___rarg___closed__1(); lean_mark_persistent(l_Lean_ImportM_runCoreM___rarg___closed__1); return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Data/Json/FromToJson.c b/stage0/stdlib/Lean/Data/Json/FromToJson.c index 23a0e245dcc1..ee99fee2eb71 100644 --- a/stage0/stdlib/Lean/Data/Json/FromToJson.c +++ b/stage0/stdlib/Lean/Data/Json/FromToJson.c @@ -17,7 +17,6 @@ lean_object* l_Lean_JsonNumber_fromNat(lean_object*); LEAN_EXPORT lean_object* l_Lean_instToJsonRBMapString___rarg(lean_object*, lean_object*); static lean_object* l_Lean_instFromJsonArray___rarg___closed__2; lean_object* l_Lean_Json_getObj_x3f(lean_object*); -lean_object* l_Lean_Syntax_decodeNameLit(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instFromJsonList___spec__1___rarg(lean_object*, size_t, size_t, lean_object*); extern lean_object* l_UInt64_size; LEAN_EXPORT lean_object* l_Lean_instFromJsonUInt64___lambda__1(lean_object*, lean_object*); @@ -48,6 +47,7 @@ LEAN_EXPORT lean_object* l_Lean_instToJsonRBMapString___boxed(lean_object*, lean static lean_object* l_Lean_instFromJsonFloat___closed__1; LEAN_EXPORT lean_object* l_Lean_instToJsonOption___rarg(lean_object*, lean_object*); lean_object* l_Lean_JsonNumber_fromFloat_x3f(double); +uint8_t l_Lean_Name_isAnonymous(lean_object*); static lean_object* l_Lean_instFromJsonFloat___closed__8; static double l_Lean_instFromJsonFloat___closed__13; lean_object* lean_array_fget(lean_object*, lean_object*); @@ -56,6 +56,7 @@ LEAN_EXPORT lean_object* l_Lean_instToJsonJsonNumber(lean_object*); LEAN_EXPORT lean_object* l_Lean_instToJsonBool(uint8_t); LEAN_EXPORT lean_object* l_Lean_instToJsonName(lean_object*); LEAN_EXPORT lean_object* l_Lean_instFromJsonUSize___lambda__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_instFromJsonName___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instFromJsonName(lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_opt___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_instFromJsonName___closed__2; @@ -73,7 +74,6 @@ lean_object* l_Lean_Json_getObjVal_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instToJsonUSize(size_t); static double l_Lean_instFromJsonFloat___closed__9; LEAN_EXPORT lean_object* l_Lean_instToJsonArray___rarg(lean_object*, lean_object*); -static lean_object* l_Lean_instFromJsonName___closed__4; lean_object* l_Lean_Json_getStr_x3f(lean_object*); LEAN_EXPORT lean_object* l_Lean_instToJsonProd___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f(lean_object*, lean_object*, lean_object*, lean_object*); @@ -124,6 +124,7 @@ LEAN_EXPORT lean_object* l_Lean_instFromJsonJson(lean_object*); static lean_object* l_Lean_instToJsonJson___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instToJsonArray___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_instFromJsonList___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_instFromJsonName___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instFromJsonOption___rarg(lean_object*, lean_object*); static double l_Lean_instFromJsonFloat___closed__6; lean_object* l_Lean_Syntax_decodeNatLitVal_x3f(lean_object*); @@ -152,6 +153,7 @@ LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___boxed(lean_object*, lean_ static lean_object* l_Lean_instFromJsonJsonNumber___closed__1; lean_object* l_id___rarg___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_instFromJsonRBMapString___rarg(lean_object*, lean_object*, lean_object*); +lean_object* l_String_toName(lean_object*); LEAN_EXPORT lean_object* l_Lean_instFromJsonList___rarg(lean_object*, lean_object*); extern lean_object* l_USize_size; static double l_Lean_instFromJsonFloat___closed__7; @@ -1152,23 +1154,24 @@ x_3 = lean_alloc_closure((void*)(l_Lean_instToJsonProd___rarg), 3, 0); return x_3; } } -static lean_object* _init_l_Lean_instFromJsonName___closed__1() { +LEAN_EXPORT lean_object* l_Lean_instFromJsonName___lambda__1(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("[anonymous]", 11); -return x_1; +lean_object* x_3; +x_3 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_3, 0, x_1); +return x_3; } } -static lean_object* _init_l_Lean_instFromJsonName___closed__2() { +static lean_object* _init_l_Lean_instFromJsonName___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("`", 1); +x_1 = lean_mk_string_from_bytes("[anonymous]", 11); return x_1; } } -static lean_object* _init_l_Lean_instFromJsonName___closed__3() { +static lean_object* _init_l_Lean_instFromJsonName___closed__2() { _start: { lean_object* x_1; @@ -1176,7 +1179,7 @@ x_1 = lean_mk_string_from_bytes("expected a `Name`, got '", 24); return x_1; } } -static lean_object* _init_l_Lean_instFromJsonName___closed__4() { +static lean_object* _init_l_Lean_instFromJsonName___closed__3() { _start: { lean_object* x_1; lean_object* x_2; @@ -1223,97 +1226,97 @@ x_8 = l_Lean_instFromJsonName___closed__1; x_9 = lean_string_dec_eq(x_7, x_8); if (x_9 == 0) { -lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = l_Lean_instFromJsonName___closed__2; -x_11 = lean_string_append(x_10, x_7); -lean_dec(x_7); -x_12 = l_Lean_Syntax_decodeNameLit(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; -x_13 = lean_unsigned_to_nat(80u); -x_14 = l_Lean_Json_pretty(x_1, x_13); -x_15 = l_Lean_instFromJsonName___closed__3; -x_16 = lean_string_append(x_15, x_14); -lean_dec(x_14); -x_17 = l_Lean_instFromJsonArray___rarg___closed__2; -x_18 = lean_string_append(x_16, x_17); -lean_ctor_set_tag(x_2, 0); -lean_ctor_set(x_2, 0, x_18); +lean_object* x_10; uint8_t x_11; +x_10 = l_String_toName(x_7); +x_11 = l_Lean_Name_isAnonymous(x_10); +if (x_11 == 0) +{ +lean_dec(x_1); +lean_ctor_set(x_2, 0, x_10); return x_2; } else { -lean_object* x_19; -lean_dec(x_1); -x_19 = lean_ctor_get(x_12, 0); -lean_inc(x_19); -lean_dec(x_12); -lean_ctor_set(x_2, 0, x_19); +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_dec(x_10); +x_12 = lean_unsigned_to_nat(80u); +x_13 = l_Lean_Json_pretty(x_1, x_12); +x_14 = l_Lean_instFromJsonName___closed__2; +x_15 = lean_string_append(x_14, x_13); +lean_dec(x_13); +x_16 = l_Lean_instFromJsonArray___rarg___closed__2; +x_17 = lean_string_append(x_15, x_16); +lean_ctor_set_tag(x_2, 0); +lean_ctor_set(x_2, 0, x_17); return x_2; } } else { -lean_object* x_20; +lean_object* x_18; lean_free_object(x_2); lean_dec(x_7); lean_dec(x_1); -x_20 = l_Lean_instFromJsonName___closed__4; -return x_20; +x_18 = l_Lean_instFromJsonName___closed__3; +return x_18; } } else { -lean_object* x_21; lean_object* x_22; uint8_t x_23; -x_21 = lean_ctor_get(x_2, 0); -lean_inc(x_21); +lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_19 = lean_ctor_get(x_2, 0); +lean_inc(x_19); lean_dec(x_2); -x_22 = l_Lean_instFromJsonName___closed__1; -x_23 = lean_string_dec_eq(x_21, x_22); +x_20 = l_Lean_instFromJsonName___closed__1; +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_String_toName(x_19); +x_23 = l_Lean_Name_isAnonymous(x_22); if (x_23 == 0) { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = l_Lean_instFromJsonName___closed__2; -x_25 = lean_string_append(x_24, x_21); -lean_dec(x_21); -x_26 = l_Lean_Syntax_decodeNameLit(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; lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_27 = lean_unsigned_to_nat(80u); -x_28 = l_Lean_Json_pretty(x_1, x_27); -x_29 = l_Lean_instFromJsonName___closed__3; -x_30 = lean_string_append(x_29, x_28); -lean_dec(x_28); -x_31 = l_Lean_instFromJsonArray___rarg___closed__2; -x_32 = lean_string_append(x_30, x_31); -x_33 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_33, 0, x_32); -return x_33; +lean_object* x_24; +lean_dec(x_1); +x_24 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_24, 0, x_22); +return x_24; } else { -lean_object* x_34; lean_object* x_35; -lean_dec(x_1); -x_34 = lean_ctor_get(x_26, 0); -lean_inc(x_34); +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +lean_dec(x_22); +x_25 = lean_unsigned_to_nat(80u); +x_26 = l_Lean_Json_pretty(x_1, x_25); +x_27 = l_Lean_instFromJsonName___closed__2; +x_28 = lean_string_append(x_27, x_26); lean_dec(x_26); -x_35 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_35, 0, x_34); -return x_35; +x_29 = l_Lean_instFromJsonArray___rarg___closed__2; +x_30 = lean_string_append(x_28, x_29); +x_31 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_31, 0, x_30); +return x_31; } } else { -lean_object* x_36; -lean_dec(x_21); +lean_object* x_32; +lean_dec(x_19); lean_dec(x_1); -x_36 = l_Lean_instFromJsonName___closed__4; -return x_36; +x_32 = l_Lean_instFromJsonName___closed__3; +return x_32; +} +} } } } +LEAN_EXPORT lean_object* l_Lean_instFromJsonName___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_instFromJsonName___lambda__1(x_1, x_2); +lean_dec(x_2); +return x_3; } } LEAN_EXPORT lean_object* l_Lean_instToJsonName(lean_object* x_1) { @@ -2977,8 +2980,6 @@ l_Lean_instFromJsonName___closed__2 = _init_l_Lean_instFromJsonName___closed__2( lean_mark_persistent(l_Lean_instFromJsonName___closed__2); l_Lean_instFromJsonName___closed__3 = _init_l_Lean_instFromJsonName___closed__3(); lean_mark_persistent(l_Lean_instFromJsonName___closed__3); -l_Lean_instFromJsonName___closed__4 = _init_l_Lean_instFromJsonName___closed__4(); -lean_mark_persistent(l_Lean_instFromJsonName___closed__4); l_Lean_bignumFromJson_x3f___closed__1 = _init_l_Lean_bignumFromJson_x3f___closed__1(); lean_mark_persistent(l_Lean_bignumFromJson_x3f___closed__1); l_Lean_instFromJsonUSize___closed__1 = _init_l_Lean_instFromJsonUSize___closed__1(); diff --git a/stage0/stdlib/Lean/Data/Lsp/Extra.c b/stage0/stdlib/Lean/Data/Lsp/Extra.c index 53718be5d99e..cf9c00d5eaf1 100644 --- a/stage0/stdlib/Lean/Data/Lsp/Extra.c +++ b/stage0/stdlib/Lean/Data/Lsp/Extra.c @@ -24,7 +24,6 @@ static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonWaitForD lean_object* l_List_join___rarg(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanFileProgressProcessingInfo____x40_Lean_Data_Lsp_Extra___hyg_422____closed__11; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonRpcReleaseParams____x40_Lean_Data_Lsp_Extra___hyg_2251_(lean_object*); -lean_object* l_Lean_Syntax_decodeNameLit(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainGoalParams____x40_Lean_Data_Lsp_Extra___hyg_795____closed__10; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainGoalParams____x40_Lean_Data_Lsp_Extra___hyg_795____closed__11; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_1743____closed__15; @@ -75,7 +74,9 @@ static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLineRang lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5568____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_reprLineRange____x40_Lean_Data_Lsp_Extra___hyg_2525____closed__11; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainTermGoalParams____x40_Lean_Data_Lsp_Extra___hyg_1148____closed__1; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_1743____spec__1___lambda__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanFileProgressParams____x40_Lean_Data_Lsp_Extra___hyg_601____closed__3; +uint8_t l_Lean_Name_isAnonymous(lean_object*); 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_Extra_0__Lean_Lsp_fromJsonPlainGoal____x40_Lean_Data_Lsp_Extra___hyg_971____closed__7; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_reprLineRange____x40_Lean_Data_Lsp_Extra___hyg_2525_(lean_object*, lean_object*); @@ -131,6 +132,7 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Ext static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainTermGoal____x40_Lean_Data_Lsp_Extra___hyg_1321____closed__1; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcConnected____x40_Lean_Data_Lsp_Extra___hyg_1600____closed__6; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainGoal____x40_Lean_Data_Lsp_Extra___hyg_971____closed__11; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_1743____spec__1___lambda__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonRpcCallParams___closed__1; lean_object* l_Lean_Json_getStr_x3f(lean_object*); static lean_object* l_Lean_Lsp_instFromJsonPlainTermGoal___closed__1; @@ -189,7 +191,6 @@ static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_reprLineRange___ LEAN_EXPORT lean_object* l_Lean_Lsp_LeanFileProgressKind_toCtorIdx___boxed(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcReleaseParams____x40_Lean_Data_Lsp_Extra___hyg_2099____closed__13; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonLeanFileProgressProcessingInfo____x40_Lean_Data_Lsp_Extra___hyg_531____closed__2; -static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_1743____spec__1___closed__4; LEAN_EXPORT lean_object* l_Lean_Lsp_LeanFileProgressKind_noConfusion(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonPlainTermGoal____x40_Lean_Data_Lsp_Extra___hyg_1321____boxed(lean_object*); static lean_object* l_Lean_Lsp_instToJsonPlainGoal___closed__1; @@ -310,6 +311,7 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Ext static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_1743____spec__1___closed__1; static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcConnected____x40_Lean_Data_Lsp_Extra___hyg_1600____closed__9; LEAN_EXPORT lean_object* l_Lean_Lsp_instBEqLeanFileProgressKind; +lean_object* l_String_toName(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_1743____closed__8; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_toJsonLeanFileProgressParams____x40_Lean_Data_Lsp_Extra___hyg_729_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonLeanFileProgressKind___boxed(lean_object*); @@ -3999,23 +4001,24 @@ x_1 = l_Lean_Lsp_instToJsonRpcConnected___closed__1; return x_1; } } -static lean_object* _init_l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_1743____spec__1___closed__1() { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_1743____spec__1___lambda__1(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("[anonymous]", 11); -return x_1; +lean_object* x_3; +x_3 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_3, 0, x_1); +return x_3; } } -static lean_object* _init_l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_1743____spec__1___closed__2() { +static lean_object* _init_l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_1743____spec__1___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("`", 1); +x_1 = lean_mk_string_from_bytes("[anonymous]", 11); return x_1; } } -static lean_object* _init_l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_1743____spec__1___closed__3() { +static lean_object* _init_l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_1743____spec__1___closed__2() { _start: { lean_object* x_1; @@ -4023,7 +4026,7 @@ x_1 = lean_mk_string_from_bytes("expected a `Name`, got '", 24); return x_1; } } -static lean_object* _init_l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_1743____spec__1___closed__4() { +static lean_object* _init_l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_1743____spec__1___closed__3() { _start: { lean_object* x_1; lean_object* x_2; @@ -4071,94 +4074,85 @@ x_10 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Ls x_11 = lean_string_dec_eq(x_9, x_10); if (x_11 == 0) { -lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_12 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_1743____spec__1___closed__2; -x_13 = lean_string_append(x_12, x_9); -lean_dec(x_9); -x_14 = l_Lean_Syntax_decodeNameLit(x_13); -if (lean_obj_tag(x_14) == 0) +lean_object* x_12; uint8_t x_13; +x_12 = l_String_toName(x_9); +x_13 = l_Lean_Name_isAnonymous(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; -x_15 = lean_unsigned_to_nat(80u); -x_16 = l_Lean_Json_pretty(x_3, x_15); -x_17 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_1743____spec__1___closed__3; -x_18 = lean_string_append(x_17, x_16); -lean_dec(x_16); -x_19 = l_Lean_Lsp_instFromJsonLeanFileProgressKind___closed__2; -x_20 = lean_string_append(x_18, x_19); -lean_ctor_set_tag(x_4, 0); -lean_ctor_set(x_4, 0, x_20); +lean_dec(x_3); +lean_ctor_set(x_4, 0, x_12); return x_4; } else { -lean_object* x_21; -lean_dec(x_3); -x_21 = lean_ctor_get(x_14, 0); -lean_inc(x_21); -lean_dec(x_14); -lean_ctor_set(x_4, 0, x_21); +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_12); +x_14 = lean_unsigned_to_nat(80u); +x_15 = l_Lean_Json_pretty(x_3, x_14); +x_16 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_1743____spec__1___closed__2; +x_17 = lean_string_append(x_16, x_15); +lean_dec(x_15); +x_18 = l_Lean_Lsp_instFromJsonLeanFileProgressKind___closed__2; +x_19 = lean_string_append(x_17, x_18); +lean_ctor_set_tag(x_4, 0); +lean_ctor_set(x_4, 0, x_19); return x_4; } } else { -lean_object* x_22; +lean_object* x_20; lean_free_object(x_4); lean_dec(x_9); lean_dec(x_3); -x_22 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_1743____spec__1___closed__4; -return x_22; +x_20 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_1743____spec__1___closed__3; +return x_20; } } else { -lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_23 = lean_ctor_get(x_4, 0); -lean_inc(x_23); +lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_21 = lean_ctor_get(x_4, 0); +lean_inc(x_21); lean_dec(x_4); -x_24 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_1743____spec__1___closed__1; -x_25 = lean_string_dec_eq(x_23, x_24); -if (x_25 == 0) +x_22 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_1743____spec__1___closed__1; +x_23 = lean_string_dec_eq(x_21, x_22); +if (x_23 == 0) { -lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_26 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_1743____spec__1___closed__2; -x_27 = lean_string_append(x_26, x_23); -lean_dec(x_23); -x_28 = l_Lean_Syntax_decodeNameLit(x_27); -if (lean_obj_tag(x_28) == 0) +lean_object* x_24; uint8_t x_25; +x_24 = l_String_toName(x_21); +x_25 = l_Lean_Name_isAnonymous(x_24); +if (x_25 == 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; -x_29 = lean_unsigned_to_nat(80u); -x_30 = l_Lean_Json_pretty(x_3, x_29); -x_31 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_1743____spec__1___closed__3; -x_32 = lean_string_append(x_31, x_30); -lean_dec(x_30); -x_33 = l_Lean_Lsp_instFromJsonLeanFileProgressKind___closed__2; -x_34 = lean_string_append(x_32, x_33); -x_35 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_35, 0, x_34); -return x_35; +lean_object* x_26; +lean_dec(x_3); +x_26 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_26, 0, x_24); +return x_26; } else { -lean_object* x_36; lean_object* x_37; -lean_dec(x_3); -x_36 = lean_ctor_get(x_28, 0); -lean_inc(x_36); +lean_object* x_27; lean_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_24); +x_27 = lean_unsigned_to_nat(80u); +x_28 = l_Lean_Json_pretty(x_3, x_27); +x_29 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_1743____spec__1___closed__2; +x_30 = lean_string_append(x_29, x_28); lean_dec(x_28); -x_37 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_37, 0, x_36); -return x_37; +x_31 = l_Lean_Lsp_instFromJsonLeanFileProgressKind___closed__2; +x_32 = lean_string_append(x_30, x_31); +x_33 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_33, 0, x_32); +return x_33; } } else { -lean_object* x_38; -lean_dec(x_23); +lean_object* x_34; +lean_dec(x_21); lean_dec(x_3); -x_38 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_1743____spec__1___closed__4; -return x_38; +x_34 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_1743____spec__1___closed__3; +return x_34; } } } @@ -4575,6 +4569,15 @@ return x_57; } } } +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_1743____spec__1___lambda__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_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_1743____spec__1___lambda__1(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_1743____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { @@ -6417,8 +6420,6 @@ l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJ lean_mark_persistent(l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_1743____spec__1___closed__2); l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_1743____spec__1___closed__3 = _init_l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_1743____spec__1___closed__3(); lean_mark_persistent(l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_1743____spec__1___closed__3); -l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_1743____spec__1___closed__4 = _init_l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_1743____spec__1___closed__4(); -lean_mark_persistent(l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_1743____spec__1___closed__4); l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_1743____closed__1 = _init_l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_1743____closed__1(); lean_mark_persistent(l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_1743____closed__1); l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_1743____closed__2 = _init_l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcCallParams____x40_Lean_Data_Lsp_Extra___hyg_1743____closed__2(); diff --git a/stage0/stdlib/Lean/Data/Lsp/Internal.c b/stage0/stdlib/Lean/Data/Lsp/Internal.c index 2ac330edf7b4..1521ad3ea2d1 100644 --- a/stage0/stdlib/Lean/Data/Lsp/Internal.c +++ b/stage0/stdlib/Lean/Data/Lsp/Internal.c @@ -13,56 +13,56 @@ #ifdef __cplusplus extern "C" { #endif +static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__10; +static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__6; lean_object* l_Lean_JsonNumber_fromNat(lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_Lsp_instToJsonModuleRefs___spec__3(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__17; lean_object* l_List_join___rarg(lean_object*); lean_object* l_Lean_Json_getObj_x3f(lean_object*); -lean_object* l_Lean_Syntax_decodeNameLit(lean_object*); static lean_object* l_Lean_Lsp_instHashableRefIdent___closed__1; -static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__13; LEAN_EXPORT lean_object* l_Lean_Lsp_instInhabitedRefIdent; lean_object* lean_mk_empty_array_with_capacity(lean_object*); LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_Lsp_instFromJsonModuleRefs___spec__7(lean_object*); lean_object* l_Lean_Json_mkObj(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__5; LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_Lsp_instToJsonModuleRefs___spec__3___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__7; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instToJsonModuleRefs___spec__1(size_t, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Lsp_instToJsonModuleRefs___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Lsp_RefIdent_fromString___closed__3; uint64_t lean_uint64_mix_hash(uint64_t, uint64_t); lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_548____spec__1(lean_object*, lean_object*); lean_object* l_Lean_Name_toString(lean_object*, uint8_t); static lean_object* l_Lean_Lsp_instBEqRefIdent___closed__1; static lean_object* l_Lean_Lsp_RefIdent_toString___closed__1; lean_object* lean_array_push(lean_object*, lean_object*); +static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____spec__1___closed__1; lean_object* lean_mk_array(lean_object*, lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); -static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__14; +uint8_t l_Lean_Name_isAnonymous(lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_RefIdent_fromString(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonLeanIleanInfoParams; static lean_object* l_Lean_Lsp_RefIdent_fromString___lambda__1___closed__2; -static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__18; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonModuleRefs(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_hashRefIdent____x40_Lean_Data_Lsp_Internal___hyg_129____boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_replace___at_Lean_Lsp_instFromJsonModuleRefs___spec__6(lean_object*, lean_object*, lean_object*); uint8_t lean_string_dec_eq(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__2; static lean_object* l_Lean_Lsp_instInhabitedRefIdent___closed__1; static lean_object* l_Lean_Lsp_RefIdent_fromString___closed__1; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_toJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1412____spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_AssocList_contains___at_Lean_Lsp_instFromJsonModuleRefs___spec__2(lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_RefIdent_toString___closed__3; size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instFromJsonRefInfo___spec__5(size_t, size_t, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__4; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__13; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instFromJsonRefInfo___spec__4___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Json_getObjValD(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__18; LEAN_EXPORT lean_object* l_Lean_Lsp_RefIdent_toString(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonLeanIleanInfoParams; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210_(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instToJsonRefInfo___spec__2(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instToJsonRefInfo___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonRefInfo(lean_object*); @@ -72,56 +72,56 @@ LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Lsp_instToJsonModuleRefs___ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Lsp_instToJsonModuleRefs___spec__4(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_Lsp_instFromJsonRefInfo___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instToJsonRefInfo___spec__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__8; lean_object* l_List_lengthTRAux___rarg(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__1; uint8_t lean_name_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_RefIdent_fromString___lambda__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonRefInfo___closed__4; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonRefInfo___closed__3; -static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__15; -static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__17; static lean_object* l_Lean_Lsp_RefIdent_fromString___closed__2; +static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__12; LEAN_EXPORT lean_object* l_Lean_HashMapImp_expand___at_Lean_Lsp_instFromJsonModuleRefs___spec__3(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__9; size_t lean_data_hashmap_mk_idx(lean_object*, uint64_t); -static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__11; LEAN_EXPORT lean_object* l_Lean_Lsp_instHashableRefIdent; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instToJsonRefInfo___spec__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_beqRefIdent____x40_Lean_Data_Lsp_Internal___hyg_28_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_Lsp_instFromJsonModuleRefs___spec__7___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonModuleRefs(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_RefIdent_fromString___lambda__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_Lsp_instFromJsonRefInfo___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonRefInfo(lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMap_toList___at_Lean_Lsp_instToJsonModuleRefs___spec__2(lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_toJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1412____spec__1(size_t, size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instFromJsonRefInfo___spec__2(size_t, size_t, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__10; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonRefInfo___lambda__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__14; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_Lsp_instFromJsonRefInfo___spec__3(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__16; static lean_object* l_Lean_Lsp_RefIdent_fromString___lambda__1___closed__1; uint64_t l_Lean_Name_hash___override(lean_object*); +lean_object* l_String_toName(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instBEqRefIdent; +static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__5; lean_object* lean_nat_mul(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_toJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1395_(lean_object*); lean_object* l_Lean_Json_getNat_x3f(lean_object*); static lean_object* l_Lean_Lsp_instFromJsonLeanIleanInfoParams___closed__1; static lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_Lsp_instFromJsonRefInfo___spec__1___closed__3; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_toJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1395____spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_toJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1395____spec__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_hashRefIdent____x40_Lean_Data_Lsp_Internal___hyg_129_(lean_object*); lean_object* l_List_reverse___rarg(lean_object*); -LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_toJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1412____spec__2(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_Lsp_instFromJsonRefInfo___spec__1___closed__1; static lean_object* l_Lean_Lsp_instToJsonRefInfo___closed__2; lean_object* l_Lean_mkHashMapImp___rarg(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__11; LEAN_EXPORT lean_object* l_Lean_HashMapImp_moveEntries___at_Lean_Lsp_instFromJsonModuleRefs___spec__4(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__8; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_beqRefIdent____x40_Lean_Data_Lsp_Internal___hyg_28____boxed(lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); -static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____spec__1___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__15; static lean_object* l_Lean_Lsp_instToJsonRefInfo___closed__1; lean_object* lean_array_uget(lean_object*, size_t); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instFromJsonRefInfo___spec__2___boxed(lean_object*, lean_object*, lean_object*); @@ -129,34 +129,34 @@ static lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instFromJsonRefInfo___s uint64_t l___private_Lean_Expr_0__Lean_hashFVarId____x40_Lean_Expr___hyg_1933_(lean_object*); lean_object* l_List_redLength___rarg(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instFromJsonRefInfo___spec__4(size_t, size_t, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__9; lean_object* lean_string_append(lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__3; uint8_t lean_nat_dec_le(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_toJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1395____spec__1(size_t, size_t, size_t, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); lean_object* lean_nat_add(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__1; +static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__16; LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at_Lean_Lsp_instFromJsonModuleRefs___spec__8(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMap_insert___at_Lean_Lsp_instFromJsonModuleRefs___spec__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_Lsp_instFromJsonRefInfo___spec__1___closed__2; lean_object* l_Lean_Json_pretty(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_toJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1412_(lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_Lsp_instFromJsonModuleRefs___spec__5(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__6; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); lean_object* l_String_drop(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_contains___at_Lean_Lsp_instFromJsonModuleRefs___spec__2___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instToJsonRefInfo___spec__1(size_t, size_t, lean_object*); static lean_object* l_Lean_Lsp_instToJsonLeanIleanInfoParams___closed__1; lean_object* l_Nat_repr(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instFromJsonRefInfo___spec__5___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instToJsonRefInfo___spec__3(size_t, size_t, size_t, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__3; static lean_object* l_Lean_Lsp_RefIdent_toString___closed__2; +static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__4; lean_object* l_List_toArrayAux___rarg(lean_object*, lean_object*); lean_object* l___private_Lean_Data_HashMap_0__Lean_numBucketsForCapacity(lean_object*); lean_object* l_String_take(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__12; +static lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__7; LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_beqRefIdent____x40_Lean_Data_Lsp_Internal___hyg_28_(lean_object* x_1, lean_object* x_2) { _start: { @@ -418,14 +418,6 @@ static lean_object* _init_l_Lean_Lsp_RefIdent_fromString___closed__2() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("`", 1); -return x_1; -} -} -static lean_object* _init_l_Lean_Lsp_RefIdent_fromString___closed__3() { -_start: -{ -lean_object* x_1; x_1 = lean_mk_string_from_bytes("expected a Name, got ", 21); return x_1; } @@ -437,20 +429,29 @@ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_2 = lean_unsigned_to_nat(2u); lean_inc(x_1); x_3 = l_String_take(x_1, x_2); +lean_inc(x_1); x_4 = l_String_drop(x_1, x_2); x_5 = l_Lean_Lsp_RefIdent_fromString___closed__1; x_6 = lean_string_dec_eq(x_4, x_5); if (x_6 == 0) { -lean_object* x_7; lean_object* x_8; lean_object* x_9; -x_7 = l_Lean_Lsp_RefIdent_fromString___closed__2; -x_8 = lean_string_append(x_7, x_4); -x_9 = l_Lean_Syntax_decodeNameLit(x_8); -if (lean_obj_tag(x_9) == 0) +lean_object* x_7; uint8_t x_8; +x_7 = l_String_toName(x_1); +x_8 = l_Lean_Name_isAnonymous(x_7); +if (x_8 == 0) +{ +lean_object* x_9; +lean_dec(x_4); +x_9 = l_Lean_Lsp_RefIdent_fromString___lambda__1(x_3, x_7); +lean_dec(x_3); +return x_9; +} +else { lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +lean_dec(x_7); lean_dec(x_3); -x_10 = l_Lean_Lsp_RefIdent_fromString___closed__3; +x_10 = l_Lean_Lsp_RefIdent_fromString___closed__2; x_11 = lean_string_append(x_10, x_4); lean_dec(x_4); x_12 = l_Lean_Lsp_RefIdent_toString___closed__2; @@ -459,28 +460,18 @@ x_14 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_14, 0, x_13); return x_14; } +} else { lean_object* x_15; lean_object* x_16; lean_dec(x_4); -x_15 = lean_ctor_get(x_9, 0); -lean_inc(x_15); -lean_dec(x_9); +lean_dec(x_1); +x_15 = lean_box(0); x_16 = l_Lean_Lsp_RefIdent_fromString___lambda__1(x_3, x_15); lean_dec(x_3); return x_16; } } -else -{ -lean_object* x_17; lean_object* x_18; -lean_dec(x_4); -x_17 = lean_box(0); -x_18 = l_Lean_Lsp_RefIdent_fromString___lambda__1(x_3, x_17); -lean_dec(x_3); -return x_18; -} -} } LEAN_EXPORT lean_object* l_Lean_Lsp_RefIdent_fromString___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { _start: @@ -3614,7 +3605,7 @@ lean_dec(x_1); return x_2; } } -static lean_object* _init_l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____spec__1___closed__1() { +static lean_object* _init_l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____spec__1___closed__1() { _start: { lean_object* x_1; lean_object* x_2; @@ -3623,7 +3614,7 @@ x_2 = l_Lean_mkHashMapImp___rarg(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; @@ -3655,13 +3646,13 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; x_8 = lean_ctor_get(x_4, 0); lean_inc(x_8); lean_dec(x_4); -x_9 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____spec__1___closed__1; +x_9 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____spec__1___closed__1; x_10 = l_Lean_RBNode_foldM___at_Lean_Lsp_instFromJsonModuleRefs___spec__8(x_9, x_8); return x_10; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__1() { _start: { lean_object* x_1; @@ -3669,7 +3660,7 @@ x_1 = lean_mk_string_from_bytes("version", 7); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__2() { _start: { lean_object* x_1; @@ -3677,7 +3668,7 @@ x_1 = lean_mk_string_from_bytes("Lean", 4); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__3() { _start: { lean_object* x_1; @@ -3685,7 +3676,7 @@ x_1 = lean_mk_string_from_bytes("Lsp", 3); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__4() { _start: { lean_object* x_1; @@ -3693,28 +3684,28 @@ x_1 = lean_mk_string_from_bytes("LeanIleanInfoParams", 19); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__2; -x_2 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__3; -x_3 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__4; +x_1 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__2; +x_2 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__3; +x_3 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__4; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__6() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__5; +x_1 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__5; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__7() { _start: { lean_object* x_1; @@ -3722,47 +3713,47 @@ x_1 = lean_mk_string_from_bytes(".", 1); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__6; -x_2 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__7; +x_1 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__6; +x_2 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____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_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__1; +x_2 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__10() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__9; +x_1 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__9; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__8; -x_2 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__10; +x_1 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__8; +x_2 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__10; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__12() { _start: { lean_object* x_1; @@ -3770,17 +3761,17 @@ x_1 = lean_mk_string_from_bytes(": ", 2); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__11; -x_2 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__12; +x_1 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__11; +x_2 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__14() { +static lean_object* _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__14() { _start: { lean_object* x_1; @@ -3788,51 +3779,51 @@ x_1 = lean_mk_string_from_bytes("references", 10); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__15() { +static lean_object* _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____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_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__14; +x_2 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__14; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__16() { +static lean_object* _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__16() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__15; +x_1 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__15; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__17() { +static lean_object* _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__8; -x_2 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__16; +x_1 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__8; +x_2 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__16; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__18() { +static lean_object* _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__17; -x_2 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__12; +x_1 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__17; +x_2 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__1; +x_2 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__1; x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_548____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { @@ -3842,7 +3833,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__13; +x_6 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__13; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -3854,7 +3845,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__13; +x_9 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__13; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -3868,8 +3859,8 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__14; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____spec__1(x_1, x_13); +x_13 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__14; +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -3879,7 +3870,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__18; +x_17 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__18; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -3891,7 +3882,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__18; +x_20 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__18; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -3930,21 +3921,21 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____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_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____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_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____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_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227_(x_1); +x_2 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210_(x_1); lean_dec(x_1); return x_2; } @@ -3953,7 +3944,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonLeanIleanInfoParams___closed__1 _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____boxed), 1, 0); return x_1; } } @@ -3965,7 +3956,7 @@ x_1 = l_Lean_Lsp_instFromJsonLeanIleanInfoParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_toJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1412____spec__1(size_t 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_Internal_0__Lean_Lsp_toJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1395____spec__1(size_t x_1, size_t x_2, size_t x_3, lean_object* x_4) { _start: { uint8_t x_5; @@ -3999,7 +3990,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_toJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1412____spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_toJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1395____spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_2) == 0) @@ -4038,7 +4029,7 @@ x_17 = l_Array_mapMUnsafe_map___at_Lean_Lsp_instToJsonRefInfo___spec__1(x_15, x_ x_18 = lean_array_get_size(x_17); x_19 = lean_usize_of_nat(x_18); lean_dec(x_18); -x_20 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_toJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1412____spec__1(x_16, x_19, x_16, x_17); +x_20 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_toJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1395____spec__1(x_16, x_19, x_16, x_17); x_21 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_21, 0, x_20); x_22 = l_Lean_Lsp_instToJsonRefInfo___closed__1; @@ -4138,7 +4129,7 @@ x_61 = l_Array_mapMUnsafe_map___at_Lean_Lsp_instToJsonRefInfo___spec__1(x_59, x_ x_62 = lean_array_get_size(x_61); x_63 = lean_usize_of_nat(x_62); lean_dec(x_62); -x_64 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_toJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1412____spec__1(x_60, x_63, x_60, x_61); +x_64 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_toJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1395____spec__1(x_60, x_63, x_60, x_61); x_65 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_65, 0, x_64); x_66 = l_Lean_Lsp_instToJsonRefInfo___closed__1; @@ -4252,7 +4243,7 @@ x_108 = l_Array_mapMUnsafe_map___at_Lean_Lsp_instToJsonRefInfo___spec__1(x_106, x_109 = lean_array_get_size(x_108); x_110 = lean_usize_of_nat(x_109); lean_dec(x_109); -x_111 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_toJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1412____spec__1(x_107, x_110, x_107, x_108); +x_111 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_toJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1395____spec__1(x_107, x_110, x_107, x_108); x_112 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_112, 0, x_111); x_113 = l_Lean_Lsp_instToJsonRefInfo___closed__1; @@ -4339,7 +4330,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_toJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1412_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_toJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1395_(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; @@ -4348,7 +4339,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_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__1; +x_5 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____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); @@ -4360,9 +4351,9 @@ x_9 = lean_ctor_get(x_1, 1); lean_inc(x_9); lean_dec(x_1); x_10 = l_Lean_HashMap_toList___at_Lean_Lsp_instToJsonModuleRefs___spec__2(x_9); -x_11 = l_List_mapTR_loop___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_toJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1412____spec__2(x_7, x_10, x_7); +x_11 = l_List_mapTR_loop___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_toJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1395____spec__2(x_7, x_10, x_7); x_12 = l_Lean_Json_mkObj(x_11); -x_13 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__14; +x_13 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__14; x_14 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_14, 0, x_13); lean_ctor_set(x_14, 1, x_12); @@ -4380,7 +4371,7 @@ x_19 = l_Lean_Json_mkObj(x_18); return x_19; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_toJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1412____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_mapMUnsafe_map___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_toJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1395____spec__1___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; size_t x_7; lean_object* x_8; @@ -4390,7 +4381,7 @@ x_6 = lean_unbox_usize(x_2); lean_dec(x_2); x_7 = lean_unbox_usize(x_3); lean_dec(x_3); -x_8 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_toJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1412____spec__1(x_5, x_6, x_7, x_4); +x_8 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_toJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1395____spec__1(x_5, x_6, x_7, x_4); return x_8; } } @@ -4398,7 +4389,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonLeanIleanInfoParams___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_toJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1412_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_toJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1395_), 1, 0); return x_1; } } @@ -4453,8 +4444,6 @@ l_Lean_Lsp_RefIdent_fromString___closed__1 = _init_l_Lean_Lsp_RefIdent_fromStrin lean_mark_persistent(l_Lean_Lsp_RefIdent_fromString___closed__1); l_Lean_Lsp_RefIdent_fromString___closed__2 = _init_l_Lean_Lsp_RefIdent_fromString___closed__2(); lean_mark_persistent(l_Lean_Lsp_RefIdent_fromString___closed__2); -l_Lean_Lsp_RefIdent_fromString___closed__3 = _init_l_Lean_Lsp_RefIdent_fromString___closed__3(); -lean_mark_persistent(l_Lean_Lsp_RefIdent_fromString___closed__3); l_Lean_Lsp_instToJsonRefInfo___closed__1 = _init_l_Lean_Lsp_instToJsonRefInfo___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonRefInfo___closed__1); l_Lean_Lsp_instToJsonRefInfo___closed__2 = _init_l_Lean_Lsp_instToJsonRefInfo___closed__2(); @@ -4471,44 +4460,44 @@ l_Lean_Json_getObjValAs_x3f___at_Lean_Lsp_instFromJsonRefInfo___spec__1___closed lean_mark_persistent(l_Lean_Json_getObjValAs_x3f___at_Lean_Lsp_instFromJsonRefInfo___spec__1___closed__3); l_Array_mapMUnsafe_map___at_Lean_Lsp_instFromJsonRefInfo___spec__5___closed__1 = _init_l_Array_mapMUnsafe_map___at_Lean_Lsp_instFromJsonRefInfo___spec__5___closed__1(); lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Lsp_instFromJsonRefInfo___spec__5___closed__1); -l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____spec__1___closed__1 = _init_l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____spec__1___closed__1(); -lean_mark_persistent(l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____spec__1___closed__1); -l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__1 = _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__1); -l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__2 = _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__2); -l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__3 = _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__3); -l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__4 = _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__4); -l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__5 = _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__5); -l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__6 = _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__6); -l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__7 = _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__7); -l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__8 = _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__8); -l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__9 = _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__9); -l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__10 = _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__10); -l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__11 = _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__11); -l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__12 = _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__12); -l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__13 = _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__13); -l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__14 = _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__14(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__14); -l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__15 = _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__15(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__15); -l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__16 = _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__16(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__16); -l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__17 = _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__17(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__17); -l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__18 = _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__18(); -lean_mark_persistent(l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____closed__18); +l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____spec__1___closed__1 = _init_l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____spec__1___closed__1(); +lean_mark_persistent(l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____spec__1___closed__1); +l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__1 = _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__1); +l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__2 = _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__2); +l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__3 = _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__3); +l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__4 = _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__4); +l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__5 = _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__5); +l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__6 = _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__6); +l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__7 = _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__7); +l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__8 = _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__8); +l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__9 = _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__9); +l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__10 = _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__10); +l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__11 = _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__11); +l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__12 = _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__12); +l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__13 = _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__13); +l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__14 = _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__14(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__14); +l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__15 = _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__15(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__15); +l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__16 = _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__16(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__16); +l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__17 = _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__17(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__17); +l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__18 = _init_l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__18(); +lean_mark_persistent(l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____closed__18); l_Lean_Lsp_instFromJsonLeanIleanInfoParams___closed__1 = _init_l_Lean_Lsp_instFromJsonLeanIleanInfoParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonLeanIleanInfoParams___closed__1); l_Lean_Lsp_instFromJsonLeanIleanInfoParams = _init_l_Lean_Lsp_instFromJsonLeanIleanInfoParams(); diff --git a/stage0/stdlib/Lean/Data/Name.c b/stage0/stdlib/Lean/Data/Name.c index 610ba0fd65bd..46a1794211d7 100644 --- a/stage0/stdlib/Lean/Data/Name.c +++ b/stage0/stdlib/Lean/Data/Name.c @@ -24,7 +24,6 @@ LEAN_EXPORT uint8_t l_Lean_Name_isStr(lean_object*); LEAN_EXPORT lean_object* l_Lean_Name_lt___boxed(lean_object*, lean_object*); uint8_t lean_uint64_dec_lt(uint64_t, uint64_t); LEAN_EXPORT uint8_t l_Lean_Name_isAnonymous(lean_object*); -lean_object* l_String_trim(lean_object*); LEAN_EXPORT lean_object* l_Lean_Name_quickCmpAux___boxed(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Name_quickCmpAux(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Name_quickLt(lean_object*, lean_object*); @@ -45,11 +44,9 @@ LEAN_EXPORT lean_object* l_Lean_Name_getPrefix(lean_object*); LEAN_EXPORT lean_object* l_Lean_Name_components(lean_object*); lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); uint8_t l_String_startsWith(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_foldl___at_String_toName___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Name_isSuffixOf___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Name_componentsRev(lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_foldl___at_String_toName___spec__1___boxed(lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Name_getNumParts___boxed(lean_object*); lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -68,9 +65,7 @@ LEAN_EXPORT lean_object* l_Lean_Name_isImplementationDetail___boxed(lean_object* uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Name_eqStr___boxed(lean_object*, lean_object*); uint8_t lean_uint32_dec_eq(uint32_t, uint32_t); -static lean_object* l_String_toName___closed__1; uint64_t l_Lean_Name_hash___override(lean_object*); -LEAN_EXPORT lean_object* l_String_toName(lean_object*); static lean_object* l_Lean_Name_getString_x21___closed__2; lean_object* lean_panic_fn(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Name_getString_x21(lean_object*); @@ -86,7 +81,6 @@ lean_object* lean_nat_add(lean_object*, lean_object*); uint8_t lean_string_dec_lt(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Name_isInternal(lean_object*); LEAN_EXPORT uint8_t l_Lean_Name_lt(lean_object*, lean_object*); -lean_object* l_String_splitOn(lean_object*, lean_object*); lean_object* l_Lean_Name_mkSimple(lean_object*); static lean_object* _init_l_Lean_instCoeStringName___closed__1() { _start: @@ -1138,55 +1132,6 @@ x_4 = lean_box(x_3); return x_4; } } -LEAN_EXPORT lean_object* l_List_foldl___at_String_toName___spec__1(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; lean_object* x_6; -x_3 = lean_ctor_get(x_2, 0); -x_4 = lean_ctor_get(x_2, 1); -x_5 = l_String_trim(x_3); -x_6 = l_Lean_Name_str___override(x_1, x_5); -x_1 = x_6; -x_2 = x_4; -goto _start; -} -} -} -static lean_object* _init_l_String_toName___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes(".", 1); -return x_1; -} -} -LEAN_EXPORT lean_object* l_String_toName(lean_object* x_1) { -_start: -{ -lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_String_toName___closed__1; -x_3 = l_String_splitOn(x_1, x_2); -x_4 = lean_box(0); -x_5 = l_List_foldl___at_String_toName___spec__1(x_4, x_3); -lean_dec(x_3); -return x_5; -} -} -LEAN_EXPORT lean_object* l_List_foldl___at_String_toName___spec__1___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = l_List_foldl___at_String_toName___spec__1(x_1, x_2); -lean_dec(x_2); -return x_3; -} -} lean_object* initialize_Init(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Data_Name(uint8_t builtin, lean_object* w) { @@ -1210,8 +1155,6 @@ l_Lean_Name_getString_x21___closed__4 = _init_l_Lean_Name_getString_x21___closed lean_mark_persistent(l_Lean_Name_getString_x21___closed__4); l_Lean_Name_isImplementationDetail___closed__1 = _init_l_Lean_Name_isImplementationDetail___closed__1(); lean_mark_persistent(l_Lean_Name_isImplementationDetail___closed__1); -l_String_toName___closed__1 = _init_l_String_toName___closed__1(); -lean_mark_persistent(l_String_toName___closed__1); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lean/Data/Options.c b/stage0/stdlib/Lean/Data/Options.c index ec12aa47814f..e8f3ee6168e5 100644 --- a/stage0/stdlib/Lean/Data/Options.c +++ b/stage0/stdlib/Lean/Data/Options.c @@ -14,13 +14,11 @@ extern "C" { #endif static lean_object* l___auto____x40_Lean_Data_Options___hyg_103____closed__9; -LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_setOptionFromString___spec__1(lean_object*, lean_object*); lean_object* l_Lean_Syntax_mkNameLit(lean_object*, lean_object*); static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__registerBuiltinOption__1___closed__29; static lean_object* l_Lean_Option_registerBuiltinOption___closed__10; lean_object* lean_mk_empty_array_with_capacity(lean_object*); LEAN_EXPORT lean_object* l_Lean_getBoolOption___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_setOptionFromString___closed__1; static lean_object* l_Lean_Option_registerBuiltinOption___closed__14; static lean_object* l___auto____x40_Lean_Data_Options___hyg_103____closed__5; static lean_object* l___auto____x40_Lean_Data_Options___hyg_103____closed__19; @@ -28,8 +26,6 @@ static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Le LEAN_EXPORT uint8_t l_Lean_Options_getInPattern(lean_object*); LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Data_Options___hyg_220_(lean_object*); static lean_object* l_Lean_registerOption___closed__1; -static lean_object* l_Lean_setOptionFromString___closed__7; -lean_object* l_String_toNat_x3f(lean_object*); static lean_object* l___auto____x40_Lean_Data_Options___hyg_103____closed__26; lean_object* l_Lean_Name_toString(lean_object*, uint8_t); static lean_object* l_Lean_Option_registerBuiltinOption___closed__21; @@ -40,10 +36,8 @@ lean_object* lean_array_push(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__registerOption__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Option_registerBuiltinOption___closed__7; static lean_object* l___auto____x40_Lean_Data_Options___hyg_103____closed__21; -lean_object* l_String_trim(lean_object*); lean_object* l_Lean_KVMap_getNat(lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Data_Options___hyg_103____closed__28; -lean_object* l_Lean_KVMap_setNat(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Option_registerBuiltinOption___closed__1; static lean_object* l___auto____x40_Lean_Data_Options___hyg_103____closed__29; static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__registerBuiltinOption__1___closed__18; @@ -51,12 +45,9 @@ static lean_object* l___auto____x40_Lean_Data_Options___hyg_103____closed__1; static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__registerBuiltinOption__1___closed__37; static lean_object* l___auto____x40_Lean_Data_Options___hyg_103____closed__2; LEAN_EXPORT lean_object* l_Lean_Option_registerOption; -static lean_object* l_Lean_setOptionFromString___closed__9; static lean_object* l_Lean_Option_registerBuiltinOption___closed__11; static lean_object* l_Lean_Option_registerBuiltinOption___closed__4; uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___auto____x40_Lean_Data_Options___hyg_1592_; -static lean_object* l_Lean_setOptionFromString___closed__4; lean_object* l_Lean_KVMap_instToStringKVMap(lean_object*); static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__registerBuiltinOption__1___closed__38; lean_object* l_Lean_KVMap_insertCore(lean_object*, lean_object*, lean_object*); @@ -66,7 +57,6 @@ static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Le LEAN_EXPORT lean_object* l_Lean_Option_set(lean_object*); LEAN_EXPORT lean_object* l_Lean_Option_get___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instToStringOptions; -lean_object* l_String_toInt_x3f(lean_object*); 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__registerBuiltinOption__1___closed__8; lean_object* l_Lean_initializing(lean_object*); @@ -74,12 +64,10 @@ static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Le LEAN_EXPORT lean_object* l_Lean_Option_get_x3f___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Data_Options___hyg_103____closed__22; LEAN_EXPORT lean_object* l_Lean_Option_get_x3f___rarg___boxed(lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_KVMap_setName(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instForInOptionsProdNameDataValue(lean_object*); LEAN_EXPORT lean_object* l_Lean_Option_registerBuiltinOption; static lean_object* l___auto____x40_Lean_Data_Options___hyg_103____closed__27; -static lean_object* l_Lean_setOptionFromString___closed__8; LEAN_EXPORT lean_object* l_Lean_instMonadWithOptions(lean_object*, lean_object*); static lean_object* l_Lean_registerOption___lambda__1___closed__1; lean_object* l_Lean_SourceInfo_fromRef(lean_object*, uint8_t); @@ -105,7 +93,6 @@ static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Le lean_object* lean_st_ref_get(lean_object*, lean_object*); static lean_object* l_Lean_Option_registerBuiltinOption___closed__17; LEAN_EXPORT lean_object* l_Lean_getNatOption___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_setOptionFromString(lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Data_Options___hyg_103____closed__4; lean_object* l_Lean_Syntax_getOptional_x3f(lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); @@ -115,7 +102,6 @@ static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Le LEAN_EXPORT lean_object* l_Lean_Option_Decl_descr___default; static lean_object* l_Lean_Option_registerBuiltinOption___closed__26; lean_object* l_Lean_Syntax_node3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_setOptionFromString___closed__5; LEAN_EXPORT lean_object* l_Lean_Option_get(lean_object*); static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__registerBuiltinOption__1___closed__36; static lean_object* l_Lean_withInPattern___rarg___closed__1; @@ -126,7 +112,6 @@ static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Le static lean_object* l_Lean_Option_registerOption___closed__4; lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__registerBuiltinOption__1___closed__10; -uint8_t lean_name_eq(lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getBoolOption___rarg(lean_object*, lean_object*, lean_object*, uint8_t); static lean_object* l___auto____x40_Lean_Data_Options___hyg_103____closed__11; @@ -135,14 +120,12 @@ lean_object* l_Lean_Syntax_node2(lean_object*, lean_object*, lean_object*, lean_ static lean_object* l_Lean_Option_registerBuiltinOption___closed__24; lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instMonadWithOptions___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_KVMap_setString(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Option_registerBuiltinOption___closed__29; LEAN_EXPORT lean_object* l_Lean_instInhabitedOptionDecls; static lean_object* l_Lean_Option_registerOption___closed__5; static lean_object* l_Lean_registerOption___closed__2; LEAN_EXPORT lean_object* l_Lean_withInPattern(lean_object*, lean_object*); static lean_object* l_Lean_registerOption___lambda__2___closed__2; -static lean_object* l_Lean_setOptionFromString___closed__2; lean_object* l_Lean_KVMap_setBool(lean_object*, lean_object*, uint8_t); static lean_object* l_Lean_Option_registerBuiltinOption___closed__30; static lean_object* l___auto____x40_Lean_Data_Options___hyg_103____closed__18; @@ -166,12 +149,10 @@ LEAN_EXPORT lean_object* l_Lean_registerOption___lambda__1___boxed(lean_object*, static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__registerBuiltinOption__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Options_getInPattern___boxed(lean_object*); static lean_object* l_Lean_withInPattern___rarg___lambda__1___closed__2; -static lean_object* l_Lean_setOptionFromString___closed__12; static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__registerBuiltinOption__1___closed__9; LEAN_EXPORT lean_object* l_Lean_Option_register___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Data_Options___hyg_103____closed__13; static lean_object* l___auto____x40_Lean_Data_Options___hyg_103____closed__15; -static lean_object* l_Lean_setOptionFromString___closed__10; static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__registerBuiltinOption__1___closed__39; static lean_object* l_Lean_Option_registerBuiltinOption___closed__2; static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__registerBuiltinOption__1___closed__13; @@ -184,18 +165,15 @@ static lean_object* l___auto____x40_Lean_Data_Options___hyg_103____closed__12; LEAN_EXPORT lean_object* l_Lean_instInhabitedOptionDecl; LEAN_EXPORT lean_object* l_Lean_registerOption___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_instInhabitedOptionDecl___closed__1; -static lean_object* l_Lean_setOptionFromString___closed__11; LEAN_EXPORT lean_object* l_Lean_instInhabitedOptions; LEAN_EXPORT lean_object* l_Lean_getBoolOption___rarg___lambda__1(lean_object*, lean_object*, uint8_t, lean_object*); lean_object* l_Lean_Name_mkStr2(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__registerBuiltinOption__1___closed__5; LEAN_EXPORT lean_object* l_Lean_getNatOption___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_node1(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_KVMap_getBool(lean_object*, lean_object*, uint8_t); static lean_object* l_Lean_instForInOptionsProdNameDataValue___closed__1; LEAN_EXPORT lean_object* lean_get_option_decls_array(lean_object*); -lean_object* l_String_toName(lean_object*); LEAN_EXPORT lean_object* l_Lean_OptionDecl_group___default; LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_getOptionDecl___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getOptionDefaultValue(lean_object*, lean_object*); @@ -215,9 +193,9 @@ LEAN_EXPORT lean_object* l_Lean_getOptionDecl(lean_object*, lean_object*); static lean_object* l_Lean_Option_registerBuiltinOption___closed__25; static lean_object* l_Lean_Option_registerBuiltinOption___closed__27; LEAN_EXPORT lean_object* l_Lean_instInhabitedOption___rarg(lean_object*); -lean_object* l_List_reverse___rarg(lean_object*); static lean_object* l_Lean_Option_registerOption___closed__9; static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__registerBuiltinOption__1___closed__1; +LEAN_EXPORT lean_object* l___auto____x40_Lean_Data_Options___hyg_1232_; static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__registerBuiltinOption__1___closed__4; lean_object* l_String_intercalate(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getNatOption___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -262,14 +240,11 @@ static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Le static lean_object* l_Lean_Option_registerOption___closed__6; LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_getOptionDecl___spec__1___boxed(lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Data_Options___hyg_103____closed__23; -lean_object* l_String_splitOn(lean_object*, lean_object*); static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__registerBuiltinOption__1___closed__2; -static lean_object* l_Lean_setOptionFromString___closed__6; static lean_object* l_Lean_instToStringOptions___closed__1; LEAN_EXPORT lean_object* l_Lean_instMonadOptions___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Option_registerBuiltinOption___closed__13; static lean_object* l_Lean_Option_registerBuiltinOption___closed__12; -static lean_object* l_Lean_setOptionFromString___closed__3; static lean_object* _init_l_Lean_Options_empty() { _start: { @@ -1297,616 +1272,6 @@ return x_14; } } } -LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_setOptionFromString___spec__1(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 -{ -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 = l_String_trim(x_5); -lean_dec(x_5); -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 -{ -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 = l_String_trim(x_9); -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_setOptionFromString___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("=", 1); -return x_1; -} -} -static lean_object* _init_l_Lean_setOptionFromString___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("invalid configuration option entry, it must be of the form ' = '", 76); -return x_1; -} -} -static lean_object* _init_l_Lean_setOptionFromString___closed__3() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_setOptionFromString___closed__2; -x_2 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_setOptionFromString___closed__4() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("true", 4); -return x_1; -} -} -static lean_object* _init_l_Lean_setOptionFromString___closed__5() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_setOptionFromString___closed__4; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_setOptionFromString___closed__6() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("false", 5); -return x_1; -} -} -static lean_object* _init_l_Lean_setOptionFromString___closed__7() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_setOptionFromString___closed__6; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_setOptionFromString___closed__8() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("invalid Bool option value '", 27); -return x_1; -} -} -static lean_object* _init_l_Lean_setOptionFromString___closed__9() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("invalid Nat option value '", 26); -return x_1; -} -} -static lean_object* _init_l_Lean_setOptionFromString___closed__10() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("invalid Int option value '", 26); -return x_1; -} -} -static lean_object* _init_l_Lean_setOptionFromString___closed__11() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("invalid Syntax option value", 27); -return x_1; -} -} -static lean_object* _init_l_Lean_setOptionFromString___closed__12() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_setOptionFromString___closed__11; -x_2 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_setOptionFromString(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; -x_4 = l_Lean_setOptionFromString___closed__1; -x_5 = l_String_splitOn(x_2, x_4); -x_6 = lean_box(0); -x_7 = l_List_mapTR_loop___at_Lean_setOptionFromString___spec__1(x_5, x_6); -if (lean_obj_tag(x_7) == 0) -{ -lean_object* x_8; lean_object* x_9; -lean_dec(x_1); -x_8 = l_Lean_setOptionFromString___closed__3; -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; -x_10 = lean_ctor_get(x_7, 1); -lean_inc(x_10); -if (lean_obj_tag(x_10) == 0) -{ -lean_object* x_11; lean_object* x_12; -lean_dec(x_7); -lean_dec(x_1); -x_11 = l_Lean_setOptionFromString___closed__3; -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; -x_13 = lean_ctor_get(x_10, 1); -lean_inc(x_13); -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_7, 0); -lean_inc(x_14); -lean_dec(x_7); -x_15 = lean_ctor_get(x_10, 0); -lean_inc(x_15); -lean_dec(x_10); -x_16 = lean_box(0); -x_17 = l_Lean_Name_str___override(x_16, x_14); -lean_inc(x_17); -x_18 = l_Lean_getOptionDefaultValue(x_17, x_3); -if (lean_obj_tag(x_18) == 0) -{ -lean_object* x_19; -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -switch (lean_obj_tag(x_19)) { -case 0: -{ -uint8_t x_20; -lean_dec(x_19); -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 = l_Lean_KVMap_setString(x_1, x_17, x_15); -lean_ctor_set(x_18, 0, x_22); -return x_18; -} -else -{ -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 = l_Lean_KVMap_setString(x_1, x_17, x_15); -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; -} -} -case 1: -{ -uint8_t x_26; -lean_dec(x_19); -x_26 = !lean_is_exclusive(x_18); -if (x_26 == 0) -{ -lean_object* x_27; lean_object* x_28; uint8_t x_29; -x_27 = lean_ctor_get(x_18, 0); -lean_dec(x_27); -x_28 = l_Lean_setOptionFromString___closed__5; -x_29 = lean_name_eq(x_17, x_28); -if (x_29 == 0) -{ -lean_object* x_30; uint8_t x_31; -x_30 = l_Lean_setOptionFromString___closed__7; -x_31 = lean_name_eq(x_17, x_30); -if (x_31 == 0) -{ -lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; -lean_dec(x_17); -lean_dec(x_1); -x_32 = l_Lean_setOptionFromString___closed__8; -x_33 = lean_string_append(x_32, x_15); -lean_dec(x_15); -x_34 = l_Lean_getOptionDecl___closed__2; -x_35 = lean_string_append(x_33, x_34); -x_36 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_36, 0, x_35); -lean_ctor_set_tag(x_18, 1); -lean_ctor_set(x_18, 0, x_36); -return x_18; -} -else -{ -uint8_t x_37; lean_object* x_38; -lean_dec(x_15); -x_37 = 0; -x_38 = l_Lean_KVMap_setBool(x_1, x_17, x_37); -lean_ctor_set(x_18, 0, x_38); -return x_18; -} -} -else -{ -uint8_t x_39; lean_object* x_40; -lean_dec(x_15); -x_39 = 1; -x_40 = l_Lean_KVMap_setBool(x_1, x_17, x_39); -lean_ctor_set(x_18, 0, x_40); -return x_18; -} -} -else -{ -lean_object* x_41; lean_object* x_42; uint8_t x_43; -x_41 = lean_ctor_get(x_18, 1); -lean_inc(x_41); -lean_dec(x_18); -x_42 = l_Lean_setOptionFromString___closed__5; -x_43 = lean_name_eq(x_17, x_42); -if (x_43 == 0) -{ -lean_object* x_44; uint8_t x_45; -x_44 = l_Lean_setOptionFromString___closed__7; -x_45 = lean_name_eq(x_17, 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_dec(x_17); -lean_dec(x_1); -x_46 = l_Lean_setOptionFromString___closed__8; -x_47 = lean_string_append(x_46, x_15); -lean_dec(x_15); -x_48 = l_Lean_getOptionDecl___closed__2; -x_49 = lean_string_append(x_47, x_48); -x_50 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_50, 0, x_49); -x_51 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_51, 0, x_50); -lean_ctor_set(x_51, 1, x_41); -return x_51; -} -else -{ -uint8_t x_52; lean_object* x_53; lean_object* x_54; -lean_dec(x_15); -x_52 = 0; -x_53 = l_Lean_KVMap_setBool(x_1, x_17, 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_41); -return x_54; -} -} -else -{ -uint8_t x_55; lean_object* x_56; lean_object* x_57; -lean_dec(x_15); -x_55 = 1; -x_56 = l_Lean_KVMap_setBool(x_1, x_17, 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_41); -return x_57; -} -} -} -case 2: -{ -uint8_t x_58; -lean_dec(x_19); -x_58 = !lean_is_exclusive(x_18); -if (x_58 == 0) -{ -lean_object* x_59; lean_object* x_60; lean_object* x_61; -x_59 = lean_ctor_get(x_18, 0); -lean_dec(x_59); -x_60 = l_String_toName(x_15); -x_61 = l_Lean_KVMap_setName(x_1, x_17, x_60); -lean_ctor_set(x_18, 0, x_61); -return x_18; -} -else -{ -lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; -x_62 = lean_ctor_get(x_18, 1); -lean_inc(x_62); -lean_dec(x_18); -x_63 = l_String_toName(x_15); -x_64 = l_Lean_KVMap_setName(x_1, x_17, 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_62); -return x_65; -} -} -case 3: -{ -uint8_t x_66; -lean_dec(x_19); -x_66 = !lean_is_exclusive(x_18); -if (x_66 == 0) -{ -lean_object* x_67; lean_object* x_68; -x_67 = lean_ctor_get(x_18, 0); -lean_dec(x_67); -x_68 = l_String_toNat_x3f(x_15); -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_dec(x_17); -lean_dec(x_1); -x_69 = l_Lean_setOptionFromString___closed__9; -x_70 = lean_string_append(x_69, x_15); -lean_dec(x_15); -x_71 = l_Lean_getOptionDecl___closed__2; -x_72 = lean_string_append(x_70, x_71); -x_73 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_73, 0, x_72); -lean_ctor_set_tag(x_18, 1); -lean_ctor_set(x_18, 0, x_73); -return x_18; -} -else -{ -lean_object* x_74; lean_object* x_75; -lean_dec(x_15); -x_74 = lean_ctor_get(x_68, 0); -lean_inc(x_74); -lean_dec(x_68); -x_75 = l_Lean_KVMap_setNat(x_1, x_17, x_74); -lean_ctor_set(x_18, 0, x_75); -return x_18; -} -} -else -{ -lean_object* x_76; lean_object* x_77; -x_76 = lean_ctor_get(x_18, 1); -lean_inc(x_76); -lean_dec(x_18); -x_77 = l_String_toNat_x3f(x_15); -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; lean_object* x_83; -lean_dec(x_17); -lean_dec(x_1); -x_78 = l_Lean_setOptionFromString___closed__9; -x_79 = lean_string_append(x_78, x_15); -lean_dec(x_15); -x_80 = l_Lean_getOptionDecl___closed__2; -x_81 = lean_string_append(x_79, x_80); -x_82 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_82, 0, x_81); -x_83 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_83, 0, x_82); -lean_ctor_set(x_83, 1, x_76); -return x_83; -} -else -{ -lean_object* x_84; lean_object* x_85; lean_object* x_86; -lean_dec(x_15); -x_84 = lean_ctor_get(x_77, 0); -lean_inc(x_84); -lean_dec(x_77); -x_85 = l_Lean_KVMap_setNat(x_1, x_17, 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_76); -return x_86; -} -} -} -case 4: -{ -uint8_t x_87; -lean_dec(x_19); -x_87 = !lean_is_exclusive(x_18); -if (x_87 == 0) -{ -lean_object* x_88; lean_object* x_89; -x_88 = lean_ctor_get(x_18, 0); -lean_dec(x_88); -lean_inc(x_15); -x_89 = l_String_toInt_x3f(x_15); -if (lean_obj_tag(x_89) == 0) -{ -lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; -lean_dec(x_17); -lean_dec(x_1); -x_90 = l_Lean_setOptionFromString___closed__10; -x_91 = lean_string_append(x_90, x_15); -lean_dec(x_15); -x_92 = l_Lean_getOptionDecl___closed__2; -x_93 = lean_string_append(x_91, x_92); -x_94 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_94, 0, x_93); -lean_ctor_set_tag(x_18, 1); -lean_ctor_set(x_18, 0, x_94); -return x_18; -} -else -{ -lean_object* x_95; lean_object* x_96; -lean_dec(x_15); -x_95 = lean_ctor_get(x_89, 0); -lean_inc(x_95); -lean_dec(x_89); -x_96 = l_Lean_KVMap_setInt(x_1, x_17, x_95); -lean_ctor_set(x_18, 0, x_96); -return x_18; -} -} -else -{ -lean_object* x_97; lean_object* x_98; -x_97 = lean_ctor_get(x_18, 1); -lean_inc(x_97); -lean_dec(x_18); -lean_inc(x_15); -x_98 = l_String_toInt_x3f(x_15); -if (lean_obj_tag(x_98) == 0) -{ -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_17); -lean_dec(x_1); -x_99 = l_Lean_setOptionFromString___closed__10; -x_100 = lean_string_append(x_99, x_15); -lean_dec(x_15); -x_101 = l_Lean_getOptionDecl___closed__2; -x_102 = lean_string_append(x_100, x_101); -x_103 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_103, 0, x_102); -x_104 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_104, 0, x_103); -lean_ctor_set(x_104, 1, x_97); -return x_104; -} -else -{ -lean_object* x_105; lean_object* x_106; lean_object* x_107; -lean_dec(x_15); -x_105 = lean_ctor_get(x_98, 0); -lean_inc(x_105); -lean_dec(x_98); -x_106 = l_Lean_KVMap_setInt(x_1, x_17, 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_97); -return x_107; -} -} -} -default: -{ -uint8_t x_108; -lean_dec(x_19); -lean_dec(x_17); -lean_dec(x_15); -lean_dec(x_1); -x_108 = !lean_is_exclusive(x_18); -if (x_108 == 0) -{ -lean_object* x_109; lean_object* x_110; -x_109 = lean_ctor_get(x_18, 0); -lean_dec(x_109); -x_110 = l_Lean_setOptionFromString___closed__12; -lean_ctor_set_tag(x_18, 1); -lean_ctor_set(x_18, 0, x_110); -return x_18; -} -else -{ -lean_object* x_111; lean_object* x_112; lean_object* x_113; -x_111 = lean_ctor_get(x_18, 1); -lean_inc(x_111); -lean_dec(x_18); -x_112 = l_Lean_setOptionFromString___closed__12; -x_113 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_113, 0, x_112); -lean_ctor_set(x_113, 1, x_111); -return x_113; -} -} -} -} -else -{ -uint8_t x_114; -lean_dec(x_17); -lean_dec(x_15); -lean_dec(x_1); -x_114 = !lean_is_exclusive(x_18); -if (x_114 == 0) -{ -return x_18; -} -else -{ -lean_object* x_115; lean_object* x_116; lean_object* x_117; -x_115 = lean_ctor_get(x_18, 0); -x_116 = lean_ctor_get(x_18, 1); -lean_inc(x_116); -lean_inc(x_115); -lean_dec(x_18); -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_118; lean_object* x_119; -lean_dec(x_13); -lean_dec(x_10); -lean_dec(x_7); -lean_dec(x_1); -x_118 = l_Lean_setOptionFromString___closed__3; -x_119 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_119, 0, x_118); -lean_ctor_set(x_119, 1, x_3); -return x_119; -} -} -} -} -} LEAN_EXPORT lean_object* l_Lean_instMonadOptions___rarg(lean_object* x_1, lean_object* x_2) { _start: { @@ -2323,7 +1688,7 @@ x_2 = lean_alloc_closure((void*)(l_Lean_Option_setIfNotSet___rarg), 4, 0); return x_2; } } -static lean_object* _init_l___auto____x40_Lean_Data_Options___hyg_1592_() { +static lean_object* _init_l___auto____x40_Lean_Data_Options___hyg_1232_() { _start: { lean_object* x_1; @@ -3874,30 +3239,6 @@ l_Lean_getOptionDecl___closed__1 = _init_l_Lean_getOptionDecl___closed__1(); lean_mark_persistent(l_Lean_getOptionDecl___closed__1); l_Lean_getOptionDecl___closed__2 = _init_l_Lean_getOptionDecl___closed__2(); lean_mark_persistent(l_Lean_getOptionDecl___closed__2); -l_Lean_setOptionFromString___closed__1 = _init_l_Lean_setOptionFromString___closed__1(); -lean_mark_persistent(l_Lean_setOptionFromString___closed__1); -l_Lean_setOptionFromString___closed__2 = _init_l_Lean_setOptionFromString___closed__2(); -lean_mark_persistent(l_Lean_setOptionFromString___closed__2); -l_Lean_setOptionFromString___closed__3 = _init_l_Lean_setOptionFromString___closed__3(); -lean_mark_persistent(l_Lean_setOptionFromString___closed__3); -l_Lean_setOptionFromString___closed__4 = _init_l_Lean_setOptionFromString___closed__4(); -lean_mark_persistent(l_Lean_setOptionFromString___closed__4); -l_Lean_setOptionFromString___closed__5 = _init_l_Lean_setOptionFromString___closed__5(); -lean_mark_persistent(l_Lean_setOptionFromString___closed__5); -l_Lean_setOptionFromString___closed__6 = _init_l_Lean_setOptionFromString___closed__6(); -lean_mark_persistent(l_Lean_setOptionFromString___closed__6); -l_Lean_setOptionFromString___closed__7 = _init_l_Lean_setOptionFromString___closed__7(); -lean_mark_persistent(l_Lean_setOptionFromString___closed__7); -l_Lean_setOptionFromString___closed__8 = _init_l_Lean_setOptionFromString___closed__8(); -lean_mark_persistent(l_Lean_setOptionFromString___closed__8); -l_Lean_setOptionFromString___closed__9 = _init_l_Lean_setOptionFromString___closed__9(); -lean_mark_persistent(l_Lean_setOptionFromString___closed__9); -l_Lean_setOptionFromString___closed__10 = _init_l_Lean_setOptionFromString___closed__10(); -lean_mark_persistent(l_Lean_setOptionFromString___closed__10); -l_Lean_setOptionFromString___closed__11 = _init_l_Lean_setOptionFromString___closed__11(); -lean_mark_persistent(l_Lean_setOptionFromString___closed__11); -l_Lean_setOptionFromString___closed__12 = _init_l_Lean_setOptionFromString___closed__12(); -lean_mark_persistent(l_Lean_setOptionFromString___closed__12); l_Lean_withInPattern___rarg___lambda__1___closed__1 = _init_l_Lean_withInPattern___rarg___lambda__1___closed__1(); lean_mark_persistent(l_Lean_withInPattern___rarg___lambda__1___closed__1); l_Lean_withInPattern___rarg___lambda__1___closed__2 = _init_l_Lean_withInPattern___rarg___lambda__1___closed__2(); @@ -3908,8 +3249,8 @@ 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___auto____x40_Lean_Data_Options___hyg_1592_ = _init_l___auto____x40_Lean_Data_Options___hyg_1592_(); -lean_mark_persistent(l___auto____x40_Lean_Data_Options___hyg_1592_); +l___auto____x40_Lean_Data_Options___hyg_1232_ = _init_l___auto____x40_Lean_Data_Options___hyg_1232_(); +lean_mark_persistent(l___auto____x40_Lean_Data_Options___hyg_1232_); l_Lean_Option_registerBuiltinOption___closed__1 = _init_l_Lean_Option_registerBuiltinOption___closed__1(); lean_mark_persistent(l_Lean_Option_registerBuiltinOption___closed__1); l_Lean_Option_registerBuiltinOption___closed__2 = _init_l_Lean_Option_registerBuiltinOption___closed__2(); diff --git a/stage0/stdlib/Lean/Data/Parsec.c b/stage0/stdlib/Lean/Data/Parsec.c index 43b4883da7d6..e183b37ca7d1 100644 --- a/stage0/stdlib/Lean/Data/Parsec.c +++ b/stage0/stdlib/Lean/Data/Parsec.c @@ -58,12 +58,12 @@ static lean_object* l_Lean_Parsec_instMonadParsec___closed__5; LEAN_EXPORT lean_object* l_Lean_Parsec_instAlternativeParsec___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parsec_digit___closed__1; static lean_object* l_Lean_Parsec_unexpectedEndOfInput___closed__1; -uint8_t l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(lean_object*, lean_object*); static lean_object* l_Lean_Parsec_instAlternativeParsec___closed__3; LEAN_EXPORT lean_object* l_Lean_Parsec_hexDigit(lean_object*); uint8_t lean_uint32_dec_le(uint32_t, uint32_t); lean_object* lean_nat_to_int(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parsec_skip(lean_object*); +uint8_t l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(lean_object*, lean_object*); static lean_object* l_Lean_Parsec_run___rarg___closed__2; LEAN_EXPORT lean_object* l_Lean_Parsec_instReprParseResult___rarg(lean_object*); uint8_t l_String_Iterator_hasNext(lean_object*); @@ -1152,7 +1152,7 @@ if (x_9 == 0) lean_object* x_10; lean_object* x_11; uint8_t x_12; x_10 = lean_ctor_get(x_4, 0); x_11 = lean_ctor_get(x_4, 1); -x_12 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_10); +x_12 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_10); if (x_12 == 0) { lean_dec(x_3); @@ -1178,7 +1178,7 @@ x_16 = lean_ctor_get(x_4, 1); lean_inc(x_16); lean_inc(x_15); lean_dec(x_4); -x_17 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_15); +x_17 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_15); if (x_17 == 0) { lean_object* x_18; @@ -1323,7 +1323,7 @@ if (x_10 == 0) lean_object* x_11; lean_object* x_12; uint8_t x_13; x_11 = lean_ctor_get(x_5, 0); x_12 = lean_ctor_get(x_5, 1); -x_13 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_4, x_11); +x_13 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_4, x_11); if (x_13 == 0) { lean_dec(x_4); @@ -1349,7 +1349,7 @@ x_17 = lean_ctor_get(x_5, 1); lean_inc(x_17); lean_inc(x_16); lean_dec(x_5); -x_18 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_4, x_16); +x_18 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_4, x_16); if (x_18 == 0) { lean_object* x_19; @@ -1579,7 +1579,7 @@ if (x_13 == 0) lean_object* x_14; lean_object* x_15; uint8_t x_16; x_14 = lean_ctor_get(x_8, 0); x_15 = lean_ctor_get(x_8, 1); -x_16 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_14); +x_16 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_14); if (x_16 == 0) { lean_dec(x_3); @@ -1604,7 +1604,7 @@ x_18 = lean_ctor_get(x_8, 1); lean_inc(x_18); lean_inc(x_17); lean_dec(x_8); -x_19 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_17); +x_19 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_17); if (x_19 == 0) { lean_object* x_20; @@ -1638,7 +1638,7 @@ if (x_22 == 0) lean_object* x_23; lean_object* x_24; uint8_t x_25; x_23 = lean_ctor_get(x_4, 0); x_24 = lean_ctor_get(x_4, 1); -x_25 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_23); +x_25 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_23); if (x_25 == 0) { lean_dec(x_3); @@ -1663,7 +1663,7 @@ x_27 = lean_ctor_get(x_4, 1); lean_inc(x_27); lean_inc(x_26); lean_dec(x_4); -x_28 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_26); +x_28 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_26); if (x_28 == 0) { lean_object* x_29; @@ -1836,7 +1836,7 @@ if (x_14 == 0) lean_object* x_15; lean_object* x_16; uint8_t x_17; x_15 = lean_ctor_get(x_9, 0); x_16 = lean_ctor_get(x_9, 1); -x_17 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_15); +x_17 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_15); if (x_17 == 0) { lean_dec(x_3); @@ -1861,7 +1861,7 @@ x_19 = lean_ctor_get(x_9, 1); lean_inc(x_19); lean_inc(x_18); lean_dec(x_9); -x_20 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_18); +x_20 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_18); if (x_20 == 0) { lean_object* x_21; @@ -1895,7 +1895,7 @@ if (x_23 == 0) lean_object* x_24; lean_object* x_25; uint8_t x_26; x_24 = lean_ctor_get(x_4, 0); x_25 = lean_ctor_get(x_4, 1); -x_26 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_24); +x_26 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_24); if (x_26 == 0) { lean_dec(x_3); @@ -1920,7 +1920,7 @@ x_28 = lean_ctor_get(x_4, 1); lean_inc(x_28); lean_inc(x_27); lean_dec(x_4); -x_29 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_27); +x_29 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_27); if (x_29 == 0) { lean_object* x_30; diff --git a/stage0/stdlib/Lean/Data/Xml/Parser.c b/stage0/stdlib/Lean/Data/Xml/Parser.c index 2fcb83a1d454..d90d26f54e8c 100644 --- a/stage0/stdlib/Lean/Data/Xml/Parser.c +++ b/stage0/stdlib/Lean/Data/Xml/Parser.c @@ -133,7 +133,6 @@ LEAN_EXPORT lean_object* l_Lean_Parsec_manyCore___at_Lean_Xml_Parser_VersionNum_ LEAN_EXPORT lean_object* l_Lean_Xml_Parser_EntityValue(lean_object*); LEAN_EXPORT lean_object* l_Lean_Xml_Parser_SDDecl___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Xml_Parser_NameStartChar(lean_object*); -uint8_t l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(lean_object*, lean_object*); static lean_object* l_Lean_Xml_Parser_endl___closed__5; static lean_object* l_Lean_Xml_Parser_SDDecl___closed__1; LEAN_EXPORT lean_object* l_Array_filterMapM___at_Lean_Xml_Parser_content___spec__2___boxed(lean_object*, lean_object*, lean_object*); @@ -152,6 +151,7 @@ lean_object* l_Lean_RBNode_setBlack___rarg(lean_object*); static lean_object* l_Lean_Xml_Parser_endl___closed__7; static lean_object* l___private_Lean_Data_Xml_Parser_0__Lean_Xml_Parser_nameStartCharRanges___closed__20; static lean_object* l_Lean_Xml_Parser_Mixed___closed__2; +uint8_t l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(lean_object*, lean_object*); static lean_object* l_Lean_Xml_Parser_doctypedecl___closed__6; static lean_object* l_Lean_Parsec_manyCore___at_Lean_Xml_Parser_Mixed___spec__1___closed__3; static lean_object* l_Lean_Xml_Parser_VersionInfo___closed__1; @@ -543,7 +543,7 @@ if (x_26 == 0) lean_object* x_27; lean_object* x_28; uint8_t x_29; x_27 = lean_ctor_get(x_19, 0); x_28 = lean_ctor_get(x_19, 1); -x_29 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_27); +x_29 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_27); if (x_29 == 0) { lean_dec(x_1); @@ -605,7 +605,7 @@ x_39 = lean_ctor_get(x_19, 1); lean_inc(x_39); lean_inc(x_38); lean_dec(x_19); -x_40 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_38); +x_40 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_38); if (x_40 == 0) { lean_object* x_41; @@ -665,7 +665,7 @@ return x_50; block_17: { uint8_t x_4; -x_4 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_2); +x_4 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_2); if (x_4 == 0) { lean_object* x_5; @@ -934,7 +934,7 @@ goto block_42; block_42: { uint8_t x_6; -x_6 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_4); +x_6 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_4); if (x_6 == 0) { lean_object* x_7; @@ -1331,7 +1331,7 @@ return x_49; block_21: { uint8_t x_4; -x_4 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_2); +x_4 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_2); if (x_4 == 0) { lean_object* x_5; @@ -1351,7 +1351,7 @@ x_7 = l_String_Iterator_hasNext(x_1); if (x_7 == 0) { uint8_t x_8; -x_8 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_1); +x_8 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_1); if (x_8 == 0) { lean_object* x_9; lean_object* x_10; @@ -1379,7 +1379,7 @@ if (x_14 == 0) { uint8_t x_15; lean_dec(x_12); -x_15 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_1); +x_15 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_1); if (x_15 == 0) { lean_object* x_16; lean_object* x_17; @@ -1494,7 +1494,7 @@ if (x_12 == 0) lean_object* x_13; lean_object* x_14; uint8_t x_15; x_13 = lean_ctor_get(x_7, 0); x_14 = lean_ctor_get(x_7, 1); -x_15 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_13); +x_15 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_13); if (x_15 == 0) { lean_dec(x_3); @@ -1519,7 +1519,7 @@ x_17 = lean_ctor_get(x_7, 1); lean_inc(x_17); lean_inc(x_16); lean_dec(x_7); -x_18 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_16); +x_18 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_16); if (x_18 == 0) { lean_object* x_19; @@ -1546,7 +1546,7 @@ return x_20; block_27: { uint8_t x_24; -x_24 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_22); +x_24 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_22); if (x_24 == 0) { lean_object* x_25; @@ -1571,7 +1571,7 @@ return x_26; block_45: { uint8_t x_30; -x_30 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_28); +x_30 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_28); if (x_30 == 0) { x_22 = x_28; @@ -1607,7 +1607,7 @@ lean_inc(x_35); x_36 = lean_ctor_get(x_31, 1); lean_inc(x_36); lean_dec(x_31); -x_37 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_35); +x_37 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_35); if (x_37 == 0) { x_22 = x_35; @@ -1743,7 +1743,7 @@ if (x_12 == 0) lean_object* x_13; lean_object* x_14; uint8_t x_15; x_13 = lean_ctor_get(x_7, 0); x_14 = lean_ctor_get(x_7, 1); -x_15 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_13); +x_15 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_13); if (x_15 == 0) { lean_dec(x_3); @@ -1768,7 +1768,7 @@ x_17 = lean_ctor_get(x_7, 1); lean_inc(x_17); lean_inc(x_16); lean_dec(x_7); -x_18 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_16); +x_18 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_16); if (x_18 == 0) { lean_object* x_19; @@ -1795,7 +1795,7 @@ return x_20; block_27: { uint8_t x_24; -x_24 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_22); +x_24 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_22); if (x_24 == 0) { lean_object* x_25; @@ -1820,7 +1820,7 @@ return x_26; block_45: { uint8_t x_30; -x_30 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_28); +x_30 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_28); if (x_30 == 0) { x_22 = x_28; @@ -1856,7 +1856,7 @@ lean_inc(x_35); x_36 = lean_ctor_get(x_31, 1); lean_inc(x_36); lean_dec(x_31); -x_37 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_35); +x_37 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_35); if (x_37 == 0) { x_22 = x_35; @@ -1992,7 +1992,7 @@ if (x_12 == 0) lean_object* x_13; lean_object* x_14; uint8_t x_15; x_13 = lean_ctor_get(x_7, 0); x_14 = lean_ctor_get(x_7, 1); -x_15 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_13); +x_15 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_13); if (x_15 == 0) { lean_dec(x_3); @@ -2017,7 +2017,7 @@ x_17 = lean_ctor_get(x_7, 1); lean_inc(x_17); lean_inc(x_16); lean_dec(x_7); -x_18 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_16); +x_18 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_16); if (x_18 == 0) { lean_object* x_19; @@ -2044,7 +2044,7 @@ return x_20; block_27: { uint8_t x_24; -x_24 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_22); +x_24 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_22); if (x_24 == 0) { lean_object* x_25; @@ -2069,7 +2069,7 @@ return x_26; block_45: { uint8_t x_30; -x_30 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_28); +x_30 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_28); if (x_30 == 0) { x_22 = x_28; @@ -2105,7 +2105,7 @@ lean_inc(x_35); x_36 = lean_ctor_get(x_31, 1); lean_inc(x_36); lean_dec(x_31); -x_37 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_35); +x_37 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_35); if (x_37 == 0) { x_22 = x_35; @@ -2232,7 +2232,7 @@ return x_50; block_42: { uint8_t x_5; -x_5 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_3); +x_5 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_3); if (x_5 == 0) { lean_object* x_6; @@ -2274,7 +2274,7 @@ if (x_14 == 0) lean_object* x_15; lean_object* x_16; uint8_t x_17; x_15 = lean_ctor_get(x_7, 0); x_16 = lean_ctor_get(x_7, 1); -x_17 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_15); +x_17 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_15); if (x_17 == 0) { lean_dec(x_1); @@ -2331,7 +2331,7 @@ x_28 = lean_ctor_get(x_7, 1); lean_inc(x_28); lean_inc(x_27); lean_dec(x_7); -x_29 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_27); +x_29 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_27); if (x_29 == 0) { lean_object* x_30; @@ -2475,7 +2475,7 @@ if (x_33 == 0) lean_object* x_34; lean_object* x_35; uint8_t x_36; x_34 = lean_ctor_get(x_31, 0); x_35 = lean_ctor_get(x_31, 1); -x_36 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_34); +x_36 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_34); if (x_36 == 0) { lean_dec(x_1); @@ -2498,7 +2498,7 @@ x_38 = lean_ctor_get(x_31, 1); lean_inc(x_38); lean_inc(x_37); lean_dec(x_31); -x_39 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_37); +x_39 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_37); if (x_39 == 0) { lean_object* x_40; @@ -2590,7 +2590,7 @@ if (x_19 == 0) lean_object* x_20; lean_object* x_21; uint8_t x_22; x_20 = lean_ctor_get(x_12, 0); x_21 = lean_ctor_get(x_12, 1); -x_22 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_7, x_20); +x_22 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_7, x_20); if (x_22 == 0) { lean_dec(x_7); @@ -2616,7 +2616,7 @@ x_25 = lean_ctor_get(x_12, 1); lean_inc(x_25); lean_inc(x_24); lean_dec(x_12); -x_26 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_7, x_24); +x_26 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_7, x_24); if (x_26 == 0) { lean_object* x_27; @@ -3422,7 +3422,7 @@ if (lean_is_exclusive(x_2)) { lean_dec_ref(x_2); x_9 = lean_box(0); } -x_10 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_7); +x_10 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_7); if (x_10 == 0) { lean_object* x_11; @@ -3500,7 +3500,7 @@ return x_101; block_89: { uint8_t x_14; -x_14 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_12); +x_14 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_12); if (x_14 == 0) { lean_object* x_15; @@ -3562,7 +3562,7 @@ return x_88; block_80: { uint8_t x_19; -x_19 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_17); +x_19 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_17); if (x_19 == 0) { lean_object* x_20; @@ -3624,7 +3624,7 @@ return x_79; block_71: { uint8_t x_24; -x_24 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_22); +x_24 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_22); if (x_24 == 0) { lean_object* x_25; @@ -3686,7 +3686,7 @@ return x_70; block_62: { uint8_t x_29; -x_29 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_27); +x_29 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_27); if (x_29 == 0) { lean_object* x_30; @@ -3923,7 +3923,7 @@ if (x_13 == 0) lean_object* x_14; lean_object* x_15; uint8_t x_16; x_14 = lean_ctor_get(x_8, 0); x_15 = lean_ctor_get(x_8, 1); -x_16 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_14); +x_16 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_14); if (x_16 == 0) { lean_dec(x_2); @@ -3948,7 +3948,7 @@ x_18 = lean_ctor_get(x_8, 1); lean_inc(x_18); lean_inc(x_17); lean_dec(x_8); -x_19 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_17); +x_19 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_17); if (x_19 == 0) { lean_object* x_20; @@ -3981,7 +3981,7 @@ if (x_22 == 0) lean_object* x_23; lean_object* x_24; uint8_t x_25; x_23 = lean_ctor_get(x_3, 0); x_24 = lean_ctor_get(x_3, 1); -x_25 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_23); +x_25 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_23); if (x_25 == 0) { lean_dec(x_2); @@ -4006,7 +4006,7 @@ x_27 = lean_ctor_get(x_3, 1); lean_inc(x_27); lean_inc(x_26); lean_dec(x_3); -x_28 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_26); +x_28 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_26); if (x_28 == 0) { lean_object* x_29; @@ -4083,7 +4083,7 @@ x_3 = l_String_Iterator_hasNext(x_2); if (x_3 == 0) { uint8_t x_4; -x_4 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_2); +x_4 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_2); if (x_4 == 0) { lean_object* x_5; lean_object* x_6; @@ -4115,7 +4115,7 @@ if (x_11 == 0) { uint8_t x_12; lean_dec(x_8); -x_12 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_2); +x_12 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_2); if (x_12 == 0) { lean_object* x_13; lean_object* x_14; @@ -4144,7 +4144,7 @@ if (x_17 == 0) { uint8_t x_18; lean_dec(x_8); -x_18 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_2); +x_18 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_2); if (x_18 == 0) { lean_object* x_19; lean_object* x_20; @@ -4204,7 +4204,7 @@ if (x_29 == 0) lean_object* x_30; lean_object* x_31; uint8_t x_32; x_30 = lean_ctor_get(x_24, 0); x_31 = lean_ctor_get(x_24, 1); -x_32 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_30); +x_32 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_30); if (x_32 == 0) { lean_dec(x_2); @@ -4229,7 +4229,7 @@ x_34 = lean_ctor_get(x_24, 1); lean_inc(x_34); lean_inc(x_33); lean_dec(x_24); -x_35 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_33); +x_35 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_33); if (x_35 == 0) { lean_object* x_36; @@ -4828,7 +4828,7 @@ if (x_11 == 0) lean_object* x_12; lean_object* x_13; uint8_t x_14; x_12 = lean_ctor_get(x_6, 0); x_13 = lean_ctor_get(x_6, 1); -x_14 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_12); +x_14 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_12); if (x_14 == 0) { lean_dec(x_2); @@ -4853,7 +4853,7 @@ x_16 = lean_ctor_get(x_6, 1); lean_inc(x_16); lean_inc(x_15); lean_dec(x_6); -x_17 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_15); +x_17 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_15); if (x_17 == 0) { lean_object* x_18; @@ -4880,7 +4880,7 @@ return x_19; block_26: { uint8_t x_23; -x_23 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_21); +x_23 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_21); if (x_23 == 0) { lean_object* x_24; @@ -4905,7 +4905,7 @@ return x_25; block_37: { uint8_t x_29; -x_29 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_27); +x_29 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_27); if (x_29 == 0) { x_21 = x_27; @@ -4957,7 +4957,7 @@ goto block_20; block_48: { uint8_t x_40; -x_40 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_38); +x_40 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_38); if (x_40 == 0) { x_21 = x_38; @@ -5009,7 +5009,7 @@ goto block_20; block_59: { uint8_t x_51; -x_51 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_49); +x_51 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_49); if (x_51 == 0) { x_21 = x_49; @@ -5061,7 +5061,7 @@ goto block_20; block_73: { uint8_t x_62; -x_62 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_60); +x_62 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_60); if (x_62 == 0) { x_21 = x_60; @@ -5421,7 +5421,7 @@ if (x_8 == 0) lean_object* x_9; lean_object* x_10; uint8_t x_11; x_9 = lean_ctor_get(x_3, 0); x_10 = lean_ctor_get(x_3, 1); -x_11 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_9); +x_11 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_9); if (x_11 == 0) { lean_dec(x_1); @@ -5446,7 +5446,7 @@ x_15 = lean_ctor_get(x_3, 1); lean_inc(x_15); lean_inc(x_14); lean_dec(x_3); -x_16 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_14); +x_16 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_14); if (x_16 == 0) { lean_object* x_17; @@ -5650,7 +5650,7 @@ if (x_57 == 0) lean_object* x_58; lean_object* x_59; uint8_t x_60; x_58 = lean_ctor_get(x_55, 0); x_59 = lean_ctor_get(x_55, 1); -x_60 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_54, x_58); +x_60 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_54, x_58); if (x_60 == 0) { lean_dec(x_54); @@ -5673,7 +5673,7 @@ x_62 = lean_ctor_get(x_55, 1); lean_inc(x_62); lean_inc(x_61); lean_dec(x_55); -x_63 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_54, x_61); +x_63 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_54, x_61); if (x_63 == 0) { lean_object* x_64; @@ -5702,7 +5702,7 @@ if (x_65 == 0) lean_object* x_66; lean_object* x_67; uint8_t x_68; x_66 = lean_ctor_get(x_53, 0); x_67 = lean_ctor_get(x_53, 1); -x_68 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_52, x_66); +x_68 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_52, x_66); if (x_68 == 0) { lean_dec(x_52); @@ -5735,7 +5735,7 @@ if (x_71 == 0) lean_object* x_72; lean_object* x_73; uint8_t x_74; x_72 = lean_ctor_get(x_69, 0); x_73 = lean_ctor_get(x_69, 1); -x_74 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_52, x_72); +x_74 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_52, x_72); if (x_74 == 0) { lean_dec(x_52); @@ -5758,7 +5758,7 @@ x_76 = lean_ctor_get(x_69, 1); lean_inc(x_76); lean_inc(x_75); lean_dec(x_69); -x_77 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_52, x_75); +x_77 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_52, x_75); if (x_77 == 0) { lean_object* x_78; @@ -5787,7 +5787,7 @@ x_80 = lean_ctor_get(x_53, 1); lean_inc(x_80); lean_inc(x_79); lean_dec(x_53); -x_81 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_52, x_79); +x_81 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_52, x_79); if (x_81 == 0) { lean_object* x_82; @@ -5829,7 +5829,7 @@ if (lean_is_exclusive(x_83)) { lean_dec_ref(x_83); x_87 = lean_box(0); } -x_88 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_52, x_85); +x_88 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_52, x_85); if (x_88 == 0) { lean_object* x_89; @@ -5973,7 +5973,7 @@ if (x_17 == 0) lean_object* x_18; lean_object* x_19; uint8_t x_20; x_18 = lean_ctor_get(x_3, 0); x_19 = lean_ctor_get(x_3, 1); -x_20 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_18); +x_20 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_18); if (x_20 == 0) { lean_dec(x_2); @@ -6045,7 +6045,7 @@ x_34 = lean_ctor_get(x_3, 1); lean_inc(x_34); lean_inc(x_33); lean_dec(x_3); -x_35 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_33); +x_35 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_33); if (x_35 == 0) { lean_object* x_36; @@ -6197,7 +6197,7 @@ if (x_11 == 0) lean_object* x_12; lean_object* x_13; uint8_t x_14; x_12 = lean_ctor_get(x_6, 0); x_13 = lean_ctor_get(x_6, 1); -x_14 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_12); +x_14 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_12); if (x_14 == 0) { lean_dec(x_2); @@ -6222,7 +6222,7 @@ x_16 = lean_ctor_get(x_6, 1); lean_inc(x_16); lean_inc(x_15); lean_dec(x_6); -x_17 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_15); +x_17 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_15); if (x_17 == 0) { lean_object* x_18; @@ -6249,7 +6249,7 @@ return x_19; block_26: { uint8_t x_23; -x_23 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_21); +x_23 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_21); if (x_23 == 0) { lean_object* x_24; @@ -6274,7 +6274,7 @@ return x_25; block_48: { uint8_t x_29; -x_29 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_27); +x_29 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_27); if (x_29 == 0) { x_21 = x_27; @@ -6913,7 +6913,7 @@ goto block_51; block_69: { uint8_t x_58; -x_58 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_53, x_56); +x_58 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_53, x_56); if (x_58 == 0) { lean_object* x_59; @@ -7052,7 +7052,7 @@ return x_26; block_18: { uint8_t x_6; -x_6 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_4); +x_6 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_4); if (x_6 == 0) { lean_object* x_7; @@ -7153,7 +7153,7 @@ goto block_27; block_43: { uint8_t x_32; -x_32 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_28, x_30); +x_32 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_28, x_30); if (x_32 == 0) { lean_object* x_33; @@ -7230,7 +7230,7 @@ x_6 = lean_ctor_get(x_4, 1); lean_dec(x_6); x_7 = lean_ctor_get(x_4, 0); lean_dec(x_7); -x_8 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_2); +x_8 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_2); if (x_8 == 0) { lean_object* x_9; @@ -7252,7 +7252,7 @@ else { uint8_t x_10; lean_dec(x_4); -x_10 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_2); +x_10 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_2); if (x_10 == 0) { lean_object* x_11; lean_object* x_12; @@ -7325,7 +7325,7 @@ if (x_24 == 0) lean_object* x_25; lean_object* x_26; uint8_t x_27; x_25 = lean_ctor_get(x_19, 0); x_26 = lean_ctor_get(x_19, 1); -x_27 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_25); +x_27 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_25); if (x_27 == 0) { lean_dec(x_2); @@ -7350,7 +7350,7 @@ x_29 = lean_ctor_get(x_19, 1); lean_inc(x_29); lean_inc(x_28); lean_dec(x_19); -x_30 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_28); +x_30 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_28); if (x_30 == 0) { lean_object* x_31; @@ -7383,7 +7383,7 @@ if (x_33 == 0) lean_object* x_34; lean_object* x_35; uint8_t x_36; x_34 = lean_ctor_get(x_14, 0); x_35 = lean_ctor_get(x_14, 1); -x_36 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_34); +x_36 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_34); if (x_36 == 0) { lean_dec(x_2); @@ -7408,7 +7408,7 @@ x_38 = lean_ctor_get(x_14, 1); lean_inc(x_38); lean_inc(x_37); lean_dec(x_14); -x_39 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_37); +x_39 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_37); if (x_39 == 0) { lean_object* x_40; @@ -7538,7 +7538,7 @@ if (x_24 == 0) lean_object* x_25; lean_object* x_26; uint8_t x_27; x_25 = lean_ctor_get(x_10, 0); x_26 = lean_ctor_get(x_10, 1); -x_27 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_6, x_25); +x_27 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_6, x_25); if (x_27 == 0) { lean_dec(x_6); @@ -7610,7 +7610,7 @@ x_41 = lean_ctor_get(x_10, 1); lean_inc(x_41); lean_inc(x_40); lean_dec(x_10); -x_42 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_6, x_40); +x_42 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_6, x_40); if (x_42 == 0) { lean_object* x_43; @@ -7687,7 +7687,7 @@ if (x_54 == 0) lean_object* x_55; lean_object* x_56; uint8_t x_57; x_55 = lean_ctor_get(x_7, 0); x_56 = lean_ctor_get(x_7, 1); -x_57 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_6, x_55); +x_57 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_6, x_55); if (x_57 == 0) { lean_dec(x_6); @@ -7759,7 +7759,7 @@ x_71 = lean_ctor_get(x_7, 1); lean_inc(x_71); lean_inc(x_70); lean_dec(x_7); -x_72 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_6, x_70); +x_72 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_6, x_70); if (x_72 == 0) { lean_object* x_73; @@ -7916,7 +7916,7 @@ if (x_9 == 0) lean_object* x_10; lean_object* x_11; uint8_t x_12; x_10 = lean_ctor_get(x_2, 0); x_11 = lean_ctor_get(x_2, 1); -x_12 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_10); +x_12 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_10); if (x_12 == 0) { lean_dec(x_1); @@ -7962,7 +7962,7 @@ if (x_18 == 0) lean_object* x_19; lean_object* x_20; uint8_t x_21; x_19 = lean_ctor_get(x_13, 0); x_20 = lean_ctor_get(x_13, 1); -x_21 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_19); +x_21 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_19); if (x_21 == 0) { lean_dec(x_1); @@ -8033,7 +8033,7 @@ x_34 = lean_ctor_get(x_13, 1); lean_inc(x_34); lean_inc(x_33); lean_dec(x_13); -x_35 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_33); +x_35 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_33); if (x_35 == 0) { lean_object* x_36; @@ -8109,7 +8109,7 @@ x_47 = lean_ctor_get(x_2, 1); lean_inc(x_47); lean_inc(x_46); lean_dec(x_2); -x_48 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_46); +x_48 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_46); if (x_48 == 0) { lean_object* x_49; @@ -8166,7 +8166,7 @@ if (lean_is_exclusive(x_50)) { lean_dec_ref(x_50); x_57 = lean_box(0); } -x_58 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_55); +x_58 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_55); if (x_58 == 0) { lean_object* x_59; @@ -8249,7 +8249,7 @@ x_4 = l_String_Iterator_hasNext(x_3); if (x_4 == 0) { uint8_t x_5; -x_5 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_3); +x_5 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_3); if (x_5 == 0) { lean_object* x_6; lean_object* x_7; @@ -8281,7 +8281,7 @@ if (x_12 == 0) { uint8_t x_13; lean_dec(x_9); -x_13 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_3); +x_13 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_3); if (x_13 == 0) { lean_object* x_14; lean_object* x_15; @@ -8340,7 +8340,7 @@ if (x_23 == 0) lean_object* x_24; lean_object* x_25; uint8_t x_26; x_24 = lean_ctor_get(x_18, 0); x_25 = lean_ctor_get(x_18, 1); -x_26 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_24); +x_26 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_24); if (x_26 == 0) { lean_dec(x_3); @@ -8365,7 +8365,7 @@ x_28 = lean_ctor_get(x_18, 1); lean_inc(x_28); lean_inc(x_27); lean_dec(x_18); -x_29 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_27); +x_29 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_27); if (x_29 == 0) { lean_object* x_30; @@ -8400,7 +8400,7 @@ x_4 = l_String_Iterator_hasNext(x_3); if (x_4 == 0) { uint8_t x_5; -x_5 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_3); +x_5 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_3); if (x_5 == 0) { lean_object* x_6; lean_object* x_7; @@ -8432,7 +8432,7 @@ if (x_12 == 0) { uint8_t x_13; lean_dec(x_9); -x_13 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_3); +x_13 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_3); if (x_13 == 0) { lean_object* x_14; lean_object* x_15; @@ -8491,7 +8491,7 @@ if (x_23 == 0) lean_object* x_24; lean_object* x_25; uint8_t x_26; x_24 = lean_ctor_get(x_18, 0); x_25 = lean_ctor_get(x_18, 1); -x_26 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_24); +x_26 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_24); if (x_26 == 0) { lean_dec(x_3); @@ -8516,7 +8516,7 @@ x_28 = lean_ctor_get(x_18, 1); lean_inc(x_28); lean_inc(x_27); lean_dec(x_18); -x_29 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_27); +x_29 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_27); if (x_29 == 0) { lean_object* x_30; @@ -8692,7 +8692,7 @@ goto block_26; block_26: { uint8_t x_5; -x_5 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_3); +x_5 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_3); if (x_5 == 0) { lean_object* x_6; @@ -8953,7 +8953,7 @@ return x_81; block_55: { uint8_t x_4; -x_4 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_2); +x_4 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_2); if (x_4 == 0) { lean_object* x_5; @@ -9026,7 +9026,7 @@ return x_54; block_42: { uint8_t x_8; -x_8 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_6); +x_8 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_6); if (x_8 == 0) { lean_object* x_9; @@ -9075,7 +9075,7 @@ if (x_15 == 0) lean_object* x_16; lean_object* x_17; uint8_t x_18; x_16 = lean_ctor_get(x_10, 0); x_17 = lean_ctor_get(x_10, 1); -x_18 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_16); +x_18 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_16); if (x_18 == 0) { lean_dec(x_1); @@ -9133,7 +9133,7 @@ x_28 = lean_ctor_get(x_10, 1); lean_inc(x_28); lean_inc(x_27); lean_dec(x_10); -x_29 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_27); +x_29 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_27); if (x_29 == 0) { lean_object* x_30; @@ -9263,7 +9263,7 @@ if (x_17 == 0) lean_object* x_18; lean_object* x_19; uint8_t x_20; x_18 = lean_ctor_get(x_12, 0); x_19 = lean_ctor_get(x_12, 1); -x_20 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_18); +x_20 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_18); if (x_20 == 0) { lean_dec(x_3); @@ -9288,7 +9288,7 @@ x_22 = lean_ctor_get(x_12, 1); lean_inc(x_22); lean_inc(x_21); lean_dec(x_12); -x_23 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_21); +x_23 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_21); if (x_23 == 0) { lean_object* x_24; @@ -9317,7 +9317,7 @@ else uint8_t x_26; lean_dec(x_7); lean_dec(x_6); -x_26 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_3); +x_26 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_3); if (x_26 == 0) { lean_object* x_27; @@ -9395,7 +9395,7 @@ if (lean_is_exclusive(x_34)) { lean_dec_ref(x_34); x_41 = lean_box(0); } -x_42 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_39); +x_42 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_39); if (x_42 == 0) { lean_object* x_43; @@ -9432,7 +9432,7 @@ else uint8_t x_45; lean_dec(x_29); lean_dec(x_28); -x_45 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_3); +x_45 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_3); if (x_45 == 0) { lean_object* x_46; lean_object* x_47; @@ -9464,7 +9464,7 @@ lean_object* x_50; lean_object* x_51; uint8_t x_52; x_50 = lean_ctor_get(x_4, 1); x_51 = lean_ctor_get(x_4, 0); lean_dec(x_51); -x_52 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_3); +x_52 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_3); if (x_52 == 0) { lean_dec(x_2); @@ -9486,7 +9486,7 @@ lean_object* x_53; uint8_t x_54; x_53 = lean_ctor_get(x_4, 1); lean_inc(x_53); lean_dec(x_4); -x_54 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_3); +x_54 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_3); if (x_54 == 0) { lean_object* x_55; @@ -9561,7 +9561,7 @@ if (x_13 == 0) lean_object* x_14; lean_object* x_15; uint8_t x_16; x_14 = lean_ctor_get(x_8, 0); x_15 = lean_ctor_get(x_8, 1); -x_16 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_14); +x_16 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_14); if (x_16 == 0) { lean_dec(x_2); @@ -9586,7 +9586,7 @@ x_18 = lean_ctor_get(x_8, 1); lean_inc(x_18); lean_inc(x_17); lean_dec(x_8); -x_19 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_17); +x_19 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_17); if (x_19 == 0) { lean_object* x_20; @@ -9619,7 +9619,7 @@ if (x_22 == 0) lean_object* x_23; lean_object* x_24; uint8_t x_25; x_23 = lean_ctor_get(x_3, 0); x_24 = lean_ctor_get(x_3, 1); -x_25 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_23); +x_25 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_23); if (x_25 == 0) { lean_dec(x_2); @@ -9644,7 +9644,7 @@ x_27 = lean_ctor_get(x_3, 1); lean_inc(x_27); lean_inc(x_26); lean_dec(x_3); -x_28 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_26); +x_28 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_26); if (x_28 == 0) { lean_object* x_29; @@ -9818,7 +9818,7 @@ goto block_42; block_42: { uint8_t x_5; -x_5 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_3); +x_5 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_3); if (x_5 == 0) { lean_object* x_6; @@ -10071,7 +10071,7 @@ if (x_14 == 0) lean_object* x_15; lean_object* x_16; uint8_t x_17; x_15 = lean_ctor_get(x_7, 0); x_16 = lean_ctor_get(x_7, 1); -x_17 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_15); +x_17 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_15); if (x_17 == 0) { lean_dec(x_1); @@ -10263,7 +10263,7 @@ x_55 = lean_ctor_get(x_7, 1); lean_inc(x_55); lean_inc(x_54); lean_dec(x_7); -x_56 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_54); +x_56 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_54); if (x_56 == 0) { lean_object* x_57; @@ -10468,7 +10468,7 @@ if (x_92 == 0) lean_object* x_93; lean_object* x_94; uint8_t x_95; x_93 = lean_ctor_get(x_5, 0); x_94 = lean_ctor_get(x_5, 1); -x_95 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_93); +x_95 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_93); if (x_95 == 0) { lean_dec(x_1); @@ -10660,7 +10660,7 @@ x_133 = lean_ctor_get(x_5, 1); lean_inc(x_133); lean_inc(x_132); lean_dec(x_5); -x_134 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_132); +x_134 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_132); if (x_134 == 0) { lean_object* x_135; @@ -10865,7 +10865,7 @@ if (x_170 == 0) lean_object* x_171; lean_object* x_172; uint8_t x_173; x_171 = lean_ctor_get(x_3, 0); x_172 = lean_ctor_get(x_3, 1); -x_173 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_171); +x_173 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_171); if (x_173 == 0) { lean_dec(x_1); @@ -11057,7 +11057,7 @@ x_211 = lean_ctor_get(x_3, 1); lean_inc(x_211); lean_inc(x_210); lean_dec(x_3); -x_212 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_210); +x_212 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_210); if (x_212 == 0) { lean_object* x_213; @@ -11308,7 +11308,7 @@ if (x_84 == 0) lean_object* x_85; lean_object* x_86; uint8_t x_87; x_85 = lean_ctor_get(x_82, 0); x_86 = lean_ctor_get(x_82, 1); -x_87 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_85); +x_87 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_85); if (x_87 == 0) { lean_dec(x_2); @@ -11333,7 +11333,7 @@ x_89 = lean_ctor_get(x_82, 1); lean_inc(x_89); lean_inc(x_88); lean_dec(x_82); -x_90 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_88); +x_90 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_88); if (x_90 == 0) { lean_object* x_91; @@ -11393,7 +11393,7 @@ if (x_11 == 0) lean_object* x_12; lean_object* x_13; uint8_t x_14; x_12 = lean_ctor_get(x_6, 0); x_13 = lean_ctor_get(x_6, 1); -x_14 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_12); +x_14 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_12); if (x_14 == 0) { lean_dec(x_2); @@ -11418,7 +11418,7 @@ x_16 = lean_ctor_get(x_6, 1); lean_inc(x_16); lean_inc(x_15); lean_dec(x_6); -x_17 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_15); +x_17 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_15); if (x_17 == 0) { lean_object* x_18; @@ -11450,7 +11450,7 @@ x_23 = l_String_Iterator_hasNext(x_21); if (x_23 == 0) { uint8_t x_24; -x_24 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_21); +x_24 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_21); if (x_24 == 0) { lean_object* x_25; lean_object* x_26; @@ -11483,7 +11483,7 @@ if (x_30 == 0) { uint8_t x_31; lean_dec(x_28); -x_31 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_21); +x_31 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_21); if (x_31 == 0) { lean_object* x_32; lean_object* x_33; @@ -11540,7 +11540,7 @@ if (x_40 == 0) lean_object* x_41; lean_object* x_42; uint8_t x_43; x_41 = lean_ctor_get(x_37, 0); x_42 = lean_ctor_get(x_37, 1); -x_43 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_41); +x_43 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_41); if (x_43 == 0) { lean_dec(x_2); @@ -11565,7 +11565,7 @@ x_45 = lean_ctor_get(x_37, 1); lean_inc(x_45); lean_inc(x_44); lean_dec(x_37); -x_46 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_44); +x_46 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_44); if (x_46 == 0) { lean_object* x_47; @@ -11598,12 +11598,12 @@ if (x_49 == 0) lean_object* x_50; lean_object* x_51; uint8_t x_52; x_50 = lean_ctor_get(x_35, 0); x_51 = lean_ctor_get(x_35, 1); -x_52 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_28, x_50); +x_52 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_28, x_50); if (x_52 == 0) { uint8_t x_53; lean_dec(x_28); -x_53 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_50); +x_53 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_50); if (x_53 == 0) { lean_dec(x_2); @@ -11648,7 +11648,7 @@ if (x_57 == 0) lean_object* x_58; lean_object* x_59; uint8_t x_60; x_58 = lean_ctor_get(x_54, 0); x_59 = lean_ctor_get(x_54, 1); -x_60 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_58); +x_60 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_58); if (x_60 == 0) { lean_dec(x_2); @@ -11673,7 +11673,7 @@ x_62 = lean_ctor_get(x_54, 1); lean_inc(x_62); lean_inc(x_61); lean_dec(x_54); -x_63 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_61); +x_63 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_61); if (x_63 == 0) { lean_object* x_64; @@ -11706,12 +11706,12 @@ x_67 = lean_ctor_get(x_35, 1); lean_inc(x_67); lean_inc(x_66); lean_dec(x_35); -x_68 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_28, x_66); +x_68 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_28, x_66); if (x_68 == 0) { uint8_t x_69; lean_dec(x_28); -x_69 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_66); +x_69 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_66); if (x_69 == 0) { lean_object* x_70; @@ -11766,7 +11766,7 @@ if (lean_is_exclusive(x_72)) { lean_dec_ref(x_72); x_77 = lean_box(0); } -x_78 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_75); +x_78 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_75); if (x_78 == 0) { lean_object* x_79; @@ -11957,7 +11957,7 @@ lean_inc(x_106); x_107 = lean_ctor_get(x_99, 1); lean_inc(x_107); lean_dec(x_99); -x_108 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_95, x_106); +x_108 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_95, x_106); if (x_108 == 0) { lean_dec(x_95); @@ -12075,7 +12075,7 @@ if (x_22 == 0) lean_object* x_23; lean_object* x_24; uint8_t x_25; x_23 = lean_ctor_get(x_19, 0); x_24 = lean_ctor_get(x_19, 1); -x_25 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_18, x_23); +x_25 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_18, x_23); if (x_25 == 0) { lean_dec(x_18); @@ -12101,7 +12101,7 @@ x_28 = lean_ctor_get(x_19, 1); lean_inc(x_28); lean_inc(x_27); lean_dec(x_19); -x_29 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_18, x_27); +x_29 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_18, x_27); if (x_29 == 0) { lean_object* x_30; @@ -12150,7 +12150,7 @@ return x_35; block_60: { uint8_t x_40; -x_40 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_38); +x_40 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_38); if (x_40 == 0) { lean_object* x_41; @@ -12217,7 +12217,7 @@ if (x_52 == 0) lean_object* x_53; lean_object* x_54; uint8_t x_55; x_53 = lean_ctor_get(x_50, 0); x_54 = lean_ctor_get(x_50, 1); -x_55 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_45, x_53); +x_55 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_45, x_53); if (x_55 == 0) { lean_dec(x_45); @@ -12240,7 +12240,7 @@ x_57 = lean_ctor_get(x_50, 1); lean_inc(x_57); lean_inc(x_56); lean_dec(x_50); -x_58 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_45, x_56); +x_58 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_45, x_56); if (x_58 == 0) { lean_object* x_59; @@ -12333,7 +12333,7 @@ lean_inc(x_77); x_78 = lean_ctor_get(x_65, 1); lean_inc(x_78); lean_dec(x_65); -x_79 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_64, x_77); +x_79 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_64, x_77); if (x_79 == 0) { lean_dec(x_64); @@ -12522,7 +12522,7 @@ if (x_48 == 0) lean_object* x_49; lean_object* x_50; uint8_t x_51; x_49 = lean_ctor_get(x_45, 0); x_50 = lean_ctor_get(x_45, 1); -x_51 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_49); +x_51 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_49); if (x_51 == 0) { lean_dec(x_1); @@ -12558,7 +12558,7 @@ if (x_55 == 0) lean_object* x_56; lean_object* x_57; uint8_t x_58; x_56 = lean_ctor_get(x_52, 0); x_57 = lean_ctor_get(x_52, 1); -x_58 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_56); +x_58 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_56); if (x_58 == 0) { lean_dec(x_1); @@ -12615,7 +12615,7 @@ x_67 = lean_ctor_get(x_52, 1); lean_inc(x_67); lean_inc(x_66); lean_dec(x_52); -x_68 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_66); +x_68 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_66); if (x_68 == 0) { lean_object* x_69; @@ -12680,7 +12680,7 @@ x_78 = lean_ctor_get(x_45, 1); lean_inc(x_78); lean_inc(x_77); lean_dec(x_45); -x_79 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_77); +x_79 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_77); if (x_79 == 0) { lean_object* x_80; @@ -12725,7 +12725,7 @@ if (lean_is_exclusive(x_81)) { lean_dec_ref(x_81); x_86 = lean_box(0); } -x_87 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_84); +x_87 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_84); if (x_87 == 0) { lean_object* x_88; @@ -12831,7 +12831,7 @@ return x_43; block_22: { uint8_t x_6; -x_6 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_4); +x_6 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_4); if (x_6 == 0) { lean_object* x_7; @@ -12852,7 +12852,7 @@ x_9 = l_String_Iterator_hasNext(x_2); if (x_9 == 0) { uint8_t x_10; -x_10 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_2); +x_10 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_2); if (x_10 == 0) { lean_object* x_11; lean_object* x_12; @@ -12883,7 +12883,7 @@ if (x_16 == 0) { uint8_t x_17; lean_dec(x_14); -x_17 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_2); +x_17 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_2); if (x_17 == 0) { lean_object* x_18; lean_object* x_19; @@ -12918,7 +12918,7 @@ return x_21; block_35: { uint8_t x_25; -x_25 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_23); +x_25 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_23); if (x_25 == 0) { lean_object* x_26; @@ -13031,7 +13031,7 @@ if (x_84 == 0) lean_object* x_85; lean_object* x_86; uint8_t x_87; x_85 = lean_ctor_get(x_82, 0); x_86 = lean_ctor_get(x_82, 1); -x_87 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_85); +x_87 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_85); if (x_87 == 0) { lean_dec(x_2); @@ -13056,7 +13056,7 @@ x_89 = lean_ctor_get(x_82, 1); lean_inc(x_89); lean_inc(x_88); lean_dec(x_82); -x_90 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_88); +x_90 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_88); if (x_90 == 0) { lean_object* x_91; @@ -13116,7 +13116,7 @@ if (x_11 == 0) lean_object* x_12; lean_object* x_13; uint8_t x_14; x_12 = lean_ctor_get(x_6, 0); x_13 = lean_ctor_get(x_6, 1); -x_14 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_12); +x_14 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_12); if (x_14 == 0) { lean_dec(x_2); @@ -13141,7 +13141,7 @@ x_16 = lean_ctor_get(x_6, 1); lean_inc(x_16); lean_inc(x_15); lean_dec(x_6); -x_17 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_15); +x_17 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_15); if (x_17 == 0) { lean_object* x_18; @@ -13173,7 +13173,7 @@ x_23 = l_String_Iterator_hasNext(x_21); if (x_23 == 0) { uint8_t x_24; -x_24 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_21); +x_24 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_21); if (x_24 == 0) { lean_object* x_25; lean_object* x_26; @@ -13206,7 +13206,7 @@ if (x_30 == 0) { uint8_t x_31; lean_dec(x_28); -x_31 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_21); +x_31 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_21); if (x_31 == 0) { lean_object* x_32; lean_object* x_33; @@ -13263,7 +13263,7 @@ if (x_40 == 0) lean_object* x_41; lean_object* x_42; uint8_t x_43; x_41 = lean_ctor_get(x_37, 0); x_42 = lean_ctor_get(x_37, 1); -x_43 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_41); +x_43 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_41); if (x_43 == 0) { lean_dec(x_2); @@ -13288,7 +13288,7 @@ x_45 = lean_ctor_get(x_37, 1); lean_inc(x_45); lean_inc(x_44); lean_dec(x_37); -x_46 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_44); +x_46 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_44); if (x_46 == 0) { lean_object* x_47; @@ -13321,12 +13321,12 @@ if (x_49 == 0) lean_object* x_50; lean_object* x_51; uint8_t x_52; x_50 = lean_ctor_get(x_35, 0); x_51 = lean_ctor_get(x_35, 1); -x_52 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_28, x_50); +x_52 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_28, x_50); if (x_52 == 0) { uint8_t x_53; lean_dec(x_28); -x_53 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_50); +x_53 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_50); if (x_53 == 0) { lean_dec(x_2); @@ -13371,7 +13371,7 @@ if (x_57 == 0) lean_object* x_58; lean_object* x_59; uint8_t x_60; x_58 = lean_ctor_get(x_54, 0); x_59 = lean_ctor_get(x_54, 1); -x_60 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_58); +x_60 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_58); if (x_60 == 0) { lean_dec(x_2); @@ -13396,7 +13396,7 @@ x_62 = lean_ctor_get(x_54, 1); lean_inc(x_62); lean_inc(x_61); lean_dec(x_54); -x_63 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_61); +x_63 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_61); if (x_63 == 0) { lean_object* x_64; @@ -13429,12 +13429,12 @@ x_67 = lean_ctor_get(x_35, 1); lean_inc(x_67); lean_inc(x_66); lean_dec(x_35); -x_68 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_28, x_66); +x_68 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_28, x_66); if (x_68 == 0) { uint8_t x_69; lean_dec(x_28); -x_69 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_66); +x_69 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_66); if (x_69 == 0) { lean_object* x_70; @@ -13489,7 +13489,7 @@ if (lean_is_exclusive(x_72)) { lean_dec_ref(x_72); x_77 = lean_box(0); } -x_78 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_75); +x_78 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_75); if (x_78 == 0) { lean_object* x_79; @@ -13585,7 +13585,7 @@ if (x_51 == 0) lean_object* x_52; lean_object* x_53; uint8_t x_54; x_52 = lean_ctor_get(x_49, 0); x_53 = lean_ctor_get(x_49, 1); -x_54 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_44, x_52); +x_54 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_44, x_52); if (x_54 == 0) { lean_dec(x_44); @@ -13608,7 +13608,7 @@ x_56 = lean_ctor_get(x_49, 1); lean_inc(x_56); lean_inc(x_55); lean_dec(x_49); -x_57 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_44, x_55); +x_57 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_44, x_55); if (x_57 == 0) { lean_object* x_58; @@ -13711,7 +13711,7 @@ if (x_23 == 0) lean_object* x_24; lean_object* x_25; uint8_t x_26; x_24 = lean_ctor_get(x_21, 0); x_25 = lean_ctor_get(x_21, 1); -x_26 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_20, x_24); +x_26 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_20, x_24); if (x_26 == 0) { lean_dec(x_20); @@ -13734,7 +13734,7 @@ x_28 = lean_ctor_get(x_21, 1); lean_inc(x_28); lean_inc(x_27); lean_dec(x_21); -x_29 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_20, x_27); +x_29 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_20, x_27); if (x_29 == 0) { lean_object* x_30; @@ -13826,7 +13826,7 @@ if (x_84 == 0) lean_object* x_85; lean_object* x_86; uint8_t x_87; x_85 = lean_ctor_get(x_82, 0); x_86 = lean_ctor_get(x_82, 1); -x_87 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_85); +x_87 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_85); if (x_87 == 0) { lean_dec(x_2); @@ -13851,7 +13851,7 @@ x_89 = lean_ctor_get(x_82, 1); lean_inc(x_89); lean_inc(x_88); lean_dec(x_82); -x_90 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_88); +x_90 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_88); if (x_90 == 0) { lean_object* x_91; @@ -13911,7 +13911,7 @@ if (x_11 == 0) lean_object* x_12; lean_object* x_13; uint8_t x_14; x_12 = lean_ctor_get(x_6, 0); x_13 = lean_ctor_get(x_6, 1); -x_14 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_12); +x_14 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_12); if (x_14 == 0) { lean_dec(x_2); @@ -13936,7 +13936,7 @@ x_16 = lean_ctor_get(x_6, 1); lean_inc(x_16); lean_inc(x_15); lean_dec(x_6); -x_17 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_15); +x_17 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_15); if (x_17 == 0) { lean_object* x_18; @@ -13968,7 +13968,7 @@ x_23 = l_String_Iterator_hasNext(x_21); if (x_23 == 0) { uint8_t x_24; -x_24 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_21); +x_24 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_21); if (x_24 == 0) { lean_object* x_25; lean_object* x_26; @@ -14001,7 +14001,7 @@ if (x_30 == 0) { uint8_t x_31; lean_dec(x_28); -x_31 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_21); +x_31 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_21); if (x_31 == 0) { lean_object* x_32; lean_object* x_33; @@ -14058,7 +14058,7 @@ if (x_40 == 0) lean_object* x_41; lean_object* x_42; uint8_t x_43; x_41 = lean_ctor_get(x_37, 0); x_42 = lean_ctor_get(x_37, 1); -x_43 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_41); +x_43 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_41); if (x_43 == 0) { lean_dec(x_2); @@ -14083,7 +14083,7 @@ x_45 = lean_ctor_get(x_37, 1); lean_inc(x_45); lean_inc(x_44); lean_dec(x_37); -x_46 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_44); +x_46 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_44); if (x_46 == 0) { lean_object* x_47; @@ -14116,12 +14116,12 @@ if (x_49 == 0) lean_object* x_50; lean_object* x_51; uint8_t x_52; x_50 = lean_ctor_get(x_35, 0); x_51 = lean_ctor_get(x_35, 1); -x_52 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_28, x_50); +x_52 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_28, x_50); if (x_52 == 0) { uint8_t x_53; lean_dec(x_28); -x_53 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_50); +x_53 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_50); if (x_53 == 0) { lean_dec(x_2); @@ -14166,7 +14166,7 @@ if (x_57 == 0) lean_object* x_58; lean_object* x_59; uint8_t x_60; x_58 = lean_ctor_get(x_54, 0); x_59 = lean_ctor_get(x_54, 1); -x_60 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_58); +x_60 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_58); if (x_60 == 0) { lean_dec(x_2); @@ -14191,7 +14191,7 @@ x_62 = lean_ctor_get(x_54, 1); lean_inc(x_62); lean_inc(x_61); lean_dec(x_54); -x_63 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_61); +x_63 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_61); if (x_63 == 0) { lean_object* x_64; @@ -14224,12 +14224,12 @@ x_67 = lean_ctor_get(x_35, 1); lean_inc(x_67); lean_inc(x_66); lean_dec(x_35); -x_68 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_28, x_66); +x_68 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_28, x_66); if (x_68 == 0) { uint8_t x_69; lean_dec(x_28); -x_69 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_66); +x_69 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_66); if (x_69 == 0) { lean_object* x_70; @@ -14284,7 +14284,7 @@ if (lean_is_exclusive(x_72)) { lean_dec_ref(x_72); x_77 = lean_box(0); } -x_78 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_75); +x_78 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_75); if (x_78 == 0) { lean_object* x_79; @@ -14410,7 +14410,7 @@ if (x_92 == 0) lean_object* x_93; lean_object* x_94; uint8_t x_95; x_93 = lean_ctor_get(x_84, 0); x_94 = lean_ctor_get(x_84, 1); -x_95 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_79, x_93); +x_95 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_79, x_93); if (x_95 == 0) { lean_dec(x_79); @@ -14464,7 +14464,7 @@ x_103 = lean_ctor_get(x_84, 1); lean_inc(x_103); lean_inc(x_102); lean_dec(x_84); -x_104 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_79, x_102); +x_104 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_79, x_102); if (x_104 == 0) { lean_object* x_105; @@ -14585,7 +14585,7 @@ if (x_66 == 0) lean_object* x_67; lean_object* x_68; uint8_t x_69; x_67 = lean_ctor_get(x_64, 0); x_68 = lean_ctor_get(x_64, 1); -x_69 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_15, x_67); +x_69 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_15, x_67); if (x_69 == 0) { lean_dec(x_15); @@ -14608,7 +14608,7 @@ x_71 = lean_ctor_get(x_64, 1); lean_inc(x_71); lean_inc(x_70); lean_dec(x_64); -x_72 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_15, x_70); +x_72 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_15, x_70); if (x_72 == 0) { lean_object* x_73; @@ -14669,7 +14669,7 @@ if (x_26 == 0) lean_object* x_27; lean_object* x_28; uint8_t x_29; x_27 = lean_ctor_get(x_24, 0); x_28 = lean_ctor_get(x_24, 1); -x_29 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_23, x_27); +x_29 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_23, x_27); if (x_29 == 0) { lean_dec(x_23); @@ -14692,7 +14692,7 @@ x_31 = lean_ctor_get(x_24, 1); lean_inc(x_31); lean_inc(x_30); lean_dec(x_24); -x_32 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_23, x_30); +x_32 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_23, x_30); if (x_32 == 0) { lean_object* x_33; @@ -14814,7 +14814,7 @@ if (x_55 == 0) lean_object* x_56; lean_object* x_57; uint8_t x_58; x_56 = lean_ctor_get(x_53, 0); x_57 = lean_ctor_get(x_53, 1); -x_58 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_48, x_56); +x_58 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_48, x_56); if (x_58 == 0) { lean_dec(x_48); @@ -14837,7 +14837,7 @@ x_60 = lean_ctor_get(x_53, 1); lean_inc(x_60); lean_inc(x_59); lean_dec(x_53); -x_61 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_48, x_59); +x_61 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_48, x_59); if (x_61 == 0) { lean_object* x_62; @@ -14890,7 +14890,7 @@ if (x_48 == 0) lean_object* x_49; lean_object* x_50; uint8_t x_51; x_49 = lean_ctor_get(x_45, 0); x_50 = lean_ctor_get(x_45, 1); -x_51 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_49); +x_51 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_49); if (x_51 == 0) { lean_dec(x_1); @@ -14947,7 +14947,7 @@ x_60 = lean_ctor_get(x_45, 1); lean_inc(x_60); lean_inc(x_59); lean_dec(x_45); -x_61 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_59); +x_61 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_59); if (x_61 == 0) { lean_object* x_62; @@ -15046,7 +15046,7 @@ return x_43; block_22: { uint8_t x_6; -x_6 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_4); +x_6 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_4); if (x_6 == 0) { lean_object* x_7; @@ -15067,7 +15067,7 @@ x_9 = l_String_Iterator_hasNext(x_2); if (x_9 == 0) { uint8_t x_10; -x_10 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_2); +x_10 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_2); if (x_10 == 0) { lean_object* x_11; lean_object* x_12; @@ -15098,7 +15098,7 @@ if (x_16 == 0) { uint8_t x_17; lean_dec(x_14); -x_17 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_2); +x_17 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_2); if (x_17 == 0) { lean_object* x_18; lean_object* x_19; @@ -15133,7 +15133,7 @@ return x_21; block_35: { uint8_t x_25; -x_25 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_23); +x_25 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_23); if (x_25 == 0) { lean_object* x_26; @@ -15251,7 +15251,7 @@ if (x_10 == 0) lean_object* x_11; lean_object* x_12; uint8_t x_13; x_11 = lean_ctor_get(x_3, 0); x_12 = lean_ctor_get(x_3, 1); -x_13 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_11); +x_13 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_11); if (x_13 == 0) { lean_dec(x_1); @@ -15302,7 +15302,7 @@ if (x_22 == 0) lean_object* x_23; lean_object* x_24; uint8_t x_25; x_23 = lean_ctor_get(x_15, 0); x_24 = lean_ctor_get(x_15, 1); -x_25 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_23); +x_25 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_23); if (x_25 == 0) { lean_dec(x_1); @@ -15348,7 +15348,7 @@ if (x_31 == 0) lean_object* x_32; lean_object* x_33; uint8_t x_34; x_32 = lean_ctor_get(x_26, 0); x_33 = lean_ctor_get(x_26, 1); -x_34 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_32); +x_34 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_32); if (x_34 == 0) { lean_dec(x_1); @@ -15372,7 +15372,7 @@ x_37 = lean_ctor_get(x_26, 1); lean_inc(x_37); lean_inc(x_36); lean_dec(x_26); -x_38 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_36); +x_38 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_36); if (x_38 == 0) { lean_object* x_39; @@ -15402,7 +15402,7 @@ x_42 = lean_ctor_get(x_15, 1); lean_inc(x_42); lean_inc(x_41); lean_dec(x_15); -x_43 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_41); +x_43 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_41); if (x_43 == 0) { lean_object* x_44; @@ -15459,7 +15459,7 @@ if (lean_is_exclusive(x_45)) { lean_dec_ref(x_45); x_52 = lean_box(0); } -x_53 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_50); +x_53 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_50); if (x_53 == 0) { lean_object* x_54; @@ -15496,7 +15496,7 @@ x_57 = lean_ctor_get(x_3, 1); lean_inc(x_57); lean_inc(x_56); lean_dec(x_3); -x_58 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_56); +x_58 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_56); if (x_58 == 0) { lean_object* x_59; @@ -15553,7 +15553,7 @@ if (lean_is_exclusive(x_61)) { lean_dec_ref(x_61); x_68 = lean_box(0); } -x_69 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_66); +x_69 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_66); if (x_69 == 0) { lean_object* x_70; @@ -15615,7 +15615,7 @@ if (lean_is_exclusive(x_71)) { lean_dec_ref(x_71); x_78 = lean_box(0); } -x_79 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_76); +x_79 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_76); if (x_79 == 0) { lean_object* x_80; @@ -15738,7 +15738,7 @@ if (x_26 == 0) lean_object* x_27; lean_object* x_28; uint8_t x_29; x_27 = lean_ctor_get(x_24, 0); x_28 = lean_ctor_get(x_24, 1); -x_29 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_23, x_27); +x_29 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_23, x_27); if (x_29 == 0) { lean_dec(x_23); @@ -15761,7 +15761,7 @@ x_31 = lean_ctor_get(x_24, 1); lean_inc(x_31); lean_inc(x_30); lean_dec(x_24); -x_32 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_23, x_30); +x_32 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_23, x_30); if (x_32 == 0) { lean_object* x_33; @@ -16080,7 +16080,7 @@ if (x_10 == 0) lean_object* x_11; lean_object* x_12; uint8_t x_13; x_11 = lean_ctor_get(x_3, 0); x_12 = lean_ctor_get(x_3, 1); -x_13 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_11); +x_13 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_11); if (x_13 == 0) { lean_dec(x_1); @@ -16131,7 +16131,7 @@ if (x_22 == 0) lean_object* x_23; lean_object* x_24; uint8_t x_25; x_23 = lean_ctor_get(x_15, 0); x_24 = lean_ctor_get(x_15, 1); -x_25 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_23); +x_25 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_23); if (x_25 == 0) { lean_dec(x_1); @@ -16182,7 +16182,7 @@ if (x_34 == 0) lean_object* x_35; lean_object* x_36; uint8_t x_37; x_35 = lean_ctor_get(x_27, 0); x_36 = lean_ctor_get(x_27, 1); -x_37 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_35); +x_37 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_35); if (x_37 == 0) { lean_dec(x_1); @@ -16233,7 +16233,7 @@ if (x_46 == 0) lean_object* x_47; lean_object* x_48; uint8_t x_49; x_47 = lean_ctor_get(x_39, 0); x_48 = lean_ctor_get(x_39, 1); -x_49 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_47); +x_49 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_47); if (x_49 == 0) { lean_dec(x_1); @@ -16284,7 +16284,7 @@ if (x_58 == 0) lean_object* x_59; lean_object* x_60; uint8_t x_61; x_59 = lean_ctor_get(x_51, 0); x_60 = lean_ctor_get(x_51, 1); -x_61 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_59); +x_61 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_59); if (x_61 == 0) { lean_dec(x_1); @@ -16335,7 +16335,7 @@ if (x_70 == 0) lean_object* x_71; lean_object* x_72; uint8_t x_73; x_71 = lean_ctor_get(x_63, 0); x_72 = lean_ctor_get(x_63, 1); -x_73 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_71); +x_73 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_71); if (x_73 == 0) { lean_dec(x_1); @@ -16407,7 +16407,7 @@ x_87 = lean_ctor_get(x_63, 1); lean_inc(x_87); lean_inc(x_86); lean_dec(x_63); -x_88 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_86); +x_88 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_86); if (x_88 == 0) { lean_object* x_89; @@ -16484,7 +16484,7 @@ x_101 = lean_ctor_get(x_51, 1); lean_inc(x_101); lean_inc(x_100); lean_dec(x_51); -x_102 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_100); +x_102 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_100); if (x_102 == 0) { lean_object* x_103; @@ -16541,7 +16541,7 @@ if (lean_is_exclusive(x_105)) { lean_dec_ref(x_105); x_112 = lean_box(0); } -x_113 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_110); +x_113 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_110); if (x_113 == 0) { lean_object* x_114; @@ -16625,7 +16625,7 @@ x_126 = lean_ctor_get(x_39, 1); lean_inc(x_126); lean_inc(x_125); lean_dec(x_39); -x_127 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_125); +x_127 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_125); if (x_127 == 0) { lean_object* x_128; @@ -16682,7 +16682,7 @@ if (lean_is_exclusive(x_130)) { lean_dec_ref(x_130); x_137 = lean_box(0); } -x_138 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_135); +x_138 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_135); if (x_138 == 0) { lean_object* x_139; @@ -16744,7 +16744,7 @@ if (lean_is_exclusive(x_141)) { lean_dec_ref(x_141); x_148 = lean_box(0); } -x_149 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_146); +x_149 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_146); if (x_149 == 0) { lean_object* x_150; @@ -16830,7 +16830,7 @@ x_162 = lean_ctor_get(x_27, 1); lean_inc(x_162); lean_inc(x_161); lean_dec(x_27); -x_163 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_161); +x_163 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_161); if (x_163 == 0) { lean_object* x_164; @@ -16887,7 +16887,7 @@ if (lean_is_exclusive(x_166)) { lean_dec_ref(x_166); x_173 = lean_box(0); } -x_174 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_171); +x_174 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_171); if (x_174 == 0) { lean_object* x_175; @@ -16949,7 +16949,7 @@ if (lean_is_exclusive(x_177)) { lean_dec_ref(x_177); x_184 = lean_box(0); } -x_185 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_182); +x_185 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_182); if (x_185 == 0) { lean_object* x_186; @@ -17011,7 +17011,7 @@ if (lean_is_exclusive(x_188)) { lean_dec_ref(x_188); x_195 = lean_box(0); } -x_196 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_193); +x_196 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_193); if (x_196 == 0) { lean_object* x_197; @@ -17099,7 +17099,7 @@ x_209 = lean_ctor_get(x_15, 1); lean_inc(x_209); lean_inc(x_208); lean_dec(x_15); -x_210 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_208); +x_210 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_208); if (x_210 == 0) { lean_object* x_211; @@ -17156,7 +17156,7 @@ if (lean_is_exclusive(x_213)) { lean_dec_ref(x_213); x_220 = lean_box(0); } -x_221 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_218); +x_221 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_218); if (x_221 == 0) { lean_object* x_222; @@ -17218,7 +17218,7 @@ if (lean_is_exclusive(x_224)) { lean_dec_ref(x_224); x_231 = lean_box(0); } -x_232 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_229); +x_232 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_229); if (x_232 == 0) { lean_object* x_233; @@ -17280,7 +17280,7 @@ if (lean_is_exclusive(x_235)) { lean_dec_ref(x_235); x_242 = lean_box(0); } -x_243 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_240); +x_243 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_240); if (x_243 == 0) { lean_object* x_244; @@ -17342,7 +17342,7 @@ if (lean_is_exclusive(x_246)) { lean_dec_ref(x_246); x_253 = lean_box(0); } -x_254 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_251); +x_254 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_251); if (x_254 == 0) { lean_object* x_255; @@ -17432,7 +17432,7 @@ x_267 = lean_ctor_get(x_3, 1); lean_inc(x_267); lean_inc(x_266); lean_dec(x_3); -x_268 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_266); +x_268 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_266); if (x_268 == 0) { lean_object* x_269; @@ -17489,7 +17489,7 @@ if (lean_is_exclusive(x_271)) { lean_dec_ref(x_271); x_278 = lean_box(0); } -x_279 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_276); +x_279 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_276); if (x_279 == 0) { lean_object* x_280; @@ -17551,7 +17551,7 @@ if (lean_is_exclusive(x_282)) { lean_dec_ref(x_282); x_289 = lean_box(0); } -x_290 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_287); +x_290 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_287); if (x_290 == 0) { lean_object* x_291; @@ -17613,7 +17613,7 @@ if (lean_is_exclusive(x_293)) { lean_dec_ref(x_293); x_300 = lean_box(0); } -x_301 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_298); +x_301 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_298); if (x_301 == 0) { lean_object* x_302; @@ -17675,7 +17675,7 @@ if (lean_is_exclusive(x_304)) { lean_dec_ref(x_304); x_311 = lean_box(0); } -x_312 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_309); +x_312 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_309); if (x_312 == 0) { lean_object* x_313; @@ -17737,7 +17737,7 @@ if (lean_is_exclusive(x_315)) { lean_dec_ref(x_315); x_322 = lean_box(0); } -x_323 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_320); +x_323 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_320); if (x_323 == 0) { lean_object* x_324; @@ -17904,7 +17904,7 @@ if (x_56 == 0) lean_object* x_57; lean_object* x_58; uint8_t x_59; x_57 = lean_ctor_get(x_54, 0); x_58 = lean_ctor_get(x_54, 1); -x_59 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_50, x_57); +x_59 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_50, x_57); if (x_59 == 0) { lean_dec(x_50); @@ -17927,7 +17927,7 @@ x_61 = lean_ctor_get(x_54, 1); lean_inc(x_61); lean_inc(x_60); lean_dec(x_54); -x_62 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_50, x_60); +x_62 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_50, x_60); if (x_62 == 0) { lean_object* x_63; @@ -18014,7 +18014,7 @@ if (lean_is_exclusive(x_74)) { lean_dec_ref(x_74); x_78 = lean_box(0); } -x_79 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_69, x_76); +x_79 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_69, x_76); if (x_79 == 0) { lean_object* x_80; @@ -18168,7 +18168,7 @@ if (x_23 == 0) lean_object* x_24; lean_object* x_25; uint8_t x_26; x_24 = lean_ctor_get(x_21, 0); x_25 = lean_ctor_get(x_21, 1); -x_26 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_20, x_24); +x_26 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_20, x_24); if (x_26 == 0) { lean_dec(x_20); @@ -18191,7 +18191,7 @@ x_28 = lean_ctor_get(x_21, 1); lean_inc(x_28); lean_inc(x_27); lean_dec(x_21); -x_29 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_20, x_27); +x_29 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_20, x_27); if (x_29 == 0) { lean_object* x_30; @@ -18327,7 +18327,7 @@ if (x_84 == 0) lean_object* x_85; lean_object* x_86; uint8_t x_87; x_85 = lean_ctor_get(x_82, 0); x_86 = lean_ctor_get(x_82, 1); -x_87 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_85); +x_87 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_85); if (x_87 == 0) { lean_dec(x_2); @@ -18352,7 +18352,7 @@ x_89 = lean_ctor_get(x_82, 1); lean_inc(x_89); lean_inc(x_88); lean_dec(x_82); -x_90 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_88); +x_90 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_88); if (x_90 == 0) { lean_object* x_91; @@ -18412,7 +18412,7 @@ if (x_11 == 0) lean_object* x_12; lean_object* x_13; uint8_t x_14; x_12 = lean_ctor_get(x_6, 0); x_13 = lean_ctor_get(x_6, 1); -x_14 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_12); +x_14 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_12); if (x_14 == 0) { lean_dec(x_2); @@ -18437,7 +18437,7 @@ x_16 = lean_ctor_get(x_6, 1); lean_inc(x_16); lean_inc(x_15); lean_dec(x_6); -x_17 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_15); +x_17 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_15); if (x_17 == 0) { lean_object* x_18; @@ -18469,7 +18469,7 @@ x_23 = l_String_Iterator_hasNext(x_21); if (x_23 == 0) { uint8_t x_24; -x_24 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_21); +x_24 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_21); if (x_24 == 0) { lean_object* x_25; lean_object* x_26; @@ -18502,7 +18502,7 @@ if (x_30 == 0) { uint8_t x_31; lean_dec(x_28); -x_31 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_21); +x_31 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_21); if (x_31 == 0) { lean_object* x_32; lean_object* x_33; @@ -18559,7 +18559,7 @@ if (x_40 == 0) lean_object* x_41; lean_object* x_42; uint8_t x_43; x_41 = lean_ctor_get(x_37, 0); x_42 = lean_ctor_get(x_37, 1); -x_43 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_41); +x_43 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_41); if (x_43 == 0) { lean_dec(x_2); @@ -18584,7 +18584,7 @@ x_45 = lean_ctor_get(x_37, 1); lean_inc(x_45); lean_inc(x_44); lean_dec(x_37); -x_46 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_44); +x_46 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_44); if (x_46 == 0) { lean_object* x_47; @@ -18617,12 +18617,12 @@ if (x_49 == 0) lean_object* x_50; lean_object* x_51; uint8_t x_52; x_50 = lean_ctor_get(x_35, 0); x_51 = lean_ctor_get(x_35, 1); -x_52 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_28, x_50); +x_52 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_28, x_50); if (x_52 == 0) { uint8_t x_53; lean_dec(x_28); -x_53 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_50); +x_53 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_50); if (x_53 == 0) { lean_dec(x_2); @@ -18667,7 +18667,7 @@ if (x_57 == 0) lean_object* x_58; lean_object* x_59; uint8_t x_60; x_58 = lean_ctor_get(x_54, 0); x_59 = lean_ctor_get(x_54, 1); -x_60 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_58); +x_60 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_58); if (x_60 == 0) { lean_dec(x_2); @@ -18692,7 +18692,7 @@ x_62 = lean_ctor_get(x_54, 1); lean_inc(x_62); lean_inc(x_61); lean_dec(x_54); -x_63 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_61); +x_63 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_61); if (x_63 == 0) { lean_object* x_64; @@ -18725,12 +18725,12 @@ x_67 = lean_ctor_get(x_35, 1); lean_inc(x_67); lean_inc(x_66); lean_dec(x_35); -x_68 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_28, x_66); +x_68 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_28, x_66); if (x_68 == 0) { uint8_t x_69; lean_dec(x_28); -x_69 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_66); +x_69 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_66); if (x_69 == 0) { lean_object* x_70; @@ -18785,7 +18785,7 @@ if (lean_is_exclusive(x_72)) { lean_dec_ref(x_72); x_77 = lean_box(0); } -x_78 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_75); +x_78 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_75); if (x_78 == 0) { lean_object* x_79; @@ -18881,7 +18881,7 @@ if (x_51 == 0) lean_object* x_52; lean_object* x_53; uint8_t x_54; x_52 = lean_ctor_get(x_49, 0); x_53 = lean_ctor_get(x_49, 1); -x_54 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_44, x_52); +x_54 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_44, x_52); if (x_54 == 0) { lean_dec(x_44); @@ -18904,7 +18904,7 @@ x_56 = lean_ctor_get(x_49, 1); lean_inc(x_56); lean_inc(x_55); lean_dec(x_49); -x_57 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_44, x_55); +x_57 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_44, x_55); if (x_57 == 0) { lean_object* x_58; @@ -19007,7 +19007,7 @@ if (x_23 == 0) lean_object* x_24; lean_object* x_25; uint8_t x_26; x_24 = lean_ctor_get(x_21, 0); x_25 = lean_ctor_get(x_21, 1); -x_26 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_20, x_24); +x_26 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_20, x_24); if (x_26 == 0) { lean_dec(x_20); @@ -19030,7 +19030,7 @@ x_28 = lean_ctor_get(x_21, 1); lean_inc(x_28); lean_inc(x_27); lean_dec(x_21); -x_29 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_20, x_27); +x_29 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_20, x_27); if (x_29 == 0) { lean_object* x_30; @@ -19136,7 +19136,7 @@ if (x_7 == 0) lean_object* x_8; lean_object* x_9; uint8_t x_10; x_8 = lean_ctor_get(x_2, 0); x_9 = lean_ctor_get(x_2, 1); -x_10 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_8); +x_10 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_8); if (x_10 == 0) { lean_dec(x_1); @@ -19160,7 +19160,7 @@ x_13 = lean_ctor_get(x_2, 1); lean_inc(x_13); lean_inc(x_12); lean_dec(x_2); -x_14 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_12); +x_14 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_12); if (x_14 == 0) { lean_object* x_15; @@ -19220,7 +19220,7 @@ if (x_7 == 0) lean_object* x_8; lean_object* x_9; uint8_t x_10; x_8 = lean_ctor_get(x_2, 0); x_9 = lean_ctor_get(x_2, 1); -x_10 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_8); +x_10 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_8); if (x_10 == 0) { lean_dec(x_1); @@ -19266,7 +19266,7 @@ if (x_16 == 0) lean_object* x_17; lean_object* x_18; uint8_t x_19; x_17 = lean_ctor_get(x_11, 0); x_18 = lean_ctor_get(x_11, 1); -x_19 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_17); +x_19 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_17); if (x_19 == 0) { lean_dec(x_1); @@ -19290,7 +19290,7 @@ x_22 = lean_ctor_get(x_11, 1); lean_inc(x_22); lean_inc(x_21); lean_dec(x_11); -x_23 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_21); +x_23 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_21); if (x_23 == 0) { lean_object* x_24; @@ -19320,7 +19320,7 @@ x_27 = lean_ctor_get(x_2, 1); lean_inc(x_27); lean_inc(x_26); lean_dec(x_2); -x_28 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_26); +x_28 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_26); if (x_28 == 0) { lean_object* x_29; @@ -19377,7 +19377,7 @@ if (lean_is_exclusive(x_30)) { lean_dec_ref(x_30); x_37 = lean_box(0); } -x_38 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_35); +x_38 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_35); if (x_38 == 0) { lean_object* x_39; @@ -20047,7 +20047,7 @@ x_60 = l_String_Iterator_hasNext(x_2); if (x_60 == 0) { uint8_t x_61; -x_61 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_2); +x_61 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_2); if (x_61 == 0) { lean_object* x_62; lean_object* x_63; @@ -20236,7 +20236,7 @@ if (x_11 == 0) lean_object* x_12; lean_object* x_13; uint8_t x_14; x_12 = lean_ctor_get(x_6, 0); x_13 = lean_ctor_get(x_6, 1); -x_14 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_12); +x_14 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_12); if (x_14 == 0) { lean_dec(x_2); @@ -20261,7 +20261,7 @@ x_16 = lean_ctor_get(x_6, 1); lean_inc(x_16); lean_inc(x_15); lean_dec(x_6); -x_17 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_15); +x_17 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_15); if (x_17 == 0) { lean_object* x_18; @@ -20288,7 +20288,7 @@ return x_19; block_26: { uint8_t x_23; -x_23 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_21); +x_23 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_21); if (x_23 == 0) { lean_object* x_24; @@ -20412,7 +20412,7 @@ x_3 = l_String_Iterator_hasNext(x_2); if (x_3 == 0) { uint8_t x_4; -x_4 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_2); +x_4 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_2); if (x_4 == 0) { lean_object* x_5; lean_object* x_6; @@ -20444,7 +20444,7 @@ if (x_11 == 0) { uint8_t x_12; lean_dec(x_8); -x_12 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_2); +x_12 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_2); if (x_12 == 0) { lean_object* x_13; lean_object* x_14; @@ -20473,7 +20473,7 @@ if (x_17 == 0) { uint8_t x_18; lean_dec(x_8); -x_18 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_2); +x_18 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_2); if (x_18 == 0) { lean_object* x_19; lean_object* x_20; @@ -20536,7 +20536,7 @@ if (x_31 == 0) lean_object* x_32; lean_object* x_33; uint8_t x_34; x_32 = lean_ctor_get(x_26, 0); x_33 = lean_ctor_get(x_26, 1); -x_34 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_32); +x_34 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_32); if (x_34 == 0) { lean_dec(x_2); @@ -20561,7 +20561,7 @@ x_36 = lean_ctor_get(x_26, 1); lean_inc(x_36); lean_inc(x_35); lean_dec(x_26); -x_37 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_35); +x_37 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_35); if (x_37 == 0) { lean_object* x_38; @@ -20697,7 +20697,7 @@ if (x_138 == 0) lean_object* x_139; lean_object* x_140; uint8_t x_141; x_139 = lean_ctor_get(x_133, 0); x_140 = lean_ctor_get(x_133, 1); -x_141 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_102, x_139); +x_141 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_102, x_139); if (x_141 == 0) { lean_dec(x_120); @@ -20737,7 +20737,7 @@ x_146 = lean_ctor_get(x_133, 1); lean_inc(x_146); lean_inc(x_145); lean_dec(x_133); -x_147 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_102, x_145); +x_147 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_102, x_145); if (x_147 == 0) { lean_object* x_148; @@ -20780,7 +20780,7 @@ goto block_99; block_117: { uint8_t x_106; -x_106 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_102, x_104); +x_106 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_102, x_104); if (x_106 == 0) { lean_object* x_107; @@ -21244,7 +21244,7 @@ if (x_7 == 0) lean_object* x_8; lean_object* x_9; uint8_t x_10; x_8 = lean_ctor_get(x_2, 0); x_9 = lean_ctor_get(x_2, 1); -x_10 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_8); +x_10 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_8); if (x_10 == 0) { lean_dec(x_1); @@ -21318,7 +21318,7 @@ x_24 = lean_ctor_get(x_2, 1); lean_inc(x_24); lean_inc(x_23); lean_dec(x_2); -x_25 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_23); +x_25 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_23); if (x_25 == 0) { lean_object* x_26; @@ -21433,7 +21433,7 @@ x_28 = l_String_Iterator_hasNext(x_3); if (x_28 == 0) { uint8_t x_29; -x_29 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_3); +x_29 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_3); if (x_29 == 0) { lean_object* x_30; @@ -21487,7 +21487,7 @@ if (x_40 == 0) { uint8_t x_41; lean_dec(x_36); -x_41 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_3); +x_41 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_3); if (x_41 == 0) { lean_object* x_42; @@ -21538,7 +21538,7 @@ if (x_50 == 0) { uint8_t x_51; lean_dec(x_36); -x_51 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_3); +x_51 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_3); if (x_51 == 0) { lean_object* x_52; @@ -21588,7 +21588,7 @@ if (x_59 == 0) { uint8_t x_60; lean_dec(x_36); -x_60 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_3); +x_60 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_3); if (x_60 == 0) { lean_object* x_61; @@ -21681,7 +21681,7 @@ if (x_12 == 0) lean_object* x_13; lean_object* x_14; uint8_t x_15; x_13 = lean_ctor_get(x_7, 0); x_14 = lean_ctor_get(x_7, 1); -x_15 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_13); +x_15 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_13); if (x_15 == 0) { lean_dec(x_3); @@ -21706,7 +21706,7 @@ x_17 = lean_ctor_get(x_7, 1); lean_inc(x_17); lean_inc(x_16); lean_dec(x_7); -x_18 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_16); +x_18 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_16); if (x_18 == 0) { lean_object* x_19; @@ -21733,7 +21733,7 @@ return x_20; block_27: { uint8_t x_24; -x_24 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_22); +x_24 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_22); if (x_24 == 0) { lean_object* x_25; @@ -21765,7 +21765,7 @@ x_28 = l_String_Iterator_hasNext(x_3); if (x_28 == 0) { uint8_t x_29; -x_29 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_3); +x_29 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_3); if (x_29 == 0) { lean_object* x_30; @@ -21819,7 +21819,7 @@ if (x_40 == 0) { uint8_t x_41; lean_dec(x_36); -x_41 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_3); +x_41 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_3); if (x_41 == 0) { lean_object* x_42; @@ -21870,7 +21870,7 @@ if (x_50 == 0) { uint8_t x_51; lean_dec(x_36); -x_51 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_3); +x_51 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_3); if (x_51 == 0) { lean_object* x_52; @@ -21920,7 +21920,7 @@ if (x_59 == 0) { uint8_t x_60; lean_dec(x_36); -x_60 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_3); +x_60 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_3); if (x_60 == 0) { lean_object* x_61; @@ -22013,7 +22013,7 @@ if (x_12 == 0) lean_object* x_13; lean_object* x_14; uint8_t x_15; x_13 = lean_ctor_get(x_7, 0); x_14 = lean_ctor_get(x_7, 1); -x_15 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_13); +x_15 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_13); if (x_15 == 0) { lean_dec(x_3); @@ -22038,7 +22038,7 @@ x_17 = lean_ctor_get(x_7, 1); lean_inc(x_17); lean_inc(x_16); lean_dec(x_7); -x_18 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_16); +x_18 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_16); if (x_18 == 0) { lean_object* x_19; @@ -22065,7 +22065,7 @@ return x_20; block_27: { uint8_t x_24; -x_24 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_22); +x_24 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_22); if (x_24 == 0) { lean_object* x_25; @@ -22236,7 +22236,7 @@ return x_16; block_56: { uint8_t x_20; -x_20 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_18); +x_20 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_18); if (x_20 == 0) { lean_object* x_21; @@ -22506,7 +22506,7 @@ if (x_10 == 0) lean_object* x_11; lean_object* x_12; uint8_t x_13; x_11 = lean_ctor_get(x_3, 0); x_12 = lean_ctor_get(x_3, 1); -x_13 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_11); +x_13 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_11); if (x_13 == 0) { lean_dec(x_1); @@ -22557,7 +22557,7 @@ if (x_22 == 0) lean_object* x_23; lean_object* x_24; uint8_t x_25; x_23 = lean_ctor_get(x_15, 0); x_24 = lean_ctor_get(x_15, 1); -x_25 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_23); +x_25 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_23); if (x_25 == 0) { lean_dec(x_1); @@ -22645,7 +22645,7 @@ if (x_42 == 0) lean_object* x_43; lean_object* x_44; uint8_t x_45; x_43 = lean_ctor_get(x_29, 0); x_44 = lean_ctor_get(x_29, 1); -x_45 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_43); +x_45 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_43); if (x_45 == 0) { lean_dec(x_1); @@ -22716,7 +22716,7 @@ x_58 = lean_ctor_get(x_29, 1); lean_inc(x_58); lean_inc(x_57); lean_dec(x_29); -x_59 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_57); +x_59 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_57); if (x_59 == 0) { lean_object* x_60; @@ -22792,7 +22792,7 @@ if (x_70 == 0) lean_object* x_71; lean_object* x_72; uint8_t x_73; x_71 = lean_ctor_get(x_27, 0); x_72 = lean_ctor_get(x_27, 1); -x_73 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_71); +x_73 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_71); if (x_73 == 0) { lean_dec(x_1); @@ -22863,7 +22863,7 @@ x_86 = lean_ctor_get(x_27, 1); lean_inc(x_86); lean_inc(x_85); lean_dec(x_27); -x_87 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_85); +x_87 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_85); if (x_87 == 0) { lean_object* x_88; @@ -22939,7 +22939,7 @@ x_99 = lean_ctor_get(x_15, 1); lean_inc(x_99); lean_inc(x_98); lean_dec(x_15); -x_100 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_98); +x_100 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_98); if (x_100 == 0) { lean_object* x_101; @@ -23035,7 +23035,7 @@ if (lean_is_exclusive(x_105)) { lean_dec_ref(x_105); x_118 = lean_box(0); } -x_119 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_116); +x_119 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_116); if (x_119 == 0) { lean_object* x_120; @@ -23121,7 +23121,7 @@ if (lean_is_exclusive(x_103)) { lean_dec_ref(x_103); x_132 = lean_box(0); } -x_133 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_130); +x_133 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_130); if (x_133 == 0) { lean_object* x_134; @@ -23204,7 +23204,7 @@ x_145 = lean_ctor_get(x_3, 1); lean_inc(x_145); lean_inc(x_144); lean_dec(x_3); -x_146 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_144); +x_146 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_144); if (x_146 == 0) { lean_object* x_147; @@ -23261,7 +23261,7 @@ if (lean_is_exclusive(x_149)) { lean_dec_ref(x_149); x_156 = lean_box(0); } -x_157 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_154); +x_157 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_154); if (x_157 == 0) { lean_object* x_158; @@ -23362,7 +23362,7 @@ if (lean_is_exclusive(x_162)) { lean_dec_ref(x_162); x_175 = lean_box(0); } -x_176 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_173); +x_176 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_173); if (x_176 == 0) { lean_object* x_177; @@ -23448,7 +23448,7 @@ if (lean_is_exclusive(x_160)) { lean_dec_ref(x_160); x_189 = lean_box(0); } -x_190 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_187); +x_190 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_187); if (x_190 == 0) { lean_object* x_191; @@ -23733,7 +23733,7 @@ if (x_12 == 0) lean_object* x_13; lean_object* x_14; uint8_t x_15; x_13 = lean_ctor_get(x_7, 0); x_14 = lean_ctor_get(x_7, 1); -x_15 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_13); +x_15 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_13); if (x_15 == 0) { lean_dec(x_2); @@ -23758,7 +23758,7 @@ x_17 = lean_ctor_get(x_7, 1); lean_inc(x_17); lean_inc(x_16); lean_dec(x_7); -x_18 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_16); +x_18 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_16); if (x_18 == 0) { lean_object* x_19; @@ -23791,7 +23791,7 @@ if (x_21 == 0) lean_object* x_22; lean_object* x_23; uint8_t x_24; x_22 = lean_ctor_get(x_3, 0); x_23 = lean_ctor_get(x_3, 1); -x_24 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_22); +x_24 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_22); if (x_24 == 0) { lean_dec(x_2); @@ -23816,7 +23816,7 @@ x_26 = lean_ctor_get(x_3, 1); lean_inc(x_26); lean_inc(x_25); lean_dec(x_3); -x_27 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_25); +x_27 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_25); if (x_27 == 0) { lean_object* x_28; @@ -23904,7 +23904,7 @@ if (x_27 == 0) lean_object* x_28; lean_object* x_29; uint8_t x_30; x_28 = lean_ctor_get(x_25, 0); x_29 = lean_ctor_get(x_25, 1); -x_30 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_24, x_28); +x_30 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_24, x_28); if (x_30 == 0) { lean_dec(x_24); @@ -23927,7 +23927,7 @@ x_32 = lean_ctor_get(x_25, 1); lean_inc(x_32); lean_inc(x_31); lean_dec(x_25); -x_33 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_24, x_31); +x_33 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_24, x_31); if (x_33 == 0) { lean_object* x_34; @@ -24386,7 +24386,7 @@ if (x_12 == 0) lean_object* x_13; lean_object* x_14; uint8_t x_15; x_13 = lean_ctor_get(x_7, 0); x_14 = lean_ctor_get(x_7, 1); -x_15 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_13); +x_15 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_13); if (x_15 == 0) { lean_dec(x_3); @@ -24411,7 +24411,7 @@ x_17 = lean_ctor_get(x_7, 1); lean_inc(x_17); lean_inc(x_16); lean_dec(x_7); -x_18 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_16); +x_18 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_16); if (x_18 == 0) { lean_object* x_19; @@ -24438,7 +24438,7 @@ return x_20; block_58: { uint8_t x_24; -x_24 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_22); +x_24 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_22); if (x_24 == 0) { lean_object* x_25; @@ -24476,7 +24476,7 @@ if (x_29 == 0) lean_object* x_30; lean_object* x_31; uint8_t x_32; x_30 = lean_ctor_get(x_26, 0); x_31 = lean_ctor_get(x_26, 1); -x_32 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_30); +x_32 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_30); if (x_32 == 0) { lean_dec(x_3); @@ -24512,7 +24512,7 @@ if (x_36 == 0) lean_object* x_37; lean_object* x_38; uint8_t x_39; x_37 = lean_ctor_get(x_33, 0); x_38 = lean_ctor_get(x_33, 1); -x_39 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_37); +x_39 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_37); if (x_39 == 0) { lean_dec(x_3); @@ -24537,7 +24537,7 @@ x_41 = lean_ctor_get(x_33, 1); lean_inc(x_41); lean_inc(x_40); lean_dec(x_33); -x_42 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_40); +x_42 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_40); if (x_42 == 0) { lean_object* x_43; @@ -24570,7 +24570,7 @@ x_46 = lean_ctor_get(x_26, 1); lean_inc(x_46); lean_inc(x_45); lean_dec(x_26); -x_47 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_45); +x_47 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_45); if (x_47 == 0) { lean_object* x_48; @@ -24615,7 +24615,7 @@ if (lean_is_exclusive(x_49)) { lean_dec_ref(x_49); x_54 = lean_box(0); } -x_55 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_52); +x_55 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_52); if (x_55 == 0) { lean_object* x_56; @@ -24769,7 +24769,7 @@ if (x_12 == 0) lean_object* x_13; lean_object* x_14; uint8_t x_15; x_13 = lean_ctor_get(x_7, 0); x_14 = lean_ctor_get(x_7, 1); -x_15 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_13); +x_15 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_13); if (x_15 == 0) { lean_dec(x_3); @@ -24794,7 +24794,7 @@ x_17 = lean_ctor_get(x_7, 1); lean_inc(x_17); lean_inc(x_16); lean_dec(x_7); -x_18 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_16); +x_18 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_16); if (x_18 == 0) { lean_object* x_19; @@ -24821,7 +24821,7 @@ return x_20; block_58: { uint8_t x_24; -x_24 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_22); +x_24 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_22); if (x_24 == 0) { lean_object* x_25; @@ -24859,7 +24859,7 @@ if (x_29 == 0) lean_object* x_30; lean_object* x_31; uint8_t x_32; x_30 = lean_ctor_get(x_26, 0); x_31 = lean_ctor_get(x_26, 1); -x_32 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_30); +x_32 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_30); if (x_32 == 0) { lean_dec(x_3); @@ -24895,7 +24895,7 @@ if (x_36 == 0) lean_object* x_37; lean_object* x_38; uint8_t x_39; x_37 = lean_ctor_get(x_33, 0); x_38 = lean_ctor_get(x_33, 1); -x_39 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_37); +x_39 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_37); if (x_39 == 0) { lean_dec(x_3); @@ -24920,7 +24920,7 @@ x_41 = lean_ctor_get(x_33, 1); lean_inc(x_41); lean_inc(x_40); lean_dec(x_33); -x_42 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_40); +x_42 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_40); if (x_42 == 0) { lean_object* x_43; @@ -24953,7 +24953,7 @@ x_46 = lean_ctor_get(x_26, 1); lean_inc(x_46); lean_inc(x_45); lean_dec(x_26); -x_47 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_45); +x_47 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_45); if (x_47 == 0) { lean_object* x_48; @@ -24998,7 +24998,7 @@ if (lean_is_exclusive(x_49)) { lean_dec_ref(x_49); x_54 = lean_box(0); } -x_55 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_52); +x_55 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_52); if (x_55 == 0) { lean_object* x_56; @@ -25183,7 +25183,7 @@ return x_16; block_56: { uint8_t x_20; -x_20 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_18); +x_20 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_18); if (x_20 == 0) { lean_object* x_21; @@ -25569,7 +25569,7 @@ if (x_9 == 0) lean_object* x_10; lean_object* x_11; uint8_t x_12; x_10 = lean_ctor_get(x_2, 0); x_11 = lean_ctor_get(x_2, 1); -x_12 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_10); +x_12 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_10); if (x_12 == 0) { lean_dec(x_1); @@ -25626,7 +25626,7 @@ if (x_21 == 0) lean_object* x_22; lean_object* x_23; uint8_t x_24; x_22 = lean_ctor_get(x_16, 0); x_23 = lean_ctor_get(x_16, 1); -x_24 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_14, x_22); +x_24 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_14, x_22); if (x_24 == 0) { lean_dec(x_15); @@ -25651,7 +25651,7 @@ x_26 = lean_ctor_get(x_16, 1); lean_inc(x_26); lean_inc(x_25); lean_dec(x_16); -x_27 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_14, x_25); +x_27 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_14, x_25); if (x_27 == 0) { lean_object* x_28; @@ -25707,7 +25707,7 @@ x_35 = lean_ctor_get(x_2, 1); lean_inc(x_35); lean_inc(x_34); lean_dec(x_2); -x_36 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_34); +x_36 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_34); if (x_36 == 0) { lean_object* x_37; @@ -25771,7 +25771,7 @@ if (lean_is_exclusive(x_41)) { lean_dec_ref(x_41); x_47 = lean_box(0); } -x_48 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_39, x_45); +x_48 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_39, x_45); if (x_48 == 0) { lean_object* x_49; @@ -25901,7 +25901,7 @@ if (x_28 == 0) lean_object* x_29; lean_object* x_30; uint8_t x_31; x_29 = lean_ctor_get(x_26, 0); x_30 = lean_ctor_get(x_26, 1); -x_31 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_25, x_29); +x_31 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_25, x_29); if (x_31 == 0) { lean_dec(x_25); @@ -25924,7 +25924,7 @@ x_33 = lean_ctor_get(x_26, 1); lean_inc(x_33); lean_inc(x_32); lean_dec(x_26); -x_34 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_25, x_32); +x_34 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_25, x_32); if (x_34 == 0) { lean_object* x_35; @@ -26145,7 +26145,7 @@ if (x_9 == 0) lean_object* x_10; lean_object* x_11; uint8_t x_12; x_10 = lean_ctor_get(x_2, 0); x_11 = lean_ctor_get(x_2, 1); -x_12 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_10); +x_12 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_10); if (x_12 == 0) { lean_dec(x_1); @@ -26169,7 +26169,7 @@ x_15 = lean_ctor_get(x_2, 1); lean_inc(x_15); lean_inc(x_14); lean_dec(x_2); -x_16 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_14); +x_16 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_14); if (x_16 == 0) { lean_object* x_17; @@ -26287,7 +26287,7 @@ if (x_37 == 0) lean_object* x_38; lean_object* x_39; uint8_t x_40; x_38 = lean_ctor_get(x_35, 0); x_39 = lean_ctor_get(x_35, 1); -x_40 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_34, x_38); +x_40 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_34, x_38); if (x_40 == 0) { lean_dec(x_34); @@ -26310,7 +26310,7 @@ x_42 = lean_ctor_get(x_35, 1); lean_inc(x_42); lean_inc(x_41); lean_dec(x_35); -x_43 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_34, x_41); +x_43 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_34, x_41); if (x_43 == 0) { lean_object* x_44; @@ -26487,7 +26487,7 @@ if (lean_is_exclusive(x_73)) { lean_dec_ref(x_73); x_77 = lean_box(0); } -x_78 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_72, x_75); +x_78 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_72, x_75); if (x_78 == 0) { lean_object* x_79; @@ -26717,7 +26717,7 @@ if (x_7 == 0) lean_object* x_8; lean_object* x_9; uint8_t x_10; x_8 = lean_ctor_get(x_2, 0); x_9 = lean_ctor_get(x_2, 1); -x_10 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_8); +x_10 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_8); if (x_10 == 0) { lean_dec(x_1); @@ -26741,7 +26741,7 @@ x_13 = lean_ctor_get(x_2, 1); lean_inc(x_13); lean_inc(x_12); lean_dec(x_2); -x_14 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_12); +x_14 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_12); if (x_14 == 0) { lean_object* x_15; @@ -26934,7 +26934,7 @@ if (x_36 == 0) lean_object* x_37; lean_object* x_38; uint8_t x_39; x_37 = lean_ctor_get(x_34, 0); x_38 = lean_ctor_get(x_34, 1); -x_39 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_33, x_37); +x_39 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_33, x_37); if (x_39 == 0) { lean_dec(x_33); @@ -26988,7 +26988,7 @@ x_47 = lean_ctor_get(x_34, 1); lean_inc(x_47); lean_inc(x_46); lean_dec(x_34); -x_48 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_33, x_46); +x_48 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_33, x_46); if (x_48 == 0) { lean_object* x_49; @@ -27176,7 +27176,7 @@ if (x_18 == 0) lean_object* x_19; lean_object* x_20; uint8_t x_21; x_19 = lean_ctor_get(x_16, 0); x_20 = lean_ctor_get(x_16, 1); -x_21 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_15, x_19); +x_21 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_15, x_19); if (x_21 == 0) { lean_dec(x_15); @@ -27199,7 +27199,7 @@ x_23 = lean_ctor_get(x_16, 1); lean_inc(x_23); lean_inc(x_22); lean_dec(x_16); -x_24 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_15, x_22); +x_24 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_15, x_22); if (x_24 == 0) { lean_object* x_25; @@ -27259,7 +27259,7 @@ if (x_7 == 0) lean_object* x_8; lean_object* x_9; uint8_t x_10; x_8 = lean_ctor_get(x_2, 0); x_9 = lean_ctor_get(x_2, 1); -x_10 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_8); +x_10 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_8); if (x_10 == 0) { lean_dec(x_1); @@ -27305,7 +27305,7 @@ if (x_16 == 0) lean_object* x_17; lean_object* x_18; uint8_t x_19; x_17 = lean_ctor_get(x_11, 0); x_18 = lean_ctor_get(x_11, 1); -x_19 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_17); +x_19 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_17); if (x_19 == 0) { lean_dec(x_1); @@ -27351,7 +27351,7 @@ if (x_25 == 0) lean_object* x_26; lean_object* x_27; uint8_t x_28; x_26 = lean_ctor_get(x_20, 0); x_27 = lean_ctor_get(x_20, 1); -x_28 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_26); +x_28 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_26); if (x_28 == 0) { lean_dec(x_1); @@ -27397,7 +27397,7 @@ if (x_34 == 0) lean_object* x_35; lean_object* x_36; uint8_t x_37; x_35 = lean_ctor_get(x_29, 0); x_36 = lean_ctor_get(x_29, 1); -x_37 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_35); +x_37 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_35); if (x_37 == 0) { lean_dec(x_1); @@ -27443,7 +27443,7 @@ if (x_43 == 0) lean_object* x_44; lean_object* x_45; uint8_t x_46; x_44 = lean_ctor_get(x_38, 0); x_45 = lean_ctor_get(x_38, 1); -x_46 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_44); +x_46 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_44); if (x_46 == 0) { lean_dec(x_1); @@ -27514,7 +27514,7 @@ x_59 = lean_ctor_get(x_38, 1); lean_inc(x_59); lean_inc(x_58); lean_dec(x_38); -x_60 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_58); +x_60 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_58); if (x_60 == 0) { lean_object* x_61; @@ -27590,7 +27590,7 @@ x_72 = lean_ctor_get(x_29, 1); lean_inc(x_72); lean_inc(x_71); lean_dec(x_29); -x_73 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_71); +x_73 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_71); if (x_73 == 0) { lean_object* x_74; @@ -27647,7 +27647,7 @@ if (lean_is_exclusive(x_75)) { lean_dec_ref(x_75); x_82 = lean_box(0); } -x_83 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_80); +x_83 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_80); if (x_83 == 0) { lean_object* x_84; @@ -27730,7 +27730,7 @@ x_95 = lean_ctor_get(x_20, 1); lean_inc(x_95); lean_inc(x_94); lean_dec(x_20); -x_96 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_94); +x_96 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_94); if (x_96 == 0) { lean_object* x_97; @@ -27787,7 +27787,7 @@ if (lean_is_exclusive(x_98)) { lean_dec_ref(x_98); x_105 = lean_box(0); } -x_106 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_103); +x_106 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_103); if (x_106 == 0) { lean_object* x_107; @@ -27849,7 +27849,7 @@ if (lean_is_exclusive(x_108)) { lean_dec_ref(x_108); x_115 = lean_box(0); } -x_116 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_113); +x_116 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_113); if (x_116 == 0) { lean_object* x_117; @@ -27934,7 +27934,7 @@ x_128 = lean_ctor_get(x_11, 1); lean_inc(x_128); lean_inc(x_127); lean_dec(x_11); -x_129 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_127); +x_129 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_127); if (x_129 == 0) { lean_object* x_130; @@ -27991,7 +27991,7 @@ if (lean_is_exclusive(x_131)) { lean_dec_ref(x_131); x_138 = lean_box(0); } -x_139 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_136); +x_139 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_136); if (x_139 == 0) { lean_object* x_140; @@ -28053,7 +28053,7 @@ if (lean_is_exclusive(x_141)) { lean_dec_ref(x_141); x_148 = lean_box(0); } -x_149 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_146); +x_149 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_146); if (x_149 == 0) { lean_object* x_150; @@ -28115,7 +28115,7 @@ if (lean_is_exclusive(x_151)) { lean_dec_ref(x_151); x_158 = lean_box(0); } -x_159 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_156); +x_159 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_156); if (x_159 == 0) { lean_object* x_160; @@ -28202,7 +28202,7 @@ x_171 = lean_ctor_get(x_2, 1); lean_inc(x_171); lean_inc(x_170); lean_dec(x_2); -x_172 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_170); +x_172 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_170); if (x_172 == 0) { lean_object* x_173; @@ -28259,7 +28259,7 @@ if (lean_is_exclusive(x_174)) { lean_dec_ref(x_174); x_181 = lean_box(0); } -x_182 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_179); +x_182 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_179); if (x_182 == 0) { lean_object* x_183; @@ -28321,7 +28321,7 @@ if (lean_is_exclusive(x_184)) { lean_dec_ref(x_184); x_191 = lean_box(0); } -x_192 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_189); +x_192 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_189); if (x_192 == 0) { lean_object* x_193; @@ -28383,7 +28383,7 @@ if (lean_is_exclusive(x_194)) { lean_dec_ref(x_194); x_201 = lean_box(0); } -x_202 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_199); +x_202 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_199); if (x_202 == 0) { lean_object* x_203; @@ -28445,7 +28445,7 @@ if (lean_is_exclusive(x_204)) { lean_dec_ref(x_204); x_211 = lean_box(0); } -x_212 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_209); +x_212 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_209); if (x_212 == 0) { lean_object* x_213; @@ -28564,7 +28564,7 @@ if (x_7 == 0) lean_object* x_8; lean_object* x_9; uint8_t x_10; x_8 = lean_ctor_get(x_2, 0); x_9 = lean_ctor_get(x_2, 1); -x_10 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_8); +x_10 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_8); if (x_10 == 0) { lean_dec(x_1); @@ -28635,7 +28635,7 @@ x_23 = lean_ctor_get(x_2, 1); lean_inc(x_23); lean_inc(x_22); lean_dec(x_2); -x_24 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_22); +x_24 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_22); if (x_24 == 0) { lean_object* x_25; @@ -28753,7 +28753,7 @@ if (x_12 == 0) lean_object* x_13; lean_object* x_14; uint8_t x_15; x_13 = lean_ctor_get(x_7, 0); x_14 = lean_ctor_get(x_7, 1); -x_15 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_13); +x_15 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_13); if (x_15 == 0) { lean_dec(x_2); @@ -28778,7 +28778,7 @@ x_17 = lean_ctor_get(x_7, 1); lean_inc(x_17); lean_inc(x_16); lean_dec(x_7); -x_18 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_16); +x_18 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_16); if (x_18 == 0) { lean_object* x_19; @@ -28811,7 +28811,7 @@ if (x_21 == 0) lean_object* x_22; lean_object* x_23; uint8_t x_24; x_22 = lean_ctor_get(x_3, 0); x_23 = lean_ctor_get(x_3, 1); -x_24 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_22); +x_24 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_22); if (x_24 == 0) { lean_dec(x_2); @@ -28870,7 +28870,7 @@ if (x_34 == 0) lean_object* x_35; lean_object* x_36; uint8_t x_37; x_35 = lean_ctor_get(x_29, 0); x_36 = lean_ctor_get(x_29, 1); -x_37 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_35); +x_37 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_35); if (x_37 == 0) { lean_dec(x_2); @@ -28895,7 +28895,7 @@ x_39 = lean_ctor_get(x_29, 1); lean_inc(x_39); lean_inc(x_38); lean_dec(x_29); -x_40 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_38); +x_40 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_38); if (x_40 == 0) { lean_object* x_41; @@ -28928,7 +28928,7 @@ if (x_43 == 0) lean_object* x_44; lean_object* x_45; uint8_t x_46; x_44 = lean_ctor_get(x_25, 0); x_45 = lean_ctor_get(x_25, 1); -x_46 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_44); +x_46 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_44); if (x_46 == 0) { lean_dec(x_2); @@ -28953,7 +28953,7 @@ x_48 = lean_ctor_get(x_25, 1); lean_inc(x_48); lean_inc(x_47); lean_dec(x_25); -x_49 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_47); +x_49 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_47); if (x_49 == 0) { lean_object* x_50; @@ -28986,7 +28986,7 @@ x_53 = lean_ctor_get(x_3, 1); lean_inc(x_53); lean_inc(x_52); lean_dec(x_3); -x_54 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_52); +x_54 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_52); if (x_54 == 0) { lean_object* x_55; @@ -29056,7 +29056,7 @@ if (lean_is_exclusive(x_60)) { lean_dec_ref(x_60); x_67 = lean_box(0); } -x_68 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_65); +x_68 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_65); if (x_68 == 0) { lean_object* x_69; @@ -29103,7 +29103,7 @@ if (lean_is_exclusive(x_56)) { lean_dec_ref(x_56); x_73 = lean_box(0); } -x_74 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_71); +x_74 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_71); if (x_74 == 0) { lean_object* x_75; @@ -29317,7 +29317,7 @@ if (x_104 == 0) lean_object* x_105; lean_object* x_106; uint8_t x_107; x_105 = lean_ctor_get(x_102, 0); x_106 = lean_ctor_get(x_102, 1); -x_107 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_99, x_105); +x_107 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_99, x_105); if (x_107 == 0) { lean_dec(x_99); @@ -29340,7 +29340,7 @@ x_109 = lean_ctor_get(x_102, 1); lean_inc(x_109); lean_inc(x_108); lean_dec(x_102); -x_110 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_99, x_108); +x_110 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_99, x_108); if (x_110 == 0) { lean_object* x_111; @@ -29369,7 +29369,7 @@ if (x_112 == 0) lean_object* x_113; lean_object* x_114; uint8_t x_115; x_113 = lean_ctor_get(x_100, 0); x_114 = lean_ctor_get(x_100, 1); -x_115 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_99, x_113); +x_115 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_99, x_113); if (x_115 == 0) { lean_dec(x_99); @@ -29392,7 +29392,7 @@ x_117 = lean_ctor_get(x_100, 1); lean_inc(x_117); lean_inc(x_116); lean_dec(x_100); -x_118 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_99, x_116); +x_118 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_99, x_116); if (x_118 == 0) { lean_object* x_119; @@ -29531,7 +29531,7 @@ x_17 = l_String_Iterator_hasNext(x_15); if (x_17 == 0) { uint8_t x_18; -x_18 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_15, x_15); +x_18 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_15, x_15); if (x_18 == 0) { lean_object* x_19; lean_object* x_20; @@ -29558,7 +29558,7 @@ if (x_23 == 0) { uint8_t x_24; lean_dec(x_21); -x_24 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_15, x_15); +x_24 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_15, x_15); if (x_24 == 0) { lean_object* x_25; lean_object* x_26; @@ -29593,7 +29593,7 @@ x_32 = l_String_Iterator_hasNext(x_29); if (x_32 == 0) { uint8_t x_33; -x_33 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_15, x_29); +x_33 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_15, x_29); if (x_33 == 0) { lean_object* x_34; @@ -29622,7 +29622,7 @@ if (x_37 == 0) { uint8_t x_38; lean_dec(x_35); -x_38 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_15, x_29); +x_38 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_15, x_29); if (x_38 == 0) { lean_object* x_39; @@ -29667,12 +29667,12 @@ if (x_42 == 0) lean_object* x_43; lean_object* x_44; uint8_t x_45; x_43 = lean_ctor_get(x_40, 0); x_44 = lean_ctor_get(x_40, 1); -x_45 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_35, x_43); +x_45 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_35, x_43); if (x_45 == 0) { uint8_t x_46; lean_dec(x_35); -x_46 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_15, x_43); +x_46 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_15, x_43); if (x_46 == 0) { lean_dec(x_15); @@ -29705,12 +29705,12 @@ x_48 = lean_ctor_get(x_40, 1); lean_inc(x_48); lean_inc(x_47); lean_dec(x_40); -x_49 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_35, x_47); +x_49 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_35, x_47); if (x_49 == 0) { uint8_t x_50; lean_dec(x_35); -x_50 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_15, x_47); +x_50 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_15, x_47); if (x_50 == 0) { lean_object* x_51; @@ -29752,7 +29752,7 @@ x_54 = l_String_Iterator_hasNext(x_52); if (x_54 == 0) { uint8_t x_55; -x_55 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_15, x_52); +x_55 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_15, x_52); if (x_55 == 0) { lean_object* x_56; lean_object* x_57; @@ -29781,7 +29781,7 @@ if (x_60 == 0) { uint8_t x_61; lean_dec(x_58); -x_61 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_15, x_52); +x_61 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_15, x_52); if (x_61 == 0) { lean_object* x_62; lean_object* x_63; @@ -29831,12 +29831,12 @@ if (lean_is_exclusive(x_64)) { lean_dec_ref(x_64); x_68 = lean_box(0); } -x_69 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_58, x_66); +x_69 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_58, x_66); if (x_69 == 0) { uint8_t x_70; lean_dec(x_58); -x_70 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_15, x_66); +x_70 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_15, x_66); if (x_70 == 0) { lean_object* x_71; @@ -29882,7 +29882,7 @@ if (x_72 == 0) lean_object* x_73; lean_object* x_74; uint8_t x_75; x_73 = lean_ctor_get(x_27, 0); x_74 = lean_ctor_get(x_27, 1); -x_75 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_15, x_73); +x_75 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_15, x_73); if (x_75 == 0) { lean_dec(x_15); @@ -29905,7 +29905,7 @@ x_77 = lean_ctor_get(x_27, 1); lean_inc(x_77); lean_inc(x_76); lean_dec(x_27); -x_78 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_15, x_76); +x_78 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_15, x_76); if (x_78 == 0) { lean_object* x_79; @@ -29951,7 +29951,7 @@ if (x_84 == 0) lean_object* x_85; lean_object* x_86; uint8_t x_87; x_85 = lean_ctor_get(x_82, 0); x_86 = lean_ctor_get(x_82, 1); -x_87 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_81, x_85); +x_87 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_81, x_85); if (x_87 == 0) { lean_dec(x_81); @@ -29974,7 +29974,7 @@ x_89 = lean_ctor_get(x_82, 1); lean_inc(x_89); lean_inc(x_88); lean_dec(x_82); -x_90 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_81, x_88); +x_90 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_81, x_88); if (x_90 == 0) { lean_object* x_91; @@ -30046,7 +30046,7 @@ if (x_12 == 0) lean_object* x_13; lean_object* x_14; uint8_t x_15; x_13 = lean_ctor_get(x_7, 0); x_14 = lean_ctor_get(x_7, 1); -x_15 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_13); +x_15 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_13); if (x_15 == 0) { lean_dec(x_2); @@ -30071,7 +30071,7 @@ x_17 = lean_ctor_get(x_7, 1); lean_inc(x_17); lean_inc(x_16); lean_dec(x_7); -x_18 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_16); +x_18 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_16); if (x_18 == 0) { lean_object* x_19; @@ -30104,7 +30104,7 @@ if (x_21 == 0) lean_object* x_22; lean_object* x_23; uint8_t x_24; x_22 = lean_ctor_get(x_3, 0); x_23 = lean_ctor_get(x_3, 1); -x_24 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_22); +x_24 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_22); if (x_24 == 0) { lean_dec(x_2); @@ -30129,7 +30129,7 @@ x_26 = lean_ctor_get(x_3, 1); lean_inc(x_26); lean_inc(x_25); lean_dec(x_3); -x_27 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_25); +x_27 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_25); if (x_27 == 0) { lean_object* x_28; @@ -30179,7 +30179,7 @@ if (x_44 == 0) lean_object* x_45; lean_object* x_46; uint8_t x_47; x_45 = lean_ctor_get(x_42, 0); x_46 = lean_ctor_get(x_42, 1); -x_47 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_45); +x_47 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_45); if (x_47 == 0) { lean_dec(x_1); @@ -30202,7 +30202,7 @@ x_49 = lean_ctor_get(x_42, 1); lean_inc(x_49); lean_inc(x_48); lean_dec(x_42); -x_50 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_48); +x_50 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_48); if (x_50 == 0) { lean_object* x_51; @@ -30277,7 +30277,7 @@ if (x_15 == 0) lean_object* x_16; lean_object* x_17; uint8_t x_18; x_16 = lean_ctor_get(x_8, 0); x_17 = lean_ctor_get(x_8, 1); -x_18 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_5, x_16); +x_18 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_5, x_16); if (x_18 == 0) { lean_dec(x_5); @@ -30303,7 +30303,7 @@ x_21 = lean_ctor_get(x_8, 1); lean_inc(x_21); lean_inc(x_20); lean_dec(x_8); -x_22 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_5, x_20); +x_22 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_5, x_20); if (x_22 == 0) { lean_object* x_23; @@ -30336,7 +30336,7 @@ if (x_26 == 0) lean_object* x_27; lean_object* x_28; uint8_t x_29; x_27 = lean_ctor_get(x_6, 0); x_28 = lean_ctor_get(x_6, 1); -x_29 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_5, x_27); +x_29 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_5, x_27); if (x_29 == 0) { lean_dec(x_5); @@ -30362,7 +30362,7 @@ x_32 = lean_ctor_get(x_6, 1); lean_inc(x_32); lean_inc(x_31); lean_dec(x_6); -x_33 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_5, x_31); +x_33 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_5, x_31); if (x_33 == 0) { lean_object* x_34; @@ -33070,7 +33070,7 @@ if (x_14 == 0) lean_object* x_15; lean_object* x_16; uint8_t x_17; x_15 = lean_ctor_get(x_9, 0); x_16 = lean_ctor_get(x_9, 1); -x_17 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_15); +x_17 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_15); if (x_17 == 0) { lean_dec(x_2); @@ -33095,7 +33095,7 @@ x_19 = lean_ctor_get(x_9, 1); lean_inc(x_19); lean_inc(x_18); lean_dec(x_9); -x_20 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_18); +x_20 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_18); if (x_20 == 0) { lean_object* x_21; @@ -33128,7 +33128,7 @@ if (x_23 == 0) lean_object* x_24; lean_object* x_25; uint8_t x_26; x_24 = lean_ctor_get(x_5, 0); x_25 = lean_ctor_get(x_5, 1); -x_26 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_24); +x_26 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_24); if (x_26 == 0) { lean_dec(x_2); @@ -33153,7 +33153,7 @@ x_28 = lean_ctor_get(x_5, 1); lean_inc(x_28); lean_inc(x_27); lean_dec(x_5); -x_29 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_27); +x_29 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_27); if (x_29 == 0) { lean_object* x_30; @@ -33186,7 +33186,7 @@ if (x_32 == 0) lean_object* x_33; lean_object* x_34; uint8_t x_35; x_33 = lean_ctor_get(x_3, 0); x_34 = lean_ctor_get(x_3, 1); -x_35 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_33); +x_35 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_33); if (x_35 == 0) { lean_dec(x_2); @@ -33211,7 +33211,7 @@ x_37 = lean_ctor_get(x_3, 1); lean_inc(x_37); lean_inc(x_36); lean_dec(x_3); -x_38 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_36); +x_38 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_36); if (x_38 == 0) { lean_object* x_39; @@ -33363,7 +33363,7 @@ if (x_25 == 0) lean_object* x_26; lean_object* x_27; uint8_t x_28; x_26 = lean_ctor_get(x_18, 0); x_27 = lean_ctor_get(x_18, 1); -x_28 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_16, x_26); +x_28 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_16, x_26); if (x_28 == 0) { lean_dec(x_17); @@ -33393,7 +33393,7 @@ x_31 = lean_ctor_get(x_18, 1); lean_inc(x_31); lean_inc(x_30); lean_dec(x_18); -x_32 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_16, x_30); +x_32 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_16, x_30); if (x_32 == 0) { lean_object* x_33; @@ -33641,7 +33641,7 @@ if (x_22 == 0) lean_object* x_23; lean_object* x_24; uint8_t x_25; x_23 = lean_ctor_get(x_20, 0); x_24 = lean_ctor_get(x_20, 1); -x_25 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_19, x_23); +x_25 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_19, x_23); if (x_25 == 0) { lean_dec(x_19); @@ -33664,7 +33664,7 @@ x_27 = lean_ctor_get(x_20, 1); lean_inc(x_27); lean_inc(x_26); lean_dec(x_20); -x_28 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_19, x_26); +x_28 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_19, x_26); if (x_28 == 0) { lean_object* x_29; @@ -33920,7 +33920,7 @@ x_6 = lean_ctor_get(x_4, 1); lean_dec(x_6); x_7 = lean_ctor_get(x_4, 0); lean_dec(x_7); -x_8 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_2); +x_8 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_2); if (x_8 == 0) { lean_object* x_9; @@ -33942,7 +33942,7 @@ else { uint8_t x_10; lean_dec(x_4); -x_10 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_2); +x_10 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_2); if (x_10 == 0) { lean_object* x_11; lean_object* x_12; @@ -33978,7 +33978,7 @@ x_17 = l_String_Iterator_hasNext(x_2); if (x_17 == 0) { uint8_t x_18; -x_18 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_2); +x_18 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_2); if (x_18 == 0) { lean_object* x_19; @@ -34039,7 +34039,7 @@ if (x_28 == 0) lean_object* x_29; lean_object* x_30; uint8_t x_31; x_29 = lean_ctor_get(x_23, 0); x_30 = lean_ctor_get(x_23, 1); -x_31 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_29); +x_31 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_29); if (x_31 == 0) { lean_dec(x_2); @@ -34064,7 +34064,7 @@ x_33 = lean_ctor_get(x_23, 1); lean_inc(x_33); lean_inc(x_32); lean_dec(x_23); -x_34 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_32); +x_34 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_32); if (x_34 == 0) { lean_object* x_35; @@ -34097,7 +34097,7 @@ x_37 = l_String_Iterator_hasNext(x_2); if (x_37 == 0) { uint8_t x_38; -x_38 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_2); +x_38 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_2); if (x_38 == 0) { lean_object* x_39; lean_object* x_40; @@ -34167,7 +34167,7 @@ if (lean_is_exclusive(x_45)) { lean_dec_ref(x_45); x_52 = lean_box(0); } -x_53 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_50); +x_53 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_50); if (x_53 == 0) { lean_object* x_54; @@ -34337,7 +34337,7 @@ x_3 = l_String_Iterator_hasNext(x_2); if (x_3 == 0) { uint8_t x_4; -x_4 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_2); +x_4 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_2); if (x_4 == 0) { lean_object* x_5; lean_object* x_6; @@ -34370,7 +34370,7 @@ if (x_12 == 0) { uint8_t x_13; lean_dec(x_8); -x_13 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_2); +x_13 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_2); if (x_13 == 0) { lean_object* x_14; lean_object* x_15; @@ -34400,7 +34400,7 @@ if (x_19 == 0) { uint8_t x_20; lean_dec(x_8); -x_20 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_2); +x_20 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_2); if (x_20 == 0) { lean_object* x_21; lean_object* x_22; @@ -34459,7 +34459,7 @@ if (x_30 == 0) lean_object* x_31; lean_object* x_32; uint8_t x_33; x_31 = lean_ctor_get(x_25, 0); x_32 = lean_ctor_get(x_25, 1); -x_33 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_31); +x_33 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_31); if (x_33 == 0) { lean_dec(x_2); @@ -34484,7 +34484,7 @@ x_35 = lean_ctor_get(x_25, 1); lean_inc(x_35); lean_inc(x_34); lean_dec(x_25); -x_36 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_34); +x_36 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_34); if (x_36 == 0) { lean_object* x_37; @@ -34594,7 +34594,7 @@ lean_object* x_50; uint8_t x_51; x_50 = lean_ctor_get(x_45, 1); lean_inc(x_50); lean_dec(x_45); -x_51 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_2); +x_51 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_2); if (x_51 == 0) { lean_inc(x_2); @@ -34674,7 +34674,7 @@ lean_inc(x_69); x_70 = lean_ctor_get(x_52, 1); lean_inc(x_70); lean_dec(x_52); -x_71 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_69); +x_71 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_69); if (x_71 == 0) { x_21 = x_69; @@ -34712,7 +34712,7 @@ lean_inc(x_77); x_78 = lean_ctor_get(x_72, 1); lean_inc(x_78); lean_dec(x_72); -x_79 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_77); +x_79 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_77); if (x_79 == 0) { x_21 = x_77; @@ -34745,7 +34745,7 @@ lean_inc(x_83); x_84 = lean_ctor_get(x_80, 1); lean_inc(x_84); lean_dec(x_80); -x_85 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_83); +x_85 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_83); if (x_85 == 0) { x_21 = x_83; @@ -34834,7 +34834,7 @@ if (x_11 == 0) lean_object* x_12; lean_object* x_13; uint8_t x_14; x_12 = lean_ctor_get(x_6, 0); x_13 = lean_ctor_get(x_6, 1); -x_14 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_12); +x_14 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_12); if (x_14 == 0) { lean_dec(x_2); @@ -34859,7 +34859,7 @@ x_16 = lean_ctor_get(x_6, 1); lean_inc(x_16); lean_inc(x_15); lean_dec(x_6); -x_17 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_15); +x_17 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_15); if (x_17 == 0) { lean_object* x_18; @@ -34886,7 +34886,7 @@ return x_19; block_26: { uint8_t x_23; -x_23 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_2, x_21); +x_23 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_2, x_21); if (x_23 == 0) { lean_object* x_24; @@ -34941,7 +34941,7 @@ lean_inc(x_37); x_38 = lean_ctor_get(x_29, 1); lean_inc(x_38); lean_dec(x_29); -x_39 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_27, x_37); +x_39 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_27, x_37); if (x_39 == 0) { lean_dec(x_28); @@ -35223,7 +35223,7 @@ if (x_57 == 0) lean_object* x_58; lean_object* x_59; uint8_t x_60; x_58 = lean_ctor_get(x_52, 0); x_59 = lean_ctor_get(x_52, 1); -x_60 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_58); +x_60 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_58); if (x_60 == 0) { lean_dec(x_1); @@ -35249,7 +35249,7 @@ x_63 = lean_ctor_get(x_52, 1); lean_inc(x_63); lean_inc(x_62); lean_dec(x_52); -x_64 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_1, x_62); +x_64 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_1, x_62); if (x_64 == 0) { lean_object* x_65; @@ -35472,7 +35472,7 @@ if (x_10 == 0) lean_object* x_11; lean_object* x_12; uint8_t x_13; x_11 = lean_ctor_get(x_5, 0); x_12 = lean_ctor_get(x_5, 1); -x_13 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_11); +x_13 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_11); if (x_13 == 0) { lean_dec(x_4); @@ -35609,7 +35609,7 @@ x_39 = lean_ctor_get(x_5, 1); lean_inc(x_39); lean_inc(x_38); lean_dec(x_5); -x_40 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2167_(x_3, x_38); +x_40 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_2169_(x_3, x_38); if (x_40 == 0) { lean_object* x_41; diff --git a/stage0/stdlib/Lean/Elab/App.c b/stage0/stdlib/Lean/Elab/App.c index 38578fb1a382..ccc7a1fc2ff8 100644 --- a/stage0/stdlib/Lean/Elab/App.c +++ b/stage0/stdlib/Lean/Elab/App.c @@ -586,7 +586,6 @@ lean_object* lean_array_to_list(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabPipeProj_declRange(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit___at_Lean_Elab_Term_ElabElim_revertArgs___spec__2___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*, lean_object*); 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*); -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(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_isNextOutParamOfLocalInstanceAndResult_hasLocalInstaceWithOutParams___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Elab_Term_termElabAttribute; uint8_t l_Lean_Expr_hasLooseBVars(lean_object*); @@ -891,6 +890,7 @@ static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_prop LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabChoice_declRange(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabNamedPattern_declRange___closed__6; static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__18; +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_anyNamedArgDependsOnCurrent___spec__3___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabAppArgs___lambda__5___closed__4; LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_propagateExpectedType___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -15709,7 +15709,7 @@ size_t x_11; size_t x_12; lean_object* x_13; lean_object* x_14; lean_object* x_1 x_11 = 1; x_12 = lean_usize_sub(x_2, x_11); x_13 = lean_array_uget(x_1, x_12); -x_14 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_13, x_5, x_6, x_7, x_8, x_9); +x_14 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_13, x_5, x_6, x_7, x_8, x_9); x_15 = lean_ctor_get(x_14, 0); lean_inc(x_15); x_16 = lean_ctor_get(x_14, 1); @@ -15743,7 +15743,7 @@ lean_inc(x_22); x_23 = lean_ctor_get(x_21, 1); lean_inc(x_23); lean_dec(x_21); -x_24 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_22, x_5, x_6, x_7, x_8, x_23); +x_24 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_22, x_5, x_6, x_7, x_8, x_23); x_25 = lean_ctor_get(x_24, 0); lean_inc(x_25); x_26 = lean_ctor_get(x_24, 1); diff --git a/stage0/stdlib/Lean/Elab/Binders.c b/stage0/stdlib/Lean/Elab/Binders.c index 3bdf59947527..26115fb3c40c 100644 --- a/stage0/stdlib/Lean/Elab/Binders.c +++ b/stage0/stdlib/Lean/Elab/Binders.c @@ -15,8 +15,8 @@ extern "C" { #endif LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_elabFun___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); -uint8_t l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_expandExplicitFun(lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__5; static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_ensureAtomicBinderName___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetDecl___closed__2; lean_object* l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -61,6 +61,7 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetDelayedDecl___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Term_expandForall___closed__1; 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_12252____closed__6; static lean_object* l_Lean_Elab_Term_elabLetDeclAux___closed__7; static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__6; static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_checkLocalInstanceParameters___closed__2; @@ -93,7 +94,9 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabForall_declRange(lean LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLetDeclCore___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___regBuiltin_Lean_Elab_Term_precheckArrow___closed__1; static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__52; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__9; 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_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__4; static lean_object* l_Lean_Elab_Term_elabLetDeclAux___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBinderViews_loop___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*); static lean_object* l_Lean_Elab_Term_expandWhereDecls___closed__3; @@ -114,6 +117,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabBinder___rarg___lambda__1___boxed( lean_object* l_Lean_mkLetFunAnnotation(lean_object*); static lean_object* l_Lean_Elab_Term_expandWhereDecls___closed__10; LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBinderViews___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_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__15; LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_ensureAtomicBinderName(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_elabLetDelayedDecl_declRange(lean_object*); static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__12; @@ -133,9 +137,10 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabLetDecl(lean_object*) 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*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabArrow_declRange___closed__7; lean_object* l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(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_Binders___hyg_12252____closed__10; lean_object* lean_array_fget(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabArrow___closed__2; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__11; +LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLetDeclCore___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*); static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_ensureAtomicBinderName___closed__4; static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_registerFailedToInferBinderTypeInfo___closed__3; static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_2400____closed__6; @@ -151,7 +156,6 @@ LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at_Lean_Elab_Term_expandWher LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabBinders___spec__1(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabLetDeclAux___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___regBuiltin_Lean_Elab_Term_elabFun___closed__2; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__9; LEAN_EXPORT lean_object* l_Lean_Elab_Term_quoteAutoTactic(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandSimpleBinderWithType(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_precheckFun___spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); @@ -162,7 +166,6 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabLetDec lean_object* l_Lean_addBuiltinDeclarationRanges(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkIdentFrom(lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandFunBinders_loop___spec__3(lean_object*, size_t, size_t, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__12; static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__32; static lean_object* l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__1; static lean_object* l_Lean_Elab_Term_elabLetDeclCore___closed__3; @@ -174,6 +177,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_expandForall(lean_object* static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_toBinderViews___spec__1___rarg___closed__1; lean_object* l_Lean_Syntax_node5(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__1___closed__11; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__13; 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*); uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_precheckFun___spec__1___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -188,6 +192,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLetDeclAux___lambda__3___boxed(lea LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderType___boxed(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabFun___closed__1; static lean_object* l_Lean_Elab_Term_expandSimpleBinderWithType___closed__3; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_elabFun___closed__3; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandFun___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); uint8_t lean_string_dec_eq(lean_object*, lean_object*); @@ -201,8 +206,6 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders lean_object* l_Lean_Option_register___at_Lean_Elab_initFn____x40_Lean_Elab_AutoBound___hyg_7____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_expandWhereDecls___closed__6; static lean_object* l___regBuiltin_Lean_Elab_Term_expandExplicitFun_declRange___closed__4; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__2; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__10; static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetDelayedDecl_declRange___closed__6; static lean_object* l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*); @@ -239,7 +242,6 @@ static lean_object* l_Lean_Elab_Term_expandFunBinders_loop___closed__6; 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*); static lean_object* l___regBuiltin_Lean_Elab_Term_expandExplicitFun_declRange___closed__7; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__10; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__5; lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkFreshBinderName___at_Lean_Elab_Term_elabLetDeclCore___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -267,7 +269,6 @@ static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_2400_ lean_object* l_Lean_Macro_resolveGlobalName(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_expandFunBinders_loop___closed__13; static lean_object* l___regBuiltin_Lean_Elab_Term_elabForall___closed__4; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__3; static lean_object* l_Lean_Elab_Term_expandWhereDecls___lambda__1___closed__1; static lean_object* l_Lean_Elab_Term_expandLetEqnsDecl___closed__1; static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__5; @@ -296,7 +297,6 @@ static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchA lean_object* l_Lean_ResolveName_resolveNamespace(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__54; LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBinderViews_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_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__15; lean_object* l_Array_unzip___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_toBinderViews___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_SourceInfo_fromRef(lean_object*, uint8_t); @@ -304,7 +304,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBin LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_elabForall___spec__1___rarg(lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabDepArrow_declRange(lean_object*); static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__45; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__7; lean_object* l_Lean_Syntax_node6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_expandExplicitFun_declRange___closed__5; static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___closed__5; @@ -351,13 +350,13 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabForall___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandMatchAltsWhereDecls_loop___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_toBinderViews___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at_Lean_Elab_Term_elabLetDeclAux___spec__2___rarg___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_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__13; static lean_object* l_Lean_Elab_Term_mkFreshBinderName___at_Lean_Elab_Term_elabLetDeclCore___spec__2___rarg___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_expandExplicitFun___closed__2; static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__1; lean_object* l_Lean_Expr_forallE___override(lean_object*, lean_object*, lean_object*, uint8_t); static lean_object* l_Lean_Elab_Term_expandFunBinders_loop___closed__4; lean_object* l_ReaderT_instMonadLiftReaderT(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__8; 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*); lean_object* l_Lean_quoteNameMk(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_checkBinderAnnotations; @@ -428,7 +427,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLetDeclAux___lambda__7___boxed(lea LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkLetIdDeclView___boxed(lean_object*); 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*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabFun_declRange___closed__3; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__14; 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*); LEAN_EXPORT lean_object* l_Lean_mkFreshId___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_FunBinders_elabFunBinderViews___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__29; @@ -451,6 +449,7 @@ extern lean_object* l_Lean_Elab_macroAttribute; static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_registerFailedToInferBinderTypeInfo___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabBinders___spec__1___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_node2(lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__36; lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabLetDeclAux___lambda__3___closed__1; @@ -477,7 +476,6 @@ static lean_object* l_Lean_Elab_Term_elabLetDeclAux___lambda__1___closed__1; static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_toBinderViews___closed__5; 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*); static lean_object* l___regBuiltin_Lean_Elab_Term_expandForall_declRange___closed__2; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__6; lean_object* l_Lean_LocalContext_mkLocalDecl(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___boxed(lean_object**); static lean_object* l___regBuiltin_Lean_Elab_Term_expandFun_declRange___closed__2; @@ -510,7 +508,6 @@ static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__5; static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_toBinderViews___closed__8; LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandWhereDecls___lambda__1(lean_object*); static lean_object* l_Lean_Elab_Term_expandForall___closed__3; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__1; 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*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabForall___closed__3; static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBinderViews_loop___rarg___closed__2; @@ -525,6 +522,7 @@ lean_object* l_Lean_Syntax_setArg(lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabFun_declRange___closed__7; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabLetTmpDecl(lean_object*); static lean_object* l_Lean_Elab_Term_elabLetDeclAux___closed__5; +LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLetDeclCore___lambda__3(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_object* l_Lean_Elab_Term_addTermInfo_x27(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_Lean_Elab_Term_expandWhereDecls___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandFunBinders_loop___spec__6(lean_object*, size_t, size_t, lean_object*); @@ -547,7 +545,6 @@ static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_toBinderView LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_FunBinders_propagateExpectedType(lean_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_expandFunBinders_loop___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__4; static lean_object* l___regBuiltin_Lean_Elab_Term_expandForall___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetFunDecl_declRange___closed__6; uint8_t l_Lean_Name_isAtomic(lean_object*); @@ -580,6 +577,8 @@ LEAN_EXPORT lean_object* l_Lean_Meta_withLCtx___at___private_Lean_Elab_Binders_0 lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandForall___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT 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*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__14; +LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252_(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetDecl_declRange___closed__6; lean_object* l_Lean_indentExpr(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*); @@ -616,11 +615,13 @@ lean_object* l_Lean_ResolveName_resolveGlobalName(lean_object*, lean_object*, le static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetDelayedDecl___closed__1; lean_object* lean_nat_sub(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabLetDeclCore___closed__6; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__12; LEAN_EXPORT lean_object* l_Lean_mkFreshId___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_FunBinders_elabFunBinderViews___spec__2___rarg(lean_object*, lean_object*); 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*); lean_object* l_Lean_Elab_Term_clearInMatch(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_expandFun___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandMatchAltsWhereDecls___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__2; static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetDecl_declRange___closed__7; LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_2400_(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetDecl___closed__3; @@ -654,6 +655,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLetDeclAux___lambda__4(uint8_t, le static lean_object* l_Lean_Elab_Term_expandWhereDecls___closed__9; static lean_object* l_Lean_Elab_Term_expandLetEqnsDecl___closed__2; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__1; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__3; size_t lean_usize_add(size_t, size_t); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_expandFunBinders_loop___closed__2; @@ -706,9 +708,9 @@ static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchA static lean_object* l_Lean_Elab_Term_elabLetDeclCore___lambda__2___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_elabFun_declRange___closed__6; 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_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__7; lean_object* lean_array_get_size(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabArrow_declRange___closed__3; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__8; uint8_t lean_nat_dec_le(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabLetDeclCore___closed__2; LEAN_EXPORT 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*); @@ -757,6 +759,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabForall_declRange___closed_ 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___regBuiltin_Lean_Elab_Term_elabDepArrow_declRange___closed__3; static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__15; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__11; static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__44; static lean_object* l_Lean_Elab_Term_declareTacticSyntax___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_precheckFun___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -764,7 +767,6 @@ static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__9; 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_expandForall(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_expandFunBinders_loop___closed__9; -LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandExplicitFun(lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabForall_declRange___closed__2; static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___closed__1; @@ -5436,7 +5438,7 @@ else { lean_object* x_30; uint8_t x_31; x_30 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBinderViews_loop___rarg___closed__1; -x_31 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_24, x_30); +x_31 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_24, x_30); lean_dec(x_24); if (x_31 == 0) { @@ -22976,6 +22978,14 @@ return x_106; } } } +LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLetDeclCore___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, uint8_t x_6, uint8_t x_7, uint8_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: +{ +lean_object* x_17; +x_17 = l_Lean_Elab_Term_elabLetDeclAux(x_9, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +return x_17; +} +} static lean_object* _init_l_Lean_Elab_Term_elabLetDeclCore___closed__1() { _start: { @@ -23193,7 +23203,7 @@ 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_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; lean_dec(x_19); lean_dec(x_14); lean_dec(x_1); @@ -23208,8 +23218,65 @@ lean_inc(x_54); x_55 = lean_ctor_get(x_51, 3); lean_inc(x_55); lean_dec(x_51); -x_56 = l_Lean_Elab_Term_elabLetDeclAux(x_52, x_53, x_54, x_55, x_18, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_56; +x_56 = l_Lean_Syntax_isIdent(x_52); +if (x_56 == 0) +{ +uint8_t x_57; lean_object* x_58; +x_57 = 1; +lean_inc(x_11); +lean_inc(x_10); +x_58 = l_Lean_Elab_Term_mkFreshIdent___at_Lean_Elab_Term_elabLetDeclCore___spec__1(x_52, x_57, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_58) == 0) +{ +lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_59 = lean_ctor_get(x_58, 0); +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_Term_elabLetDeclAux(x_59, x_53, x_54, x_55, x_18, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_60); +return x_61; +} +else +{ +uint8_t x_62; +lean_dec(x_55); +lean_dec(x_54); +lean_dec(x_53); +lean_dec(x_18); +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_2); +x_62 = !lean_is_exclusive(x_58); +if (x_62 == 0) +{ +return x_58; +} +else +{ +lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_63 = lean_ctor_get(x_58, 0); +x_64 = lean_ctor_get(x_58, 1); +lean_inc(x_64); +lean_inc(x_63); +lean_dec(x_58); +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_66; +x_66 = l_Lean_Elab_Term_elabLetDeclAux(x_52, x_53, x_54, x_55, x_18, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_66; +} } } } @@ -23253,6 +23320,20 @@ x_19 = l_Lean_Elab_Term_elabLetDeclCore___lambda__2(x_1, x_2, x_3, x_4, x_16, x_ return x_19; } } +LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLetDeclCore___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: +{ +uint8_t x_17; uint8_t x_18; uint8_t x_19; lean_object* x_20; +x_17 = lean_unbox(x_6); +lean_dec(x_6); +x_18 = lean_unbox(x_7); +lean_dec(x_7); +x_19 = lean_unbox(x_8); +lean_dec(x_8); +x_20 = l_Lean_Elab_Term_elabLetDeclCore___lambda__3(x_1, x_2, x_3, x_4, x_5, x_17, x_18, x_19, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +return x_20; +} +} LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLetDeclCore___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, 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: { @@ -23333,7 +23414,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(740u); +x_1 = lean_unsigned_to_nat(741u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23345,7 +23426,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(741u); +x_1 = lean_unsigned_to_nat(742u); x_2 = lean_unsigned_to_nat(129u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23373,7 +23454,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(740u); +x_1 = lean_unsigned_to_nat(741u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23385,7 +23466,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(740u); +x_1 = lean_unsigned_to_nat(741u); x_2 = lean_unsigned_to_nat(42u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23504,7 +23585,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(743u); +x_1 = lean_unsigned_to_nat(744u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23516,7 +23597,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(744u); +x_1 = lean_unsigned_to_nat(745u); x_2 = lean_unsigned_to_nat(130u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23544,7 +23625,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(743u); +x_1 = lean_unsigned_to_nat(744u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23556,7 +23637,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(743u); +x_1 = lean_unsigned_to_nat(744u); x_2 = lean_unsigned_to_nat(49u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23676,7 +23757,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(746u); +x_1 = lean_unsigned_to_nat(747u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23688,7 +23769,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(747u); +x_1 = lean_unsigned_to_nat(748u); x_2 = lean_unsigned_to_nat(128u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23716,7 +23797,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(746u); +x_1 = lean_unsigned_to_nat(747u); x_2 = lean_unsigned_to_nat(39u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23728,7 +23809,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(746u); +x_1 = lean_unsigned_to_nat(747u); x_2 = lean_unsigned_to_nat(57u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23848,7 +23929,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(749u); +x_1 = lean_unsigned_to_nat(750u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23860,7 +23941,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(750u); +x_1 = lean_unsigned_to_nat(751u); x_2 = lean_unsigned_to_nat(128u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23888,7 +23969,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(749u); +x_1 = lean_unsigned_to_nat(750u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23900,7 +23981,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(749u); +x_1 = lean_unsigned_to_nat(750u); x_2 = lean_unsigned_to_nat(49u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23946,7 +24027,7 @@ x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__1() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -23956,7 +24037,7 @@ x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__2() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -23966,27 +24047,27 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__3() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____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_Binders___hyg_12196____closed__2; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__2; x_2 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__4() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__3; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__3; x_2 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__3; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__5() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__5() { _start: { lean_object* x_1; @@ -23994,17 +24075,17 @@ x_1 = lean_mk_string_from_bytes("initFn", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__6() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__4; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__5; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__4; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__7() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__7() { _start: { lean_object* x_1; @@ -24012,37 +24093,37 @@ x_1 = lean_mk_string_from_bytes("_@", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__8() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__6; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__7; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__6; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__7; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__9() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__8; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__8; x_2 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__10() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__9; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__9; x_2 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__11() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__11() { _start: { lean_object* x_1; @@ -24050,17 +24131,17 @@ x_1 = lean_mk_string_from_bytes("Binders", 7); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__12() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__10; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__11; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__10; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__11; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__13() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__13() { _start: { lean_object* x_1; @@ -24068,33 +24149,33 @@ x_1 = lean_mk_string_from_bytes("_hyg", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__14() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__12; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__13; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__12; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__13; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__15() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__14; -x_2 = lean_unsigned_to_nat(12196u); +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__14; +x_2 = lean_unsigned_to_nat(12252u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__1; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__1; x_3 = 0; -x_4 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__15; +x_4 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__15; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } @@ -24890,37 +24971,37 @@ 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); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__1 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__1(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__1); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__2 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__2(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__2); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__3 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__3(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__3); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__4 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__4(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__4); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__5 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__5(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__5); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__6 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__6(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__6); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__7 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__7(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__7); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__8 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__8(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__8); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__9 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__9(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__9); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__10 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__10(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__10); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__11 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__11(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__11); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__12 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__12(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__12); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__13 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__13(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__13); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__14 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__14(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__14); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__15 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__15(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196____closed__15); -res = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12196_(lean_io_mk_world()); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__1 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__1(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__1); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__2 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__2(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__2); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__3 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__3(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__3); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__4 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__4(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__4); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__5 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__5(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__5); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__6 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__6(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__6); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__7 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__7(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__7); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__8 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__8(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__8); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__9 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__9(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__9); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__10 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__10(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__10); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__11 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__11(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__11); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__12 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__12(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__12); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__13 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__13(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__13); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__14 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__14(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__14); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__15 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__15(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252____closed__15); +res = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_12252_(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/BuiltinCommand.c b/stage0/stdlib/Lean/Elab/BuiltinCommand.c index ec819294035a..399f953569c6 100644 --- a/stage0/stdlib/Lean/Elab/BuiltinCommand.c +++ b/stage0/stdlib/Lean/Elab/BuiltinCommand.c @@ -17,7 +17,6 @@ static lean_object* l___regBuiltin_Lean_Elab_Command_elabCheck_declRange___close lean_object* l_Lean_Name_getNumParts(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabAddDeclDoc___spec__13(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); -uint8_t l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabSetOption_declRange___closed__7; lean_object* lean_string_utf8_extract(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_OpenDecl_elabOpenDecl___at_Lean_Elab_Command_elabOpen___spec__1___closed__2; @@ -194,7 +193,6 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabExport___s LEAN_EXPORT lean_object* l_Lean_resolveGlobalConst___at_Lean_Elab_Command_elabCheckCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabCheck_declRange___closed__1; static lean_object* l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_mkRunMetaEval___closed__4; -lean_object* l_Lean_Option_get___at_Std_Format_pretty_x27___spec__1(lean_object*, lean_object*); lean_object* l_Lean_replaceRef(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_foldArgsM___at_Lean_Elab_Command_elabUniverse___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ofExcept___at_Lean_Elab_Command_elabEvalUnsafe___spec__4___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -438,6 +436,7 @@ static lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_Command_elabExpor LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabVariable___lambda__1___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_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__4___closed__6; lean_object* l_Lean_ScopedEnvExtension_popScope___rarg(lean_object*, lean_object*); +lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandInCmd(lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabVariable_declRange___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabSection(lean_object*, lean_object*, lean_object*, lean_object*); @@ -523,7 +522,6 @@ lean_object* lean_array_to_list(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabSynth___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_addScopes___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabSynth_declRange___closed__6; -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(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__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___regBuiltin_Lean_Elab_Command_elabEval_declRange___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabEnd___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -606,6 +604,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Command_elabAddDeclDoc_declRange___ lean_object* l_Lean_Syntax_node2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabSetOption(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabAddDeclDoc___spec__8(lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabExport___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabReduce___boxed(lean_object*, lean_object*, lean_object*, lean_object*); 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*); @@ -778,6 +777,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Command_elabSynth___closed__1; lean_object* lean_environment_main_module(lean_object*); static lean_object* l___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_typelessBinder_x3f___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Command_elabImport_declRange___closed__2; +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabSetOption_setOption___at_Lean_Elab_Command_elabSetOption___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabEnd_declRange___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabSynth___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -15466,7 +15466,6 @@ lean_dec(x_9); lean_inc(x_10); x_11 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabVariable___lambda__2___boxed), 9, 1); lean_closure_set(x_11, 0, x_10); -lean_inc(x_2); x_12 = l_Lean_Elab_Command_runTermElabM___rarg(x_11, x_2, x_3, x_4); if (lean_obj_tag(x_12) == 0) { @@ -17014,7 +17013,6 @@ x_7 = l_Lean_Syntax_isOfKind(x_2, x_6); if (x_7 == 0) { lean_object* x_8; -lean_dec(x_3); lean_dec(x_2); x_8 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabNamespace___spec__1___rarg(x_5); return x_8; @@ -17041,7 +17039,6 @@ lean_dec(x_15); x_18 = lean_ctor_get(x_16, 0); lean_inc(x_18); lean_dec(x_16); -lean_inc(x_3); x_19 = l_Lean_Elab_Command_runTermElabM___rarg(x_14, x_3, x_4, x_17); if (lean_obj_tag(x_19) == 0) { @@ -17052,7 +17049,6 @@ x_21 = lean_ctor_get(x_19, 1); lean_inc(x_21); lean_dec(x_19); x_22 = l_Lean_setEnv___at_Lean_Elab_Command_elabInitQuot___spec__1(x_18, x_3, x_4, x_21); -lean_dec(x_3); x_23 = !lean_is_exclusive(x_22); if (x_23 == 0) { @@ -17083,7 +17079,6 @@ x_28 = lean_ctor_get(x_19, 1); lean_inc(x_28); lean_dec(x_19); x_29 = l_Lean_setEnv___at_Lean_Elab_Command_elabInitQuot___spec__1(x_18, x_3, x_4, x_28); -lean_dec(x_3); x_30 = !lean_is_exclusive(x_29); if (x_30 == 0) { @@ -17280,6 +17275,7 @@ x_6 = lean_unbox(x_1); lean_dec(x_1); x_7 = l_Lean_Elab_Command_elabCheckCore(x_6, x_2, x_3, x_4, x_5); lean_dec(x_4); +lean_dec(x_3); return x_7; } } @@ -17298,6 +17294,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCheck___boxed(lean_object* x_1, lean_object* x_5; x_5 = l_Lean_Elab_Command_elabCheck(x_1, x_2, x_3, x_4); lean_dec(x_3); +lean_dec(x_2); return x_5; } } @@ -17959,7 +17956,6 @@ x_6 = l_Lean_Syntax_isOfKind(x_1, x_5); if (x_6 == 0) { lean_object* x_7; -lean_dec(x_2); lean_dec(x_1); x_7 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabNamespace___spec__1___rarg(x_4); return x_7; @@ -17984,7 +17980,6 @@ lean_dec(x_13); x_16 = lean_ctor_get(x_14, 0); lean_inc(x_16); lean_dec(x_14); -lean_inc(x_2); x_17 = l_Lean_Elab_Command_runTermElabM___rarg(x_12, x_2, x_3, x_15); if (lean_obj_tag(x_17) == 0) { @@ -17995,7 +17990,6 @@ x_19 = lean_ctor_get(x_17, 1); lean_inc(x_19); lean_dec(x_17); x_20 = l_Lean_setEnv___at_Lean_Elab_Command_elabInitQuot___spec__1(x_16, x_2, x_3, x_19); -lean_dec(x_2); x_21 = !lean_is_exclusive(x_20); if (x_21 == 0) { @@ -18026,7 +18020,6 @@ x_26 = lean_ctor_get(x_17, 1); lean_inc(x_26); lean_dec(x_17); x_27 = l_Lean_setEnv___at_Lean_Elab_Command_elabInitQuot___spec__1(x_16, x_2, x_3, x_26); -lean_dec(x_2); x_28 = !lean_is_exclusive(x_27); if (x_28 == 0) { @@ -18076,6 +18069,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabReduce___boxed(lean_object* x_1 lean_object* x_5; x_5 = l_Lean_Elab_Command_elabReduce(x_1, x_2, x_3, x_4); lean_dec(x_3); +lean_dec(x_2); return x_5; } } @@ -19814,7 +19808,7 @@ lean_inc(x_37); x_38 = lean_ctor_get(x_36, 1); lean_inc(x_38); lean_dec(x_36); -x_39 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_37, x_5, x_6, x_7, x_8, x_38); +x_39 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_37, x_5, x_6, x_7, x_8, x_38); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -20165,7 +20159,7 @@ x_23 = lean_array_push(x_22, x_8); x_24 = lean_array_push(x_23, x_15); x_25 = lean_array_push(x_24, x_21); x_26 = l_Lean_mkAppN(x_20, x_25); -x_27 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_26, x_2, x_3, x_4, x_5, x_16); +x_27 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_26, x_2, x_3, x_4, x_5, x_16); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); @@ -23096,13 +23090,13 @@ lean_object* x_22; lean_dec(x_14); x_22 = l_Lean_Elab_Command_runTermElabM___rarg(x_15, x_2, x_3, x_18); lean_dec(x_3); +lean_dec(x_2); return x_22; } else { lean_object* x_23; lean_dec(x_15); -lean_inc(x_2); x_23 = l_Lean_Elab_Command_runTermElabM___rarg(x_14, x_2, x_3, x_18); if (lean_obj_tag(x_23) == 0) { @@ -23685,7 +23679,6 @@ lean_dec(x_8); x_11 = lean_ctor_get(x_9, 0); lean_inc(x_11); lean_dec(x_9); -lean_inc(x_2); x_12 = l_Lean_Elab_Command_runTermElabM___rarg(x_7, x_2, x_3, x_10); if (lean_obj_tag(x_12) == 0) { @@ -23696,7 +23689,6 @@ x_14 = lean_ctor_get(x_12, 1); lean_inc(x_14); lean_dec(x_12); x_15 = l_Lean_setEnv___at_Lean_Elab_Command_elabInitQuot___spec__1(x_11, x_2, x_3, x_14); -lean_dec(x_2); x_16 = !lean_is_exclusive(x_15); if (x_16 == 0) { @@ -23727,7 +23719,6 @@ x_21 = lean_ctor_get(x_12, 1); lean_inc(x_21); lean_dec(x_12); x_22 = l_Lean_setEnv___at_Lean_Elab_Command_elabInitQuot___spec__1(x_11, x_2, x_3, x_21); -lean_dec(x_2); x_23 = !lean_is_exclusive(x_22); if (x_23 == 0) { @@ -23767,6 +23758,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabSynth___boxed(lean_object* x_1, lean_object* x_5; x_5 = l_Lean_Elab_Command_elabSynth(x_1, x_2, x_3, x_4); lean_dec(x_3); +lean_dec(x_2); lean_dec(x_1); return x_5; } @@ -24435,7 +24427,7 @@ lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean x_16 = lean_ctor_get(x_13, 4); lean_dec(x_16); x_17 = l_Lean_Elab_Command_elabSetOption___closed__1; -x_18 = l_Lean_Option_get___at_Std_Format_pretty_x27___spec__1(x_10, x_17); +x_18 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_10, x_17); lean_ctor_set(x_13, 4, x_18); x_19 = lean_st_ref_set(x_3, x_13, x_14); x_20 = lean_ctor_get(x_19, 1); @@ -24469,7 +24461,7 @@ lean_inc(x_24); lean_inc(x_23); lean_dec(x_13); x_31 = l_Lean_Elab_Command_elabSetOption___closed__1; -x_32 = l_Lean_Option_get___at_Std_Format_pretty_x27___spec__1(x_10, x_31); +x_32 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_10, x_31); x_33 = lean_alloc_ctor(0, 9, 0); lean_ctor_set(x_33, 0, x_23); lean_ctor_set(x_33, 1, x_24); @@ -26438,7 +26430,7 @@ x_11 = lean_ctor_get(x_10, 1); lean_inc(x_11); lean_dec(x_10); x_12 = l_Lean_logWarning___at_Lean_Elab_Command_elabExit___spec__1___closed__1; -x_13 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_11, x_12); +x_13 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_11, x_12); lean_dec(x_11); if (x_13 == 0) { diff --git a/stage0/stdlib/Lean/Elab/BuiltinNotation.c b/stage0/stdlib/Lean/Elab/BuiltinNotation.c index a00d1f28b347..56ee02a305da 100644 --- a/stage0/stdlib/Lean/Elab/BuiltinNotation.c +++ b/stage0/stdlib/Lean/Elab/BuiltinNotation.c @@ -107,7 +107,6 @@ static lean_object* l_Lean_Elab_Term_elabPanic___closed__12; LEAN_EXPORT lean_object* l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___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___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Term_expandTuple___closed__1; -static lean_object* l_Lean_Elab_Term_expandSuffices___lambda__1___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_expandTypeAscription_declRange___closed__4; static lean_object* l_Lean_Elab_Term_elabSubst___lambda__4___closed__6; lean_object* l_Lean_Name_toString(lean_object*, uint8_t); @@ -117,7 +116,6 @@ static lean_object* l_Lean_Elab_Term_elabSorry___closed__5; static lean_object* l_Lean_Elab_Term_elabLeadingParserMacro___lambda__2___closed__2; static lean_object* l_Lean_Elab_Term_expandDbgTrace___closed__18; lean_object* l_Lean_Expr_sort___override(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandSuffices___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_expandParen_declRange___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Term_expandSuffices_declRange___closed__5; lean_object* lean_array_push(lean_object*, lean_object*); @@ -152,6 +150,7 @@ static lean_object* l_Lean_Elab_Term_expandAssert___closed__9; lean_object* l_Lean_replaceRef(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_expandDbgTrace___closed__9; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_expandHave(lean_object*); +static lean_object* l_Lean_Elab_Term_expandSuffices___closed__5; static lean_object* l___regBuiltin_Lean_Elab_Term_elabCoe___closed__8; lean_object* l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___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_liftMacroM___at_Lean_Elab_Term_elabCDotFunctionAlias_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -222,9 +221,11 @@ static lean_object* l_Lean_Elab_Term_elabSubst___closed__6; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___spec__4___boxed(lean_object*, lean_object*, lean_object*); uint8_t lean_string_dec_eq(lean_object*, lean_object*); 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*); +static lean_object* l_Lean_Elab_Term_expandHave___closed__12; static lean_object* l___regBuiltin_Lean_Elab_Term_elabCoe___closed__6; static lean_object* l_Lean_Elab_Term_elabSubst___closed__11; static lean_object* l_Lean_Elab_Term_mkPairs_loop___closed__5; +static lean_object* l_Lean_Elab_Term_expandHave___closed__11; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_elabCDotFunctionAlias_x3f___spec__1___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_expandDbgTrace___closed__17; LEAN_EXPORT lean_object* l_Lean_exprDependsOn___at___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_isSubstCandidate___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -235,11 +236,11 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___lambda__1(lean_object*, le LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandCDot_x3f_go(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_expandHave___closed__1; static lean_object* l_Lean_Elab_Term_elabCDotFunctionAlias_x3f___closed__2; +static lean_object* l_Lean_Elab_Term_expandHave___closed__14; static lean_object* l___regBuiltin_Lean_Elab_Term_expandHave_declRange___closed__2; static lean_object* l_Lean_Elab_Term_expandCDot_x3f___closed__1; static lean_object* l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabTParserMacroAux___closed__3; static lean_object* l_Lean_Elab_Term_hasCDot___closed__4; -static lean_object* l_Lean_Elab_Term_expandHave___lambda__3___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_elabPanic_declRange___closed__5; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_isSubstCandidate___spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_getConstInfoCtor___at_Lean_Elab_Term_elabAnonymousCtor___spec__2___closed__4; @@ -268,9 +269,10 @@ static lean_object* l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elab LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_expandParen(lean_object*); static lean_object* l_Lean_Elab_Term_expandTuple___closed__6; static lean_object* l___regBuiltin_Lean_Elab_Term_elabTypeAscription_declRange___closed__1; -LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandHave___lambda__4(lean_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_expandHave___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___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabTParserMacroAux___closed__9; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_expandHave_declRange(lean_object*); +static lean_object* l_Lean_Elab_Term_expandSuffices___closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandUnreachable___rarg(lean_object*, lean_object*); lean_object* l_Lean_Meta_isTypeCorrect(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); @@ -308,6 +310,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabStateRefT_declRange___clos static lean_object* l_Lean_Elab_Term_expandDbgTrace___closed__3; static lean_object* l_Lean_Elab_Term_elabPanic___closed__7; lean_object* l_liftExcept___at_Lean_Elab_liftMacroM___spec__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandHave___lambda__6___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_elabShow_declRange___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSorry(lean_object*, 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*); @@ -363,11 +366,10 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabShow_declRange___closed__1 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__32; static lean_object* l___regBuiltin_Lean_Elab_Term_expandTuple_declRange___closed__4; -static lean_object* l_Lean_Elab_Term_expandDbgTrace___closed__25; 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*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabShow_declRange___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Term_expandSuffices_declRange___closed__4; -LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandHave___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandHave___lambda__5(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*); 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*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabShow___closed__1; @@ -463,7 +465,6 @@ static lean_object* l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__9; static lean_object* l_Lean_Elab_Term_elabCDotFunctionAlias_x3f___closed__4; 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*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabStateRefT_declRange___closed__5; -LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandSuffices___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabTypeAscription_declRange___closed__7; static lean_object* l_Lean_Elab_Term_elabShow___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Term_elabTypeAscription_declRange___closed__3; @@ -487,7 +488,6 @@ lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabSubst___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Term_elabTypeAscription_declRange___closed__4; static lean_object* l_Lean_Elab_Term_elabShow___closed__13; -static lean_object* l_Lean_Elab_Term_expandDbgTrace___closed__26; LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_elabCDotFunctionAlias_x3f___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_mkPairs_loop___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Term_elabAnonymousCtor_declRange___closed__1; @@ -547,13 +547,12 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabStateRefT(lean_object LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabCoe___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_elabCoe___closed__5; static lean_object* l___regBuiltin_Lean_Elab_Term_expandHave_declRange___closed__6; -static lean_object* l_Lean_Elab_Term_expandDbgTrace___closed__24; static lean_object* l_Lean_Elab_Term_elabStateRefT___lambda__1___closed__2; static lean_object* l_Lean_Elab_Term_expandSuffices___closed__4; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabCDotFunctionAlias_x3f___spec__6(lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Elab_Term_elabShow___closed__2; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_isSubstCandidate___spec__9(lean_object*, lean_object*, size_t, size_t); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandHave___lambda__5___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_expandHave___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___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_isSubstCandidate___spec__11___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabShow_declRange___closed__3; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_expandTypeAscription(lean_object*); @@ -576,6 +575,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabBorrowed___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Term_expandUnreachable___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabTypeAscription_declRange(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_expandTypeAscription___closed__1; +static lean_object* l_Lean_Elab_Term_expandHave___closed__10; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Term_elabAnonymousCtor___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_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_expandAssert_declRange(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*); @@ -617,11 +617,14 @@ static lean_object* l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elab LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandCDot_x3f(lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_expandHave_declRange___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabStateRefT___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_expandHave___lambda__7(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_expandAssert___closed__10; static lean_object* l_Lean_Elab_Term_elabStateRefT___lambda__1___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux(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_Term_expandSuffices___closed__9; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_isSubstCandidate___spec__16(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___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*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_expandHave___lambda__4___closed__1; static lean_object* l_Lean_Elab_Term_expandTuple___closed__7; static lean_object* l_Lean_Elab_Term_elabAnonymousCtor___closed__4; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___spec__3___boxed(lean_object*, lean_object*, lean_object*); @@ -660,7 +663,6 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_expandUnreachable___closed__4; static lean_object* l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Term_hasCDot___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabCDotFunctionAlias_x3f___lambda__1___boxed(lean_object*); -static lean_object* l_Lean_Elab_Term_expandSuffices___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_getRefPos___at_Lean_Elab_Term_elabPanic___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_isSubstCandidate___spec__6(lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); @@ -668,12 +670,12 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabTrailingParserMacro___lambda__1___ static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_elabCDotFunctionAlias_x3f___spec__4___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabNoindex(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_isPrivateNameFromImportedModule(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_expandHave___lambda__3___closed__2; static lean_object* l_Lean_Elab_Term_elabPanic___closed__3; static lean_object* l___regBuiltin_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*); static lean_object* l_Lean_Elab_Term_expandSuffices___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLeadingParserMacro___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_Term_expandHave___closed__15; static lean_object* l_Lean_Elab_Term_expandUnreachable___rarg___closed__2; static lean_object* l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__33; lean_object* l_Lean_indentExpr(lean_object*); @@ -702,12 +704,13 @@ lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0_ LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_elabCDotFunctionAlias_x3f___spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabAnonymousCtor___closed__2; static lean_object* l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__6; +static lean_object* l_Lean_Elab_Term_expandSuffices___closed__8; static lean_object* l___regBuiltin_Lean_Elab_Term_elabPanic_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*); static lean_object* l_Lean_Elab_Term_expandAssert___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_expandHave_declRange___closed__7; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_isSubstCandidate___spec__2___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandHave___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_Lean_Elab_Term_expandHave___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_Elab_Term_elabShow___closed__11; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabAnonymousCtor(lean_object*); static lean_object* l_Lean_Elab_Term_elabAnonymousCtor___closed__5; @@ -726,7 +729,7 @@ static lean_object* l_Lean_getConstInfoCtor___at_Lean_Elab_Term_elabAnonymousCto static lean_object* l_Lean_Elab_Term_expandSuffices___closed__1; lean_object* l_Lean_ResolveName_resolveGlobalName(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandHave___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_Term_expandHave___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_Elab_Term_expandSuffices(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_getAppFn(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabSorry_declRange___closed__4; @@ -738,7 +741,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_expandHave___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Term_expandShow_declRange___closed__5; 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___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabTParserMacroAux___closed__2; -LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandHave___lambda__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_Elab_Term_expandHave___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___regBuiltin_Lean_Elab_Term_expandParen___closed__1; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Term_hasCDot___spec__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabCoe___closed__7; @@ -772,6 +775,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabAnonymousCtor___lambda__1(lean_obj static lean_object* l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__49; LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandHave___lambda__1___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_expandSuffices___closed__3; +LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandHave___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabAnonymousCtor___lambda__4___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabAnonymousCtor___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_String_intercalate(lean_object*, lean_object*); @@ -838,7 +842,6 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_expandUnreachable(lean_ob LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_isSubstCandidate___spec__13___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabTrailingParserMacro(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_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__41; -static lean_object* l_Lean_Elab_Term_expandSuffices___lambda__1___closed__3; static lean_object* l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabTParserMacroAux___closed__6; static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_elabCDotFunctionAlias_x3f___spec__4___closed__2; static lean_object* l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__46; @@ -859,7 +862,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandAssert(lean_object*, lean_object uint8_t lean_usize_dec_lt(size_t, size_t); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabStateRefT_declRange(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabLeadingParserMacro___closed__1; -lean_object* l_Array_mkArray2___rarg(lean_object*, lean_object*); 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_elabCoe___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Term_elabBorrowed_declRange___closed__5; @@ -867,6 +869,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabSubst_declRange___closed__ lean_object* l_Lean_getExprMVarAssignment_x3f___at___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___spec__1(lean_object*, 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___regBuiltin_Lean_Elab_Term_elabNoindex___closed__5; +lean_object* l_Lean_HygieneInfo_mkIdent(lean_object*, lean_object*, uint8_t); static lean_object* l_Lean_Elab_Term_elabAnonymousCtor___closed__1; lean_object* lean_nat_add(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_expandShow_declRange___closed__3; @@ -882,6 +885,7 @@ static lean_object* l_Lean_Elab_Term_elabPanic___closed__14; static lean_object* l___regBuiltin_Lean_Elab_Term_expandParen_declRange___closed__4; static lean_object* l_Lean_Elab_Term_expandUnreachable___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandParen(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_expandSuffices___closed__7; static lean_object* l___regBuiltin_Lean_Elab_Term_elabLeadingParserMacro___closed__2; static lean_object* l_Lean_Elab_Term_elabPanic___closed__6; static lean_object* l___regBuiltin_Lean_Elab_Term_elabTrailingParserMacro___closed__2; @@ -894,6 +898,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabLeadingParserMacro(le lean_object* l_Lean_expandMacros(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabLeadingParserMacro___lambda__2___closed__1; static lean_object* l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__9; +static lean_object* l_Lean_Elab_Term_expandHave___closed__13; static lean_object* l___regBuiltin_Lean_Elab_Term_elabCoe_declRange___closed__7; static lean_object* l___regBuiltin_Lean_Elab_Term_elabCoe_declRange___closed__1; lean_object* l_String_toSubstring_x27(lean_object*); @@ -934,6 +939,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandHave(lean_object*, lean_object*, static lean_object* l_Lean_Elab_Term_hasCDot___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_elabSorry_declRange___closed__6; static lean_object* l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__48; +LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandHave___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*); static lean_object* l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__13; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabCoe___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: @@ -4449,7 +4455,7 @@ static lean_object* _init_l_Lean_Elab_Term_expandHave___lambda__2___closed__1() _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("matchAlts", 9); +x_1 = lean_mk_string_from_bytes("letEqnsDecl", 11); return x_1; } } @@ -4469,152 +4475,21 @@ static lean_object* _init_l_Lean_Elab_Term_expandHave___lambda__2___closed__3() _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("have", 4); -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandHave___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; lean_object* x_13; uint8_t x_14; -x_11 = lean_unsigned_to_nat(2u); -x_12 = l_Lean_Syntax_getArg(x_1, x_11); -x_13 = l_Lean_Elab_Term_expandHave___lambda__2___closed__2; -lean_inc(x_12); -x_14 = l_Lean_Syntax_isOfKind(x_12, x_13); -if (x_14 == 0) -{ -lean_object* x_15; lean_object* x_16; -lean_dec(x_12); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -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_10); -return x_16; -} -else -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t 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; 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_17 = lean_unsigned_to_nat(3u); -x_18 = l_Lean_Syntax_getArg(x_2, x_17); -x_19 = lean_ctor_get(x_9, 5); -lean_inc(x_19); -lean_dec(x_9); -x_20 = 0; -x_21 = l_Lean_SourceInfo_fromRef(x_19, x_20); -x_22 = l_Lean_Elab_Term_expandHave___lambda__2___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_Lean_Elab_Term_elabShow___closed__2; -x_25 = 1; -x_26 = l_Lean_mkIdentFrom(x_3, x_24, x_25); -x_27 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; -x_28 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__4; -lean_inc(x_21); -x_29 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_29, 0, x_21); -lean_ctor_set(x_29, 1, x_27); -lean_ctor_set(x_29, 2, x_28); -lean_inc(x_21); -x_30 = l_Lean_Syntax_node2(x_21, x_27, x_26, x_29); -x_31 = l_Lean_Elab_Term_elabShow___closed__13; -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); -if (lean_obj_tag(x_8) == 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; -x_33 = l_Lean_Elab_Term_expandHave___lambda__1___closed__1; -lean_inc(x_21); -x_34 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_34, 0, x_21); -lean_ctor_set(x_34, 1, x_27); -lean_ctor_set(x_34, 2, x_33); -lean_inc(x_21); -x_35 = l_Lean_Syntax_node3(x_21, x_4, x_30, x_34, x_12); -lean_inc(x_21); -x_36 = l_Lean_Syntax_node1(x_21, x_5, x_35); -x_37 = l_Lean_Syntax_node4(x_21, x_6, x_23, x_36, x_32, x_18); -x_38 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_38, 0, x_37); -lean_ctor_set(x_38, 1, x_10); -return x_38; -} -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; -x_39 = lean_ctor_get(x_8, 0); -lean_inc(x_39); -lean_dec(x_8); -x_40 = l_Lean_Elab_Term_elabShow___closed__11; -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_Lean_Elab_Term_elabShow___closed__10; -lean_inc(x_21); -x_43 = l_Lean_Syntax_node2(x_21, x_42, x_41, x_39); -x_44 = l_Array_mkArray1___rarg(x_43); -x_45 = l_Array_append___rarg(x_28, x_44); -lean_inc(x_21); -x_46 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_46, 0, x_21); -lean_ctor_set(x_46, 1, x_27); -lean_ctor_set(x_46, 2, x_45); -lean_inc(x_21); -x_47 = l_Lean_Syntax_node3(x_21, x_4, x_30, x_46, x_12); -lean_inc(x_21); -x_48 = l_Lean_Syntax_node1(x_21, x_5, x_47); -x_49 = l_Lean_Syntax_node4(x_21, x_6, x_23, x_48, x_32, x_18); -x_50 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_50, 0, x_49); -lean_ctor_set(x_50, 1, x_10); -return x_50; -} -} -} -} -static lean_object* _init_l_Lean_Elab_Term_expandHave___lambda__3___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("letEqnsDecl", 11); +x_1 = lean_mk_string_from_bytes("_", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_expandHave___lambda__3___closed__2() { -_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_elabCoe___closed__3; -x_2 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__8; -x_3 = l___regBuiltin_Lean_Elab_Term_elabCoe___closed__5; -x_4 = l_Lean_Elab_Term_expandHave___lambda__3___closed__1; -x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); -return x_5; -} -} -LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandHave___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_EXPORT lean_object* l_Lean_Elab_Term_expandHave___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) { _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; 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; -x_9 = lean_unsigned_to_nat(2u); -x_10 = l_Lean_Syntax_getArg(x_1, x_9); -x_11 = lean_unsigned_to_nat(3u); -x_12 = l_Lean_Syntax_getArg(x_2, x_11); +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; uint8_t 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_10 = lean_unsigned_to_nat(3u); +x_11 = l_Lean_Syntax_getArg(x_1, x_10); +x_12 = l_Lean_Syntax_getArg(x_2, x_10); x_13 = l_Lean_Syntax_getArgs(x_3); -x_14 = lean_ctor_get(x_7, 5); +x_14 = lean_ctor_get(x_8, 5); lean_inc(x_14); -lean_dec(x_7); +lean_dec(x_8); x_15 = 0; x_16 = l_Lean_SourceInfo_fromRef(x_14, x_15); x_17 = l_Lean_Elab_Term_elabShow___closed__3; @@ -4622,172 +4497,89 @@ 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_Elab_Term_elabAnonymousCtor___lambda__3___closed__4; -x_20 = l_Array_append___rarg(x_19, x_13); -x_21 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; -lean_inc(x_16); -x_22 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_22, 0, x_16); +x_19 = 1; +x_20 = l_Lean_SourceInfo_fromRef(x_4, x_19); +x_21 = l_Lean_Elab_Term_expandHave___lambda__2___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); -lean_ctor_set(x_22, 2, x_20); -x_23 = l_Lean_Elab_Term_elabShow___closed__13; -lean_inc(x_16); -x_24 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_24, 0, x_16); -lean_ctor_set(x_24, 1, x_23); -if (lean_obj_tag(x_6) == 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; -x_25 = l_Lean_Elab_Term_expandHave___lambda__1___closed__1; -lean_inc(x_16); -x_26 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_26, 0, x_16); -lean_ctor_set(x_26, 1, x_21); -lean_ctor_set(x_26, 2, x_25); -x_27 = l_Lean_Elab_Term_expandHave___lambda__3___closed__2; lean_inc(x_16); -x_28 = l_Lean_Syntax_node4(x_16, x_27, x_4, x_22, x_26, x_10); -x_29 = l_Lean_Elab_Term_elabShow___closed__6; +x_23 = l_Lean_Syntax_node1(x_16, x_5, x_22); +x_24 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__4; +x_25 = l_Array_append___rarg(x_24, x_13); +x_26 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; lean_inc(x_16); -x_30 = l_Lean_Syntax_node1(x_16, x_29, x_28); -x_31 = l_Lean_Elab_Term_elabShow___closed__4; -x_32 = l_Lean_Syntax_node4(x_16, x_31, x_18, x_30, x_24, x_12); -x_33 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_33, 0, x_32); -lean_ctor_set(x_33, 1, x_8); -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; -x_34 = lean_ctor_get(x_6, 0); -lean_inc(x_34); -lean_dec(x_6); -x_35 = l_Lean_Elab_Term_elabShow___closed__11; -lean_inc(x_16); -x_36 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_36, 0, x_16); -lean_ctor_set(x_36, 1, x_35); -x_37 = l_Lean_Elab_Term_elabShow___closed__10; -lean_inc(x_16); -x_38 = l_Lean_Syntax_node2(x_16, x_37, x_36, x_34); -x_39 = l_Array_mkArray1___rarg(x_38); -x_40 = l_Array_append___rarg(x_19, x_39); -lean_inc(x_16); -x_41 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_41, 0, x_16); -lean_ctor_set(x_41, 1, x_21); -lean_ctor_set(x_41, 2, x_40); -x_42 = l_Lean_Elab_Term_expandHave___lambda__3___closed__2; -lean_inc(x_16); -x_43 = l_Lean_Syntax_node4(x_16, x_42, x_4, x_22, x_41, x_10); -x_44 = l_Lean_Elab_Term_elabShow___closed__6; -lean_inc(x_16); -x_45 = l_Lean_Syntax_node1(x_16, x_44, x_43); -x_46 = l_Lean_Elab_Term_elabShow___closed__4; -x_47 = l_Lean_Syntax_node4(x_16, x_46, x_18, x_45, x_24, x_12); -x_48 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_48, 0, x_47); -lean_ctor_set(x_48, 1, x_8); -return x_48; -} -} -} -LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandHave___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) { -_start: -{ -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; uint8_t 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_11 = lean_unsigned_to_nat(3u); -x_12 = l_Lean_Syntax_getArg(x_1, x_11); -x_13 = l_Lean_Syntax_getArg(x_2, x_11); -x_14 = lean_ctor_get(x_9, 5); -lean_inc(x_14); -lean_dec(x_9); -x_15 = 0; -x_16 = l_Lean_SourceInfo_fromRef(x_14, x_15); -x_17 = l_Lean_Elab_Term_expandHave___lambda__2___closed__3; -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_Elab_Term_elabShow___closed__2; -x_20 = 1; -x_21 = l_Lean_mkIdentFrom(x_3, x_19, x_20); -x_22 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; -x_23 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__4; -lean_inc(x_16); -x_24 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_24, 0, x_16); -lean_ctor_set(x_24, 1, x_22); -lean_ctor_set(x_24, 2, x_23); -lean_inc(x_16); -x_25 = l_Lean_Syntax_node2(x_16, x_22, x_21, x_24); -x_26 = l_Lean_Elab_Term_elabShow___closed__12; -lean_inc(x_16); -x_27 = lean_alloc_ctor(2, 2, 0); +x_27 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_27, 0, x_16); lean_ctor_set(x_27, 1, x_26); +lean_ctor_set(x_27, 2, x_25); x_28 = l_Lean_Elab_Term_elabShow___closed__13; 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); -if (lean_obj_tag(x_8) == 0) +if (lean_obj_tag(x_7) == 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_30; lean_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_30 = l_Lean_Elab_Term_expandHave___lambda__1___closed__1; lean_inc(x_16); x_31 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_31, 0, x_16); -lean_ctor_set(x_31, 1, x_22); +lean_ctor_set(x_31, 1, x_26); lean_ctor_set(x_31, 2, x_30); +x_32 = l_Lean_Elab_Term_expandHave___lambda__2___closed__2; lean_inc(x_16); -x_32 = l_Lean_Syntax_node4(x_16, x_4, x_25, x_31, x_27, x_12); +x_33 = l_Lean_Syntax_node4(x_16, x_32, x_23, x_27, x_31, x_11); +x_34 = l_Lean_Elab_Term_elabShow___closed__6; lean_inc(x_16); -x_33 = l_Lean_Syntax_node1(x_16, x_5, x_32); -x_34 = l_Lean_Syntax_node4(x_16, x_6, x_18, x_33, x_29, x_13); -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_34); -lean_ctor_set(x_35, 1, x_10); -return x_35; +x_35 = l_Lean_Syntax_node1(x_16, x_34, x_33); +x_36 = l_Lean_Elab_Term_elabShow___closed__4; +x_37 = l_Lean_Syntax_node4(x_16, x_36, x_18, x_35, x_29, x_12); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_9); +return x_38; } 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_8, 0); -lean_inc(x_36); -lean_dec(x_8); -x_37 = l_Lean_Elab_Term_elabShow___closed__11; +lean_object* x_39; lean_object* x_40; lean_object* x_41; 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_39 = lean_ctor_get(x_7, 0); +lean_inc(x_39); +lean_dec(x_7); +x_40 = l_Lean_Elab_Term_elabShow___closed__11; lean_inc(x_16); -x_38 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_38, 0, x_16); -lean_ctor_set(x_38, 1, x_37); -x_39 = l_Lean_Elab_Term_elabShow___closed__10; +x_41 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_41, 0, x_16); +lean_ctor_set(x_41, 1, x_40); +x_42 = l_Lean_Elab_Term_elabShow___closed__10; lean_inc(x_16); -x_40 = l_Lean_Syntax_node2(x_16, x_39, x_38, x_36); -x_41 = l_Array_mkArray1___rarg(x_40); -x_42 = l_Array_append___rarg(x_23, x_41); +x_43 = l_Lean_Syntax_node2(x_16, x_42, x_41, x_39); +x_44 = l_Array_mkArray1___rarg(x_43); +x_45 = l_Array_append___rarg(x_24, x_44); lean_inc(x_16); -x_43 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_43, 0, x_16); -lean_ctor_set(x_43, 1, x_22); -lean_ctor_set(x_43, 2, x_42); +x_46 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_46, 0, x_16); +lean_ctor_set(x_46, 1, x_26); +lean_ctor_set(x_46, 2, x_45); +x_47 = l_Lean_Elab_Term_expandHave___lambda__2___closed__2; lean_inc(x_16); -x_44 = l_Lean_Syntax_node4(x_16, x_4, x_25, x_43, x_27, x_12); +x_48 = l_Lean_Syntax_node4(x_16, x_47, x_23, x_27, x_46, x_11); +x_49 = l_Lean_Elab_Term_elabShow___closed__6; lean_inc(x_16); -x_45 = l_Lean_Syntax_node1(x_16, x_5, x_44); -x_46 = l_Lean_Syntax_node4(x_16, x_6, x_18, x_45, x_29, x_13); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_10); -return x_47; +x_50 = l_Lean_Syntax_node1(x_16, x_49, x_48); +x_51 = l_Lean_Elab_Term_elabShow___closed__4; +x_52 = l_Lean_Syntax_node4(x_16, x_51, x_18, x_50, x_29, x_12); +x_53 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_53, 0, x_52); +lean_ctor_set(x_53, 1, x_9); +return x_53; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandHave___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_EXPORT lean_object* l_Lean_Elab_Term_expandHave___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) { _start: { -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; 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_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; 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_9 = lean_unsigned_to_nat(3u); x_10 = l_Lean_Syntax_getArg(x_1, x_9); x_11 = l_Lean_Syntax_getArg(x_2, x_9); @@ -4810,91 +4602,470 @@ x_21 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_21, 0, x_15); lean_ctor_set(x_21, 1, x_20); lean_ctor_set(x_21, 2, x_19); -x_22 = l_Lean_Elab_Term_elabShow___closed__12; +x_22 = l_Lean_Elab_Term_elabShow___closed__13; 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_Elab_Term_elabShow___closed__13; -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); if (lean_obj_tag(x_6) == 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; -x_26 = l_Lean_Elab_Term_expandHave___lambda__1___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; +x_24 = l_Lean_Elab_Term_expandHave___lambda__1___closed__1; lean_inc(x_15); -x_27 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_27, 0, x_15); -lean_ctor_set(x_27, 1, x_20); -lean_ctor_set(x_27, 2, x_26); -x_28 = l_Lean_Elab_Term_elabShow___closed__8; +x_25 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_25, 0, x_15); +lean_ctor_set(x_25, 1, x_20); +lean_ctor_set(x_25, 2, x_24); +x_26 = l_Lean_Elab_Term_expandHave___lambda__2___closed__2; lean_inc(x_15); -x_29 = l_Lean_Syntax_node5(x_15, x_28, x_4, x_21, x_27, x_23, x_10); -x_30 = l_Lean_Elab_Term_elabShow___closed__6; +x_27 = l_Lean_Syntax_node4(x_15, x_26, x_4, x_21, x_25, x_10); +x_28 = l_Lean_Elab_Term_elabShow___closed__6; lean_inc(x_15); -x_31 = l_Lean_Syntax_node1(x_15, x_30, x_29); -x_32 = l_Lean_Elab_Term_elabShow___closed__4; -x_33 = l_Lean_Syntax_node4(x_15, x_32, x_17, x_31, x_25, x_11); -x_34 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_34, 0, x_33); -lean_ctor_set(x_34, 1, x_8); -return x_34; +x_29 = l_Lean_Syntax_node1(x_15, x_28, x_27); +x_30 = l_Lean_Elab_Term_elabShow___closed__4; +x_31 = l_Lean_Syntax_node4(x_15, x_30, x_17, x_29, x_23, x_11); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_8); +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; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_35 = lean_ctor_get(x_6, 0); -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; +x_33 = lean_ctor_get(x_6, 0); +lean_inc(x_33); lean_dec(x_6); -x_36 = l_Lean_Elab_Term_elabShow___closed__11; +x_34 = l_Lean_Elab_Term_elabShow___closed__11; 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_elabShow___closed__10; +x_35 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_35, 0, x_15); +lean_ctor_set(x_35, 1, x_34); +x_36 = l_Lean_Elab_Term_elabShow___closed__10; lean_inc(x_15); -x_39 = l_Lean_Syntax_node2(x_15, x_38, x_37, x_35); -x_40 = l_Array_mkArray1___rarg(x_39); -x_41 = l_Array_append___rarg(x_18, x_40); +x_37 = l_Lean_Syntax_node2(x_15, x_36, x_35, x_33); +x_38 = l_Array_mkArray1___rarg(x_37); +x_39 = l_Array_append___rarg(x_18, x_38); lean_inc(x_15); -x_42 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_42, 0, x_15); -lean_ctor_set(x_42, 1, x_20); -lean_ctor_set(x_42, 2, x_41); -x_43 = l_Lean_Elab_Term_elabShow___closed__8; +x_40 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_40, 0, x_15); +lean_ctor_set(x_40, 1, x_20); +lean_ctor_set(x_40, 2, x_39); +x_41 = l_Lean_Elab_Term_expandHave___lambda__2___closed__2; lean_inc(x_15); -x_44 = l_Lean_Syntax_node5(x_15, x_43, x_4, x_21, x_42, x_23, x_10); -x_45 = l_Lean_Elab_Term_elabShow___closed__6; +x_42 = l_Lean_Syntax_node4(x_15, x_41, x_4, x_21, x_40, x_10); +x_43 = l_Lean_Elab_Term_elabShow___closed__6; lean_inc(x_15); -x_46 = l_Lean_Syntax_node1(x_15, x_45, x_44); -x_47 = l_Lean_Elab_Term_elabShow___closed__4; -x_48 = l_Lean_Syntax_node4(x_15, x_47, x_17, x_46, x_25, x_11); -x_49 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_49, 0, x_48); -lean_ctor_set(x_49, 1, x_8); -return x_49; +x_44 = l_Lean_Syntax_node1(x_15, x_43, x_42); +x_45 = l_Lean_Elab_Term_elabShow___closed__4; +x_46 = l_Lean_Syntax_node4(x_15, x_45, x_17, x_44, x_23, x_11); +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_8); +return x_47; } } } -static lean_object* _init_l_Lean_Elab_Term_expandHave___closed__1() { +static lean_object* _init_l_Lean_Elab_Term_expandHave___lambda__4___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___regBuiltin_Lean_Elab_Term_elabCoe___closed__3; -x_2 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__8; -x_3 = l___regBuiltin_Lean_Elab_Term_elabCoe___closed__5; -x_4 = l_Lean_Elab_Term_expandHave___lambda__2___closed__3; -x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); -return x_5; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("have", 4); +return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_expandHave___closed__2() { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandHave___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) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("haveDecl", 8); +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_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; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_12 = lean_unsigned_to_nat(3u); +x_13 = l_Lean_Syntax_getArg(x_1, x_12); +x_14 = l_Lean_Syntax_getArg(x_2, x_12); +x_15 = l_Lean_Syntax_getArgs(x_3); +x_16 = lean_ctor_get(x_10, 5); +lean_inc(x_16); +lean_dec(x_10); +x_17 = 0; +x_18 = l_Lean_SourceInfo_fromRef(x_16, x_17); +x_19 = l_Lean_Elab_Term_expandHave___lambda__4___closed__1; +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_Lean_Elab_Term_elabShow___closed__2; +x_22 = 1; +x_23 = l_Lean_HygieneInfo_mkIdent(x_4, x_21, x_22); +x_24 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__4; +x_25 = l_Array_append___rarg(x_24, x_15); +x_26 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; +lean_inc(x_18); +x_27 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_27, 0, x_18); +lean_ctor_set(x_27, 1, x_26); +lean_ctor_set(x_27, 2, x_25); +x_28 = l_Lean_Elab_Term_elabShow___closed__13; +lean_inc(x_18); +x_29 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_29, 0, x_18); +lean_ctor_set(x_29, 1, x_28); +if (lean_obj_tag(x_9) == 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 = l_Lean_Elab_Term_expandHave___lambda__1___closed__1; +lean_inc(x_18); +x_31 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_31, 0, x_18); +lean_ctor_set(x_31, 1, x_26); +lean_ctor_set(x_31, 2, x_30); +lean_inc(x_18); +x_32 = l_Lean_Syntax_node4(x_18, x_5, x_23, x_27, x_31, x_13); +lean_inc(x_18); +x_33 = l_Lean_Syntax_node1(x_18, x_6, x_32); +x_34 = l_Lean_Syntax_node4(x_18, x_7, x_20, x_33, x_29, x_14); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_11); +return x_35; +} +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_9, 0); +lean_inc(x_36); +lean_dec(x_9); +x_37 = l_Lean_Elab_Term_elabShow___closed__11; +lean_inc(x_18); +x_38 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_38, 0, x_18); +lean_ctor_set(x_38, 1, x_37); +x_39 = l_Lean_Elab_Term_elabShow___closed__10; +lean_inc(x_18); +x_40 = l_Lean_Syntax_node2(x_18, x_39, x_38, x_36); +x_41 = l_Array_mkArray1___rarg(x_40); +x_42 = l_Array_append___rarg(x_24, x_41); +lean_inc(x_18); +x_43 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_43, 0, x_18); +lean_ctor_set(x_43, 1, x_26); +lean_ctor_set(x_43, 2, x_42); +lean_inc(x_18); +x_44 = l_Lean_Syntax_node4(x_18, x_5, x_23, x_27, x_43, x_13); +lean_inc(x_18); +x_45 = l_Lean_Syntax_node1(x_18, x_6, x_44); +x_46 = l_Lean_Syntax_node4(x_18, x_7, x_20, x_45, x_29, x_14); +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_11); +return x_47; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandHave___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) { +_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; uint8_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t 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_unsigned_to_nat(4u); +x_11 = l_Lean_Syntax_getArg(x_1, x_10); +x_12 = lean_unsigned_to_nat(3u); +x_13 = l_Lean_Syntax_getArg(x_2, x_12); +x_14 = l_Lean_Syntax_getArgs(x_3); +x_15 = lean_ctor_get(x_8, 5); +lean_inc(x_15); +lean_dec(x_8); +x_16 = 0; +x_17 = l_Lean_SourceInfo_fromRef(x_15, x_16); +x_18 = l_Lean_Elab_Term_elabShow___closed__3; +lean_inc(x_17); +x_19 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_19, 1, x_18); +x_20 = 1; +x_21 = l_Lean_SourceInfo_fromRef(x_4, x_20); +x_22 = l_Lean_Elab_Term_expandHave___lambda__2___closed__3; +x_23 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +lean_inc(x_17); +x_24 = l_Lean_Syntax_node1(x_17, x_5, x_23); +x_25 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__4; +x_26 = l_Array_append___rarg(x_25, x_14); +x_27 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; +lean_inc(x_17); +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 = l_Lean_Elab_Term_elabShow___closed__12; +lean_inc(x_17); +x_30 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_30, 0, x_17); +lean_ctor_set(x_30, 1, x_29); +x_31 = l_Lean_Elab_Term_elabShow___closed__13; +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); +if (lean_obj_tag(x_7) == 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; +x_33 = l_Lean_Elab_Term_expandHave___lambda__1___closed__1; +lean_inc(x_17); +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_17); +lean_ctor_set(x_34, 1, x_27); +lean_ctor_set(x_34, 2, x_33); +x_35 = l_Lean_Elab_Term_elabShow___closed__8; +lean_inc(x_17); +x_36 = l_Lean_Syntax_node5(x_17, x_35, x_24, x_28, x_34, x_30, x_11); +x_37 = l_Lean_Elab_Term_elabShow___closed__6; +lean_inc(x_17); +x_38 = l_Lean_Syntax_node1(x_17, x_37, x_36); +x_39 = l_Lean_Elab_Term_elabShow___closed__4; +x_40 = l_Lean_Syntax_node4(x_17, x_39, x_19, x_38, x_32, x_13); +x_41 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_9); +return x_41; +} +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; 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_7, 0); +lean_inc(x_42); +lean_dec(x_7); +x_43 = l_Lean_Elab_Term_elabShow___closed__11; +lean_inc(x_17); +x_44 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_44, 0, x_17); +lean_ctor_set(x_44, 1, x_43); +x_45 = l_Lean_Elab_Term_elabShow___closed__10; +lean_inc(x_17); +x_46 = l_Lean_Syntax_node2(x_17, x_45, x_44, x_42); +x_47 = l_Array_mkArray1___rarg(x_46); +x_48 = l_Array_append___rarg(x_25, x_47); +lean_inc(x_17); +x_49 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_49, 0, x_17); +lean_ctor_set(x_49, 1, x_27); +lean_ctor_set(x_49, 2, x_48); +x_50 = l_Lean_Elab_Term_elabShow___closed__8; +lean_inc(x_17); +x_51 = l_Lean_Syntax_node5(x_17, x_50, x_24, x_28, x_49, x_30, x_11); +x_52 = l_Lean_Elab_Term_elabShow___closed__6; +lean_inc(x_17); +x_53 = l_Lean_Syntax_node1(x_17, x_52, x_51); +x_54 = l_Lean_Elab_Term_elabShow___closed__4; +x_55 = l_Lean_Syntax_node4(x_17, x_54, x_19, x_53, x_32, x_13); +x_56 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_56, 0, x_55); +lean_ctor_set(x_56, 1, x_9); +return x_56; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandHave___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) { +_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; 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; lean_object* x_25; lean_object* x_26; +x_9 = lean_unsigned_to_nat(4u); +x_10 = l_Lean_Syntax_getArg(x_1, x_9); +x_11 = lean_unsigned_to_nat(3u); +x_12 = l_Lean_Syntax_getArg(x_2, x_11); +x_13 = l_Lean_Syntax_getArgs(x_3); +x_14 = lean_ctor_get(x_7, 5); +lean_inc(x_14); +lean_dec(x_7); +x_15 = 0; +x_16 = l_Lean_SourceInfo_fromRef(x_14, x_15); +x_17 = l_Lean_Elab_Term_elabShow___closed__3; +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_Elab_Term_elabAnonymousCtor___lambda__3___closed__4; +x_20 = l_Array_append___rarg(x_19, x_13); +x_21 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; +lean_inc(x_16); +x_22 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_22, 0, x_16); +lean_ctor_set(x_22, 1, x_21); +lean_ctor_set(x_22, 2, x_20); +x_23 = l_Lean_Elab_Term_elabShow___closed__12; +lean_inc(x_16); +x_24 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_24, 0, x_16); +lean_ctor_set(x_24, 1, x_23); +x_25 = l_Lean_Elab_Term_elabShow___closed__13; +lean_inc(x_16); +x_26 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_26, 0, x_16); +lean_ctor_set(x_26, 1, x_25); +if (lean_obj_tag(x_6) == 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; +x_27 = l_Lean_Elab_Term_expandHave___lambda__1___closed__1; +lean_inc(x_16); +x_28 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_28, 0, x_16); +lean_ctor_set(x_28, 1, x_21); +lean_ctor_set(x_28, 2, x_27); +x_29 = l_Lean_Elab_Term_elabShow___closed__8; +lean_inc(x_16); +x_30 = l_Lean_Syntax_node5(x_16, x_29, x_4, x_22, x_28, x_24, x_10); +x_31 = l_Lean_Elab_Term_elabShow___closed__6; +lean_inc(x_16); +x_32 = l_Lean_Syntax_node1(x_16, x_31, x_30); +x_33 = l_Lean_Elab_Term_elabShow___closed__4; +x_34 = l_Lean_Syntax_node4(x_16, x_33, x_18, x_32, x_26, x_12); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_8); +return x_35; +} +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; +x_36 = lean_ctor_get(x_6, 0); +lean_inc(x_36); +lean_dec(x_6); +x_37 = l_Lean_Elab_Term_elabShow___closed__11; +lean_inc(x_16); +x_38 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_38, 0, x_16); +lean_ctor_set(x_38, 1, x_37); +x_39 = l_Lean_Elab_Term_elabShow___closed__10; +lean_inc(x_16); +x_40 = l_Lean_Syntax_node2(x_16, x_39, x_38, x_36); +x_41 = l_Array_mkArray1___rarg(x_40); +x_42 = l_Array_append___rarg(x_19, x_41); +lean_inc(x_16); +x_43 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_43, 0, x_16); +lean_ctor_set(x_43, 1, x_21); +lean_ctor_set(x_43, 2, x_42); +x_44 = l_Lean_Elab_Term_elabShow___closed__8; +lean_inc(x_16); +x_45 = l_Lean_Syntax_node5(x_16, x_44, x_4, x_22, x_43, x_24, x_10); +x_46 = l_Lean_Elab_Term_elabShow___closed__6; +lean_inc(x_16); +x_47 = l_Lean_Syntax_node1(x_16, x_46, x_45); +x_48 = l_Lean_Elab_Term_elabShow___closed__4; +x_49 = l_Lean_Syntax_node4(x_16, x_48, x_18, x_47, x_26, x_12); +x_50 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_50, 0, x_49); +lean_ctor_set(x_50, 1, x_8); +return x_50; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandHave___lambda__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) { +_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; uint8_t 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; +x_12 = lean_unsigned_to_nat(4u); +x_13 = l_Lean_Syntax_getArg(x_1, x_12); +x_14 = lean_unsigned_to_nat(3u); +x_15 = l_Lean_Syntax_getArg(x_2, x_14); +x_16 = l_Lean_Syntax_getArgs(x_3); +x_17 = lean_ctor_get(x_10, 5); +lean_inc(x_17); +lean_dec(x_10); +x_18 = 0; +x_19 = l_Lean_SourceInfo_fromRef(x_17, x_18); +x_20 = l_Lean_Elab_Term_expandHave___lambda__4___closed__1; +lean_inc(x_19); +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_Elab_Term_elabShow___closed__2; +x_23 = 1; +x_24 = l_Lean_HygieneInfo_mkIdent(x_4, x_22, x_23); +x_25 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__4; +x_26 = l_Array_append___rarg(x_25, x_16); +x_27 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; +lean_inc(x_19); +x_28 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_28, 0, x_19); +lean_ctor_set(x_28, 1, x_27); +lean_ctor_set(x_28, 2, x_26); +x_29 = l_Lean_Elab_Term_elabShow___closed__12; +lean_inc(x_19); +x_30 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_30, 0, x_19); +lean_ctor_set(x_30, 1, x_29); +x_31 = l_Lean_Elab_Term_elabShow___closed__13; +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); +if (lean_obj_tag(x_9) == 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; +x_33 = l_Lean_Elab_Term_expandHave___lambda__1___closed__1; +lean_inc(x_19); +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_19); +lean_ctor_set(x_34, 1, x_27); +lean_ctor_set(x_34, 2, x_33); +lean_inc(x_19); +x_35 = l_Lean_Syntax_node5(x_19, x_5, x_24, x_28, x_34, x_30, x_13); +lean_inc(x_19); +x_36 = l_Lean_Syntax_node1(x_19, x_6, x_35); +x_37 = l_Lean_Syntax_node4(x_19, x_7, x_21, x_36, x_32, x_15); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_11); +return x_38; +} +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; +x_39 = lean_ctor_get(x_9, 0); +lean_inc(x_39); +lean_dec(x_9); +x_40 = l_Lean_Elab_Term_elabShow___closed__11; +lean_inc(x_19); +x_41 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_41, 0, x_19); +lean_ctor_set(x_41, 1, x_40); +x_42 = l_Lean_Elab_Term_elabShow___closed__10; +lean_inc(x_19); +x_43 = l_Lean_Syntax_node2(x_19, x_42, x_41, x_39); +x_44 = l_Array_mkArray1___rarg(x_43); +x_45 = l_Array_append___rarg(x_25, x_44); +lean_inc(x_19); +x_46 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_46, 0, x_19); +lean_ctor_set(x_46, 1, x_27); +lean_ctor_set(x_46, 2, x_45); +lean_inc(x_19); +x_47 = l_Lean_Syntax_node5(x_19, x_5, x_24, x_28, x_46, x_30, x_13); +lean_inc(x_19); +x_48 = l_Lean_Syntax_node1(x_19, x_6, x_47); +x_49 = l_Lean_Syntax_node4(x_19, x_7, x_21, x_48, x_32, x_15); +x_50 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_50, 0, x_49); +lean_ctor_set(x_50, 1, x_11); +return x_50; +} +} +} +static lean_object* _init_l_Lean_Elab_Term_expandHave___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___regBuiltin_Lean_Elab_Term_elabCoe___closed__3; +x_2 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__8; +x_3 = l___regBuiltin_Lean_Elab_Term_elabCoe___closed__5; +x_4 = l_Lean_Elab_Term_expandHave___lambda__4___closed__1; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Elab_Term_expandHave___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("haveDecl", 8); return x_1; } } @@ -4970,6 +5141,62 @@ x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } +static lean_object* _init_l_Lean_Elab_Term_expandHave___closed__10() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("hygieneInfo", 11); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_expandHave___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_expandHave___closed__10; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Term_expandHave___closed__12() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("ident", 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_expandHave___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_expandHave___closed__12; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Term_expandHave___closed__14() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("hole", 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_expandHave___closed__15() { +_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_elabCoe___closed__3; +x_2 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__8; +x_3 = l___regBuiltin_Lean_Elab_Term_elabCoe___closed__5; +x_4 = l_Lean_Elab_Term_expandHave___closed__14; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandHave(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { @@ -4990,217 +5217,205 @@ 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); +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 = lean_unsigned_to_nat(1u); -x_11 = l_Lean_Syntax_getArg(x_1, x_10); -x_12 = l_Lean_Elab_Term_expandHave___closed__3; -lean_inc(x_11); -x_13 = l_Lean_Syntax_isOfKind(x_11, x_12); -if (x_13 == 0) +x_10 = l_Lean_Elab_Term_expandHave___closed__3; +lean_inc(x_9); +x_11 = l_Lean_Syntax_isOfKind(x_9, x_10); +if (x_11 == 0) { -lean_object* x_14; lean_object* x_15; -lean_dec(x_11); +lean_object* x_12; lean_object* x_13; 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_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_16; lean_object* x_17; uint8_t x_18; -x_16 = l_Lean_Syntax_getArg(x_11, x_8); -lean_dec(x_11); -x_17 = l_Lean_Elab_Term_expandHave___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(0u); +x_15 = l_Lean_Syntax_getArg(x_9, x_14); +lean_dec(x_9); +x_16 = l_Lean_Elab_Term_expandHave___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_Lean_Elab_Term_expandHave___closed__7; -lean_inc(x_16); -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_Lean_Elab_Term_expandHave___closed__7; +lean_inc(x_15); +x_19 = l_Lean_Syntax_isOfKind(x_15, x_18); +if (x_19 == 0) { -lean_object* x_21; uint8_t x_22; -lean_dec(x_9); -x_21 = l_Lean_Elab_Term_expandHave___closed__9; -lean_inc(x_16); -x_22 = l_Lean_Syntax_isOfKind(x_16, x_21); -if (x_22 == 0) +lean_object* x_20; uint8_t x_21; +x_20 = l_Lean_Elab_Term_expandHave___closed__9; +lean_inc(x_15); +x_21 = l_Lean_Syntax_isOfKind(x_15, x_20); +if (x_21 == 0) { -lean_object* x_23; lean_object* x_24; -lean_dec(x_16); +lean_object* x_22; lean_object* x_23; +lean_dec(x_15); 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_3); -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_3); +return x_23; } else { -lean_object* x_25; lean_object* x_26; uint8_t x_27; -x_25 = l_Lean_Syntax_getArg(x_16, x_8); -x_26 = l_Lean_Syntax_getArg(x_16, x_10); -x_27 = l_Lean_Syntax_matchesNull(x_26, x_8); -if (x_27 == 0) +lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_24 = l_Lean_Syntax_getArg(x_15, x_14); +x_25 = l_Lean_Syntax_getArg(x_15, x_8); +x_26 = l_Lean_Syntax_matchesNull(x_25, x_14); +if (x_26 == 0) { -lean_object* x_28; lean_object* x_29; -lean_dec(x_25); -lean_dec(x_16); +lean_object* x_27; lean_object* x_28; +lean_dec(x_24); +lean_dec(x_15); lean_dec(x_2); lean_dec(x_1); -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 = lean_unsigned_to_nat(2u); -x_31 = l_Lean_Syntax_getArg(x_16, x_30); -x_32 = l_Lean_Syntax_isNone(x_31); -if (x_32 == 0) +lean_object* x_29; lean_object* x_30; uint8_t x_31; +x_29 = lean_unsigned_to_nat(2u); +x_30 = l_Lean_Syntax_getArg(x_15, x_29); +x_31 = l_Lean_Syntax_isNone(x_30); +if (x_31 == 0) { -uint8_t x_33; -lean_inc(x_31); -x_33 = l_Lean_Syntax_matchesNull(x_31, x_10); -if (x_33 == 0) +uint8_t x_32; +lean_inc(x_30); +x_32 = l_Lean_Syntax_matchesNull(x_30, x_8); +if (x_32 == 0) { -lean_object* x_34; lean_object* x_35; -lean_dec(x_31); -lean_dec(x_25); -lean_dec(x_16); +lean_object* x_33; lean_object* x_34; +lean_dec(x_30); +lean_dec(x_24); +lean_dec(x_15); lean_dec(x_2); lean_dec(x_1); -x_34 = lean_box(1); -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(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; } else { -lean_object* x_36; lean_object* x_37; uint8_t x_38; -x_36 = l_Lean_Syntax_getArg(x_31, x_8); -lean_dec(x_31); -x_37 = l_Lean_Elab_Term_elabShow___closed__10; -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_30, x_14); +lean_dec(x_30); +x_36 = l_Lean_Elab_Term_elabShow___closed__10; +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_25); -lean_dec(x_16); +lean_object* x_38; lean_object* x_39; +lean_dec(x_35); +lean_dec(x_24); +lean_dec(x_15); lean_dec(x_2); -lean_dec(x_1); -x_39 = lean_box(1); -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; +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_3); +return x_39; } else { -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_41 = l_Lean_Syntax_getArg(x_36, x_10); -lean_dec(x_36); -x_42 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_42, 0, x_41); -x_43 = lean_box(0); -x_44 = l_Lean_Elab_Term_expandHave___lambda__1(x_16, x_1, x_21, x_25, x_43, x_42, x_2, x_3); +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_40 = l_Lean_Syntax_getArg(x_35, x_8); +lean_dec(x_35); +x_41 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_41, 0, x_40); +x_42 = lean_box(0); +x_43 = l_Lean_Elab_Term_expandHave___lambda__1(x_15, x_1, x_20, x_24, x_42, x_41, x_2, x_3); lean_dec(x_1); -lean_dec(x_16); -return x_44; +lean_dec(x_15); +return x_43; } } } else { -lean_object* x_45; lean_object* x_46; lean_object* x_47; -lean_dec(x_31); +lean_object* x_44; lean_object* x_45; lean_object* x_46; +lean_dec(x_30); +x_44 = lean_box(0); x_45 = lean_box(0); -x_46 = lean_box(0); -x_47 = l_Lean_Elab_Term_expandHave___lambda__1(x_16, x_1, x_21, x_25, x_46, x_45, x_2, x_3); +x_46 = l_Lean_Elab_Term_expandHave___lambda__1(x_15, x_1, x_20, x_24, x_45, x_44, x_2, x_3); lean_dec(x_1); -lean_dec(x_16); -return x_47; +lean_dec(x_15); +return x_46; } } } } else { -lean_object* x_48; lean_object* x_49; uint8_t x_50; -x_48 = l_Lean_Syntax_getArg(x_16, x_8); -x_49 = lean_unsigned_to_nat(2u); -lean_inc(x_48); -x_50 = l_Lean_Syntax_matchesNull(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_15, x_14); +x_48 = l_Lean_Elab_Term_expandHave___closed__11; +lean_inc(x_47); +x_49 = l_Lean_Syntax_isOfKind(x_47, x_48); +if (x_49 == 0) { -uint8_t x_51; -x_51 = l_Lean_Syntax_matchesNull(x_48, x_8); +lean_object* x_50; uint8_t x_51; +x_50 = l_Lean_Elab_Term_expandHave___closed__13; +lean_inc(x_47); +x_51 = l_Lean_Syntax_isOfKind(x_47, x_50); if (x_51 == 0) { -lean_object* x_52; lean_object* x_53; -lean_dec(x_16); -lean_dec(x_9); -lean_dec(x_2); -lean_dec(x_1); -x_52 = lean_box(1); -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; -} -else -{ -lean_object* x_54; uint8_t x_55; -x_54 = l_Lean_Syntax_getArg(x_16, x_10); -x_55 = l_Lean_Syntax_isNone(x_54); -if (x_55 == 0) -{ -uint8_t x_56; -lean_inc(x_54); -x_56 = l_Lean_Syntax_matchesNull(x_54, x_10); -if (x_56 == 0) +lean_object* x_52; uint8_t x_53; +x_52 = l_Lean_Elab_Term_expandHave___closed__15; +lean_inc(x_47); +x_53 = l_Lean_Syntax_isOfKind(x_47, x_52); +if (x_53 == 0) { -lean_object* x_57; lean_object* x_58; -lean_dec(x_54); -lean_dec(x_16); -lean_dec(x_9); +lean_object* x_54; lean_object* x_55; +lean_dec(x_47); +lean_dec(x_15); 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_54 = lean_box(1); +x_55 = lean_alloc_ctor(1, 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_59; lean_object* x_60; uint8_t x_61; -x_59 = l_Lean_Syntax_getArg(x_54, x_8); -lean_dec(x_54); -x_60 = l_Lean_Elab_Term_elabShow___closed__10; +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; uint8_t x_60; +x_56 = l_Lean_Syntax_getArg(x_47, x_14); +lean_dec(x_47); +x_57 = l_Lean_Syntax_getArg(x_15, x_8); +x_58 = lean_unsigned_to_nat(2u); +x_59 = l_Lean_Syntax_getArg(x_15, x_58); +x_60 = l_Lean_Syntax_isNone(x_59); +if (x_60 == 0) +{ +uint8_t x_61; lean_inc(x_59); -x_61 = l_Lean_Syntax_isOfKind(x_59, x_60); +x_61 = l_Lean_Syntax_matchesNull(x_59, x_8); if (x_61 == 0) { lean_object* x_62; lean_object* x_63; lean_dec(x_59); -lean_dec(x_16); -lean_dec(x_9); +lean_dec(x_57); +lean_dec(x_56); +lean_dec(x_15); lean_dec(x_2); lean_dec(x_1); x_62 = lean_box(1); @@ -5211,76 +5426,76 @@ return x_63; } else { -lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; -x_64 = l_Lean_Syntax_getArg(x_59, x_10); +lean_object* x_64; lean_object* x_65; uint8_t x_66; +x_64 = l_Lean_Syntax_getArg(x_59, x_14); lean_dec(x_59); -x_65 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_65, 0, x_64); -x_66 = lean_box(0); -x_67 = l_Lean_Elab_Term_expandHave___lambda__2(x_16, x_1, x_9, x_19, x_12, x_4, x_66, x_65, x_2, x_3); +x_65 = l_Lean_Elab_Term_elabShow___closed__10; +lean_inc(x_64); +x_66 = l_Lean_Syntax_isOfKind(x_64, x_65); +if (x_66 == 0) +{ +lean_object* x_67; lean_object* x_68; +lean_dec(x_64); +lean_dec(x_57); +lean_dec(x_56); +lean_dec(x_15); +lean_dec(x_2); lean_dec(x_1); -lean_dec(x_16); -return x_67; -} -} +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; } else { -lean_object* x_68; lean_object* x_69; lean_object* x_70; -lean_dec(x_54); -x_68 = lean_box(0); -x_69 = lean_box(0); -x_70 = l_Lean_Elab_Term_expandHave___lambda__2(x_16, x_1, x_9, x_19, x_12, x_4, x_69, x_68, x_2, x_3); +lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_69 = l_Lean_Syntax_getArg(x_64, x_8); +lean_dec(x_64); +x_70 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_70, 0, x_69); +x_71 = lean_box(0); +x_72 = l_Lean_Elab_Term_expandHave___lambda__2(x_15, x_1, x_57, x_56, x_52, x_71, x_70, x_2, x_3); +lean_dec(x_57); lean_dec(x_1); -lean_dec(x_16); -return x_70; +lean_dec(x_15); +return x_72; } } } else { -lean_object* x_71; lean_object* x_72; lean_object* x_73; uint8_t x_74; -lean_dec(x_9); -x_71 = l_Lean_Syntax_getArg(x_48, x_8); -x_72 = l_Lean_Syntax_getArg(x_48, x_10); -lean_dec(x_48); -x_73 = l_Lean_Syntax_getArg(x_16, x_10); -x_74 = l_Lean_Syntax_isNone(x_73); -if (x_74 == 0) -{ -uint8_t x_75; -lean_inc(x_73); -x_75 = l_Lean_Syntax_matchesNull(x_73, x_10); -if (x_75 == 0) -{ -lean_object* x_76; lean_object* x_77; -lean_dec(x_73); -lean_dec(x_72); -lean_dec(x_71); -lean_dec(x_16); -lean_dec(x_2); +lean_object* x_73; lean_object* x_74; lean_object* x_75; +lean_dec(x_59); +x_73 = lean_box(0); +x_74 = lean_box(0); +x_75 = l_Lean_Elab_Term_expandHave___lambda__2(x_15, x_1, x_57, x_56, x_52, x_74, x_73, x_2, x_3); +lean_dec(x_57); lean_dec(x_1); -x_76 = lean_box(1); -x_77 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_77, 0, x_76); -lean_ctor_set(x_77, 1, x_3); -return x_77; +lean_dec(x_15); +return x_75; +} +} } else { -lean_object* x_78; lean_object* x_79; uint8_t x_80; -x_78 = l_Lean_Syntax_getArg(x_73, x_8); -lean_dec(x_73); -x_79 = l_Lean_Elab_Term_elabShow___closed__10; +lean_object* x_76; lean_object* x_77; lean_object* x_78; uint8_t x_79; +x_76 = l_Lean_Syntax_getArg(x_15, x_8); +x_77 = lean_unsigned_to_nat(2u); +x_78 = l_Lean_Syntax_getArg(x_15, x_77); +x_79 = l_Lean_Syntax_isNone(x_78); +if (x_79 == 0) +{ +uint8_t x_80; lean_inc(x_78); -x_80 = l_Lean_Syntax_isOfKind(x_78, x_79); +x_80 = l_Lean_Syntax_matchesNull(x_78, x_8); if (x_80 == 0) { lean_object* x_81; lean_object* x_82; lean_dec(x_78); -lean_dec(x_72); -lean_dec(x_71); -lean_dec(x_16); +lean_dec(x_76); +lean_dec(x_47); +lean_dec(x_15); lean_dec(x_2); lean_dec(x_1); x_81 = lean_box(1); @@ -5291,212 +5506,413 @@ return x_82; } else { -lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; -x_83 = l_Lean_Syntax_getArg(x_78, x_10); +lean_object* x_83; lean_object* x_84; uint8_t x_85; +x_83 = l_Lean_Syntax_getArg(x_78, x_14); lean_dec(x_78); -x_84 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_84, 0, x_83); -x_85 = lean_box(0); -x_86 = l_Lean_Elab_Term_expandHave___lambda__3(x_16, x_1, x_72, x_71, x_85, x_84, x_2, x_3); -lean_dec(x_72); +x_84 = l_Lean_Elab_Term_elabShow___closed__10; +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_76); +lean_dec(x_47); +lean_dec(x_15); +lean_dec(x_2); lean_dec(x_1); -lean_dec(x_16); -return x_86; +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_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; +x_88 = l_Lean_Syntax_getArg(x_83, x_8); +lean_dec(x_83); +x_89 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_89, 0, x_88); +x_90 = lean_box(0); +x_91 = l_Lean_Elab_Term_expandHave___lambda__3(x_15, x_1, x_76, x_47, x_90, x_89, x_2, x_3); +lean_dec(x_76); +lean_dec(x_1); +lean_dec(x_15); +return x_91; } } } else { -lean_object* x_87; lean_object* x_88; lean_object* x_89; -lean_dec(x_73); -x_87 = lean_box(0); -x_88 = lean_box(0); -x_89 = l_Lean_Elab_Term_expandHave___lambda__3(x_16, x_1, x_72, x_71, x_88, x_87, x_2, x_3); -lean_dec(x_72); +lean_object* x_92; lean_object* x_93; lean_object* x_94; +lean_dec(x_78); +x_92 = lean_box(0); +x_93 = lean_box(0); +x_94 = l_Lean_Elab_Term_expandHave___lambda__3(x_15, x_1, x_76, x_47, x_93, x_92, x_2, x_3); +lean_dec(x_76); lean_dec(x_1); -lean_dec(x_16); -return x_89; +lean_dec(x_15); +return x_94; } } } +else +{ +lean_object* x_95; lean_object* x_96; lean_object* x_97; uint8_t x_98; +x_95 = l_Lean_Syntax_getArg(x_15, x_8); +x_96 = lean_unsigned_to_nat(2u); +x_97 = l_Lean_Syntax_getArg(x_15, x_96); +x_98 = l_Lean_Syntax_isNone(x_97); +if (x_98 == 0) +{ +uint8_t x_99; +lean_inc(x_97); +x_99 = l_Lean_Syntax_matchesNull(x_97, x_8); +if (x_99 == 0) +{ +lean_object* x_100; lean_object* x_101; +lean_dec(x_97); +lean_dec(x_95); +lean_dec(x_47); +lean_dec(x_15); +lean_dec(x_2); +lean_dec(x_1); +x_100 = lean_box(1); +x_101 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_101, 0, x_100); +lean_ctor_set(x_101, 1, x_3); +return x_101; } else { -lean_object* x_90; lean_object* x_91; uint8_t x_92; -x_90 = l_Lean_Syntax_getArg(x_16, x_8); -x_91 = lean_unsigned_to_nat(2u); -lean_inc(x_90); -x_92 = l_Lean_Syntax_matchesNull(x_90, x_91); -if (x_92 == 0) +lean_object* x_102; lean_object* x_103; uint8_t x_104; +x_102 = l_Lean_Syntax_getArg(x_97, x_14); +lean_dec(x_97); +x_103 = l_Lean_Elab_Term_elabShow___closed__10; +lean_inc(x_102); +x_104 = l_Lean_Syntax_isOfKind(x_102, x_103); +if (x_104 == 0) +{ +lean_object* x_105; lean_object* x_106; +lean_dec(x_102); +lean_dec(x_95); +lean_dec(x_47); +lean_dec(x_15); +lean_dec(x_2); +lean_dec(x_1); +x_105 = lean_box(1); +x_106 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_106, 0, x_105); +lean_ctor_set(x_106, 1, x_3); +return x_106; +} +else { -uint8_t x_93; -x_93 = l_Lean_Syntax_matchesNull(x_90, x_8); -if (x_93 == 0) +lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; +x_107 = l_Lean_Syntax_getArg(x_102, x_8); +lean_dec(x_102); +x_108 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_108, 0, x_107); +x_109 = lean_box(0); +x_110 = l_Lean_Elab_Term_expandHave___lambda__4(x_15, x_1, x_95, x_47, x_18, x_10, x_4, x_109, x_108, x_2, x_3); +lean_dec(x_47); +lean_dec(x_95); +lean_dec(x_1); +lean_dec(x_15); +return x_110; +} +} +} +else { -lean_object* x_94; lean_object* x_95; -lean_dec(x_16); -lean_dec(x_9); +lean_object* x_111; lean_object* x_112; lean_object* x_113; +lean_dec(x_97); +x_111 = lean_box(0); +x_112 = lean_box(0); +x_113 = l_Lean_Elab_Term_expandHave___lambda__4(x_15, x_1, x_95, x_47, x_18, x_10, x_4, x_112, x_111, x_2, x_3); +lean_dec(x_47); +lean_dec(x_95); +lean_dec(x_1); +lean_dec(x_15); +return x_113; +} +} +} +} +else +{ +lean_object* x_114; lean_object* x_115; uint8_t x_116; +x_114 = l_Lean_Syntax_getArg(x_15, x_14); +x_115 = l_Lean_Elab_Term_expandHave___closed__11; +lean_inc(x_114); +x_116 = l_Lean_Syntax_isOfKind(x_114, x_115); +if (x_116 == 0) +{ +lean_object* x_117; uint8_t x_118; +x_117 = l_Lean_Elab_Term_expandHave___closed__13; +lean_inc(x_114); +x_118 = l_Lean_Syntax_isOfKind(x_114, x_117); +if (x_118 == 0) +{ +lean_object* x_119; uint8_t x_120; +x_119 = l_Lean_Elab_Term_expandHave___closed__15; +lean_inc(x_114); +x_120 = l_Lean_Syntax_isOfKind(x_114, x_119); +if (x_120 == 0) +{ +lean_object* x_121; lean_object* x_122; +lean_dec(x_114); +lean_dec(x_15); lean_dec(x_2); lean_dec(x_1); -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; +x_121 = lean_box(1); +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; } else { -lean_object* x_96; uint8_t x_97; -x_96 = l_Lean_Syntax_getArg(x_16, x_10); -x_97 = l_Lean_Syntax_isNone(x_96); -if (x_97 == 0) +lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; uint8_t x_127; +x_123 = l_Lean_Syntax_getArg(x_114, x_14); +lean_dec(x_114); +x_124 = l_Lean_Syntax_getArg(x_15, x_8); +x_125 = lean_unsigned_to_nat(2u); +x_126 = l_Lean_Syntax_getArg(x_15, x_125); +x_127 = l_Lean_Syntax_isNone(x_126); +if (x_127 == 0) { -uint8_t x_98; -lean_inc(x_96); -x_98 = l_Lean_Syntax_matchesNull(x_96, x_10); -if (x_98 == 0) +uint8_t x_128; +lean_inc(x_126); +x_128 = l_Lean_Syntax_matchesNull(x_126, x_8); +if (x_128 == 0) { -lean_object* x_99; lean_object* x_100; -lean_dec(x_96); -lean_dec(x_16); -lean_dec(x_9); +lean_object* x_129; lean_object* x_130; +lean_dec(x_126); +lean_dec(x_124); +lean_dec(x_123); +lean_dec(x_15); lean_dec(x_2); lean_dec(x_1); -x_99 = lean_box(1); -x_100 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_100, 0, x_99); -lean_ctor_set(x_100, 1, x_3); -return x_100; +x_129 = lean_box(1); +x_130 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_130, 0, x_129); +lean_ctor_set(x_130, 1, x_3); +return x_130; +} +else +{ +lean_object* x_131; lean_object* x_132; uint8_t x_133; +x_131 = l_Lean_Syntax_getArg(x_126, x_14); +lean_dec(x_126); +x_132 = l_Lean_Elab_Term_elabShow___closed__10; +lean_inc(x_131); +x_133 = l_Lean_Syntax_isOfKind(x_131, x_132); +if (x_133 == 0) +{ +lean_object* x_134; lean_object* x_135; +lean_dec(x_131); +lean_dec(x_124); +lean_dec(x_123); +lean_dec(x_15); +lean_dec(x_2); +lean_dec(x_1); +x_134 = lean_box(1); +x_135 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_135, 0, x_134); +lean_ctor_set(x_135, 1, x_3); +return x_135; } else { -lean_object* x_101; lean_object* x_102; uint8_t x_103; -x_101 = l_Lean_Syntax_getArg(x_96, x_8); -lean_dec(x_96); -x_102 = l_Lean_Elab_Term_elabShow___closed__10; -lean_inc(x_101); -x_103 = l_Lean_Syntax_isOfKind(x_101, x_102); -if (x_103 == 0) +lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; +x_136 = l_Lean_Syntax_getArg(x_131, x_8); +lean_dec(x_131); +x_137 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_137, 0, x_136); +x_138 = lean_box(0); +x_139 = l_Lean_Elab_Term_expandHave___lambda__5(x_15, x_1, x_124, x_123, x_119, x_138, x_137, x_2, x_3); +lean_dec(x_124); +lean_dec(x_1); +lean_dec(x_15); +return x_139; +} +} +} +else { -lean_object* x_104; lean_object* x_105; -lean_dec(x_101); -lean_dec(x_16); -lean_dec(x_9); +lean_object* x_140; lean_object* x_141; lean_object* x_142; +lean_dec(x_126); +x_140 = lean_box(0); +x_141 = lean_box(0); +x_142 = l_Lean_Elab_Term_expandHave___lambda__5(x_15, x_1, x_124, x_123, x_119, x_141, x_140, x_2, x_3); +lean_dec(x_124); +lean_dec(x_1); +lean_dec(x_15); +return x_142; +} +} +} +else +{ +lean_object* x_143; lean_object* x_144; lean_object* x_145; uint8_t x_146; +x_143 = l_Lean_Syntax_getArg(x_15, x_8); +x_144 = lean_unsigned_to_nat(2u); +x_145 = l_Lean_Syntax_getArg(x_15, x_144); +x_146 = l_Lean_Syntax_isNone(x_145); +if (x_146 == 0) +{ +uint8_t x_147; +lean_inc(x_145); +x_147 = l_Lean_Syntax_matchesNull(x_145, x_8); +if (x_147 == 0) +{ +lean_object* x_148; lean_object* x_149; +lean_dec(x_145); +lean_dec(x_143); +lean_dec(x_114); +lean_dec(x_15); 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_148 = lean_box(1); +x_149 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_149, 0, x_148); +lean_ctor_set(x_149, 1, x_3); +return x_149; } else { -lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; -x_106 = l_Lean_Syntax_getArg(x_101, x_10); -lean_dec(x_101); -x_107 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_107, 0, x_106); -x_108 = lean_box(0); -x_109 = l_Lean_Elab_Term_expandHave___lambda__4(x_16, x_1, x_9, x_17, x_12, x_4, x_108, x_107, x_2, x_3); +lean_object* x_150; lean_object* x_151; uint8_t x_152; +x_150 = l_Lean_Syntax_getArg(x_145, x_14); +lean_dec(x_145); +x_151 = l_Lean_Elab_Term_elabShow___closed__10; +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_dec(x_143); +lean_dec(x_114); +lean_dec(x_15); +lean_dec(x_2); lean_dec(x_1); -lean_dec(x_16); -return x_109; +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; +} +else +{ +lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; +x_155 = l_Lean_Syntax_getArg(x_150, x_8); +lean_dec(x_150); +x_156 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_156, 0, x_155); +x_157 = lean_box(0); +x_158 = l_Lean_Elab_Term_expandHave___lambda__6(x_15, x_1, x_143, x_114, x_157, x_156, x_2, x_3); +lean_dec(x_143); +lean_dec(x_1); +lean_dec(x_15); +return x_158; } } } else { -lean_object* x_110; lean_object* x_111; lean_object* x_112; -lean_dec(x_96); -x_110 = lean_box(0); -x_111 = lean_box(0); -x_112 = l_Lean_Elab_Term_expandHave___lambda__4(x_16, x_1, x_9, x_17, x_12, x_4, x_111, x_110, x_2, x_3); +lean_object* x_159; lean_object* x_160; lean_object* x_161; +lean_dec(x_145); +x_159 = lean_box(0); +x_160 = lean_box(0); +x_161 = l_Lean_Elab_Term_expandHave___lambda__6(x_15, x_1, x_143, x_114, x_160, x_159, x_2, x_3); +lean_dec(x_143); lean_dec(x_1); -lean_dec(x_16); -return x_112; +lean_dec(x_15); +return x_161; } } } else { -lean_object* x_113; lean_object* x_114; lean_object* x_115; uint8_t x_116; -lean_dec(x_9); -x_113 = l_Lean_Syntax_getArg(x_90, x_8); -x_114 = l_Lean_Syntax_getArg(x_90, x_10); -lean_dec(x_90); -x_115 = l_Lean_Syntax_getArg(x_16, x_10); -x_116 = l_Lean_Syntax_isNone(x_115); -if (x_116 == 0) +lean_object* x_162; lean_object* x_163; lean_object* x_164; uint8_t x_165; +x_162 = l_Lean_Syntax_getArg(x_15, x_8); +x_163 = lean_unsigned_to_nat(2u); +x_164 = l_Lean_Syntax_getArg(x_15, x_163); +x_165 = l_Lean_Syntax_isNone(x_164); +if (x_165 == 0) { -uint8_t x_117; -lean_inc(x_115); -x_117 = l_Lean_Syntax_matchesNull(x_115, x_10); -if (x_117 == 0) +uint8_t x_166; +lean_inc(x_164); +x_166 = l_Lean_Syntax_matchesNull(x_164, x_8); +if (x_166 == 0) { -lean_object* x_118; lean_object* x_119; -lean_dec(x_115); +lean_object* x_167; lean_object* x_168; +lean_dec(x_164); +lean_dec(x_162); lean_dec(x_114); -lean_dec(x_113); -lean_dec(x_16); +lean_dec(x_15); lean_dec(x_2); lean_dec(x_1); -x_118 = lean_box(1); -x_119 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_119, 0, x_118); -lean_ctor_set(x_119, 1, x_3); -return x_119; +x_167 = lean_box(1); +x_168 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_168, 0, x_167); +lean_ctor_set(x_168, 1, x_3); +return x_168; } else { -lean_object* x_120; lean_object* x_121; uint8_t x_122; -x_120 = l_Lean_Syntax_getArg(x_115, x_8); -lean_dec(x_115); -x_121 = l_Lean_Elab_Term_elabShow___closed__10; -lean_inc(x_120); -x_122 = l_Lean_Syntax_isOfKind(x_120, x_121); -if (x_122 == 0) +lean_object* x_169; lean_object* x_170; uint8_t x_171; +x_169 = l_Lean_Syntax_getArg(x_164, x_14); +lean_dec(x_164); +x_170 = l_Lean_Elab_Term_elabShow___closed__10; +lean_inc(x_169); +x_171 = l_Lean_Syntax_isOfKind(x_169, x_170); +if (x_171 == 0) { -lean_object* x_123; lean_object* x_124; -lean_dec(x_120); +lean_object* x_172; lean_object* x_173; +lean_dec(x_169); +lean_dec(x_162); lean_dec(x_114); -lean_dec(x_113); -lean_dec(x_16); +lean_dec(x_15); lean_dec(x_2); lean_dec(x_1); -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; +x_172 = lean_box(1); +x_173 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_173, 0, x_172); +lean_ctor_set(x_173, 1, x_3); +return x_173; } else { -lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; -x_125 = l_Lean_Syntax_getArg(x_120, x_10); -lean_dec(x_120); -x_126 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_126, 0, x_125); -x_127 = lean_box(0); -x_128 = l_Lean_Elab_Term_expandHave___lambda__5(x_16, x_1, x_114, x_113, x_127, x_126, x_2, x_3); +lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; +x_174 = l_Lean_Syntax_getArg(x_169, x_8); +lean_dec(x_169); +x_175 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_175, 0, x_174); +x_176 = lean_box(0); +x_177 = l_Lean_Elab_Term_expandHave___lambda__7(x_15, x_1, x_162, x_114, x_16, x_10, x_4, x_176, x_175, x_2, x_3); lean_dec(x_114); +lean_dec(x_162); lean_dec(x_1); -lean_dec(x_16); -return x_128; +lean_dec(x_15); +return x_177; } } } else { -lean_object* x_129; lean_object* x_130; lean_object* x_131; -lean_dec(x_115); -x_129 = lean_box(0); -x_130 = lean_box(0); -x_131 = l_Lean_Elab_Term_expandHave___lambda__5(x_16, x_1, x_114, x_113, x_130, x_129, x_2, x_3); +lean_object* x_178; lean_object* x_179; lean_object* x_180; +lean_dec(x_164); +x_178 = lean_box(0); +x_179 = lean_box(0); +x_180 = l_Lean_Elab_Term_expandHave___lambda__7(x_15, x_1, x_162, x_114, x_16, x_10, x_4, x_179, x_178, x_2, x_3); lean_dec(x_114); +lean_dec(x_162); lean_dec(x_1); -lean_dec(x_16); -return x_131; +lean_dec(x_15); +return x_180; } } } @@ -5515,15 +5931,16 @@ lean_dec(x_1); return x_9; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandHave___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_Lean_Elab_Term_expandHave___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) { _start: { -lean_object* x_11; -x_11 = l_Lean_Elab_Term_expandHave___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); +lean_object* x_10; +x_10 = l_Lean_Elab_Term_expandHave___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_6); +lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -return x_11; +return x_10; } } LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandHave___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) { @@ -5538,22 +5955,36 @@ lean_dec(x_1); return x_9; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandHave___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_EXPORT lean_object* l_Lean_Elab_Term_expandHave___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) { _start: { -lean_object* x_11; -x_11 = l_Lean_Elab_Term_expandHave___lambda__4(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_object* x_12; +x_12 = l_Lean_Elab_Term_expandHave___lambda__4(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_8); +lean_dec(x_4); +lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -return x_11; +return x_12; +} } +LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandHave___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) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_Elab_Term_expandHave___lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_6); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_10; } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandHave___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_EXPORT lean_object* l_Lean_Elab_Term_expandHave___lambda__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_9; -x_9 = l_Lean_Elab_Term_expandHave___lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_9 = l_Lean_Elab_Term_expandHave___lambda__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); @@ -5561,6 +5992,19 @@ lean_dec(x_1); return x_9; } } +LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandHave___lambda__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) { +_start: +{ +lean_object* x_12; +x_12 = l_Lean_Elab_Term_expandHave___lambda__7(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_8); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_12; +} +} static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandHave___closed__1() { _start: { @@ -5617,8 +6061,8 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandHave_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(112u); -x_2 = lean_unsigned_to_nat(83u); +x_1 = lean_unsigned_to_nat(116u); +x_2 = lean_unsigned_to_nat(78u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -5632,7 +6076,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_expandHave_declRange___closed__1; x_2 = lean_unsigned_to_nat(39u); x_3 = l___regBuiltin_Lean_Elab_Term_expandHave_declRange___closed__2; -x_4 = lean_unsigned_to_nat(83u); +x_4 = lean_unsigned_to_nat(78u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -5703,7 +6147,65 @@ x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -static lean_object* _init_l_Lean_Elab_Term_expandSuffices___lambda__1___closed__1() { +static lean_object* _init_l_Lean_Elab_Term_expandSuffices___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("suffices", 8); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_expandSuffices___closed__2() { +_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_elabCoe___closed__3; +x_2 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__8; +x_3 = l___regBuiltin_Lean_Elab_Term_elabCoe___closed__5; +x_4 = l_Lean_Elab_Term_expandSuffices___closed__1; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Elab_Term_expandSuffices___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("sufficesDecl", 12); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_expandSuffices___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___regBuiltin_Lean_Elab_Term_elabCoe___closed__3; +x_2 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__8; +x_3 = l___regBuiltin_Lean_Elab_Term_elabCoe___closed__5; +x_4 = l_Lean_Elab_Term_expandSuffices___closed__3; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Elab_Term_expandSuffices___closed__5() { +_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_expandSuffices___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_expandSuffices___closed__5; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Term_expandSuffices___closed__7() { _start: { lean_object* x_1; @@ -5711,7 +6213,7 @@ x_1 = lean_mk_string_from_bytes("Tactic", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_expandSuffices___lambda__1___closed__2() { +static lean_object* _init_l_Lean_Elab_Term_expandSuffices___closed__8() { _start: { lean_object* x_1; @@ -5719,416 +6221,684 @@ x_1 = lean_mk_string_from_bytes("tacticSeq", 9); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_expandSuffices___lambda__1___closed__3() { +static lean_object* _init_l_Lean_Elab_Term_expandSuffices___closed__9() { _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_elabCoe___closed__3; x_2 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__8; -x_3 = l_Lean_Elab_Term_expandSuffices___lambda__1___closed__1; -x_4 = l_Lean_Elab_Term_expandSuffices___lambda__1___closed__2; +x_3 = l_Lean_Elab_Term_expandSuffices___closed__7; +x_4 = l_Lean_Elab_Term_expandSuffices___closed__8; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandSuffices___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_Lean_Elab_Term_expandSuffices(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_expandSuffices___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; uint8_t x_13; -x_8 = lean_unsigned_to_nat(1u); +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_10 = lean_unsigned_to_nat(1u); x_11 = l_Lean_Syntax_getArg(x_1, x_10); -x_12 = l_Lean_Elab_Term_expandShow___closed__6; +x_12 = l_Lean_Elab_Term_expandSuffices___closed__4; lean_inc(x_11); x_13 = l_Lean_Syntax_isOfKind(x_11, x_12); if (x_13 == 0) { -lean_object* x_14; uint8_t x_15; -x_14 = l_Lean_Elab_Term_expandShow___closed__4; -lean_inc(x_11); -x_15 = l_Lean_Syntax_isOfKind(x_11, x_14); -if (x_15 == 0) +lean_object* x_14; lean_object* x_15; +lean_dec(x_11); +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; +} +else { -lean_object* x_16; lean_object* x_17; +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_Lean_Elab_Term_expandSuffices___closed__6; +lean_inc(x_16); +x_18 = l_Lean_Syntax_isOfKind(x_16, x_17); +if (x_18 == 0) +{ +lean_object* x_19; uint8_t x_20; +x_19 = l_Lean_Elab_Term_expandHave___closed__11; +lean_inc(x_16); +x_20 = l_Lean_Syntax_isOfKind(x_16, x_19); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; +lean_dec(x_16); lean_dec(x_11); lean_dec(x_9); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -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_7); -return x_17; +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_3); +return x_22; } else { -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_11, x_18); -x_20 = l_Lean_Syntax_getArg(x_11, x_8); +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; +x_23 = l_Lean_Syntax_getArg(x_11, x_10); +x_24 = lean_unsigned_to_nat(2u); +x_25 = l_Lean_Syntax_getArg(x_11, x_24); lean_dec(x_11); -x_21 = l_Lean_Elab_Term_expandSuffices___lambda__1___closed__3; -lean_inc(x_20); -x_22 = l_Lean_Syntax_isOfKind(x_20, x_21); -if (x_22 == 0) +x_26 = l_Lean_Elab_Term_expandShow___closed__6; +lean_inc(x_25); +x_27 = l_Lean_Syntax_isOfKind(x_25, x_26); +if (x_27 == 0) { -lean_object* x_23; lean_object* x_24; -lean_dec(x_20); -lean_dec(x_19); +lean_object* x_28; uint8_t x_29; +x_28 = l_Lean_Elab_Term_expandShow___closed__4; +lean_inc(x_25); +x_29 = l_Lean_Syntax_isOfKind(x_25, x_28); +if (x_29 == 0) +{ +lean_object* x_30; lean_object* x_31; +lean_dec(x_25); +lean_dec(x_23); +lean_dec(x_16); lean_dec(x_9); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -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; +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; } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; lean_object* x_29; 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; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; 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_25 = lean_unsigned_to_nat(3u); -x_26 = l_Lean_Syntax_getArg(x_2, x_25); -x_27 = lean_ctor_get(x_6, 5); -lean_inc(x_27); -lean_dec(x_6); -x_28 = 0; -x_29 = l_Lean_SourceInfo_fromRef(x_27, x_28); -x_30 = 1; -x_31 = l_Lean_SourceInfo_fromRef(x_3, x_30); -x_32 = l_Lean_Elab_Term_expandHave___lambda__2___closed__3; -x_33 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_33, 0, x_31); -lean_ctor_set(x_33, 1, x_32); -x_34 = l_Lean_Elab_Term_elabShow___closed__11; -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_Elab_Term_elabShow___closed__10; -lean_inc(x_29); -x_37 = l_Lean_Syntax_node2(x_29, x_36, x_35, x_9); -x_38 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; -lean_inc(x_29); -x_39 = l_Lean_Syntax_node1(x_29, x_38, x_37); -x_40 = l_Lean_Elab_Term_elabShow___closed__12; -lean_inc(x_29); -x_41 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_41, 0, x_29); -lean_ctor_set(x_41, 1, x_40); -x_42 = l_Lean_Elab_Term_elabShow___closed__13; -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_SourceInfo_fromRef(x_19, x_30); -x_45 = l_Lean_Elab_Term_expandShow___closed__10; -x_46 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -x_47 = l_Lean_Elab_Term_expandShow___closed__9; -lean_inc(x_29); -x_48 = l_Lean_Syntax_node2(x_29, x_47, x_46, x_20); -if (lean_obj_tag(x_5) == 0) +lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_32 = l_Lean_Syntax_getArg(x_25, x_8); +x_33 = l_Lean_Syntax_getArg(x_25, x_10); +lean_dec(x_25); +x_34 = l_Lean_Elab_Term_expandSuffices___closed__9; +lean_inc(x_33); +x_35 = l_Lean_Syntax_isOfKind(x_33, x_34); +if (x_35 == 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; -x_49 = l_Lean_Elab_Term_expandHave___lambda__1___closed__1; -lean_inc(x_29); -x_50 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_50, 0, x_29); -lean_ctor_set(x_50, 1, x_38); -lean_ctor_set(x_50, 2, x_49); -x_51 = l_Lean_Elab_Term_expandHave___closed__5; -lean_inc(x_29); -x_52 = l_Lean_Syntax_node4(x_29, x_51, x_50, x_39, x_41, x_26); -x_53 = l_Lean_Elab_Term_expandHave___closed__3; -lean_inc(x_29); -x_54 = l_Lean_Syntax_node1(x_29, x_53, x_52); -x_55 = l_Lean_Elab_Term_expandHave___closed__1; -x_56 = l_Lean_Syntax_node4(x_29, x_55, x_33, x_54, x_43, x_48); -x_57 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_57, 0, x_56); -lean_ctor_set(x_57, 1, x_7); -return x_57; +lean_object* x_36; lean_object* x_37; +lean_dec(x_33); +lean_dec(x_32); +lean_dec(x_23); +lean_dec(x_16); +lean_dec(x_9); +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; } 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; -x_58 = lean_ctor_get(x_5, 0); -lean_inc(x_58); -lean_dec(x_5); -x_59 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__4; -lean_inc(x_29); -x_60 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_60, 0, x_29); -lean_ctor_set(x_60, 1, x_38); -lean_ctor_set(x_60, 2, x_59); -x_61 = l_Array_mkArray2___rarg(x_58, x_60); -x_62 = l_Array_append___rarg(x_59, x_61); -lean_inc(x_29); -x_63 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_63, 0, x_29); -lean_ctor_set(x_63, 1, x_38); -lean_ctor_set(x_63, 2, x_62); -x_64 = l_Lean_Elab_Term_expandHave___closed__5; -lean_inc(x_29); -x_65 = l_Lean_Syntax_node4(x_29, x_64, x_63, x_39, x_41, x_26); -x_66 = l_Lean_Elab_Term_expandHave___closed__3; -lean_inc(x_29); -x_67 = l_Lean_Syntax_node1(x_29, x_66, x_65); +lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; lean_object* x_42; uint8_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; 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; +x_38 = lean_unsigned_to_nat(3u); +x_39 = l_Lean_Syntax_getArg(x_1, x_38); +lean_dec(x_1); +x_40 = lean_ctor_get(x_2, 5); +lean_inc(x_40); +lean_dec(x_2); +x_41 = 0; +x_42 = l_Lean_SourceInfo_fromRef(x_40, x_41); +x_43 = 1; +x_44 = l_Lean_SourceInfo_fromRef(x_9, x_43); +x_45 = l_Lean_Elab_Term_expandHave___lambda__4___closed__1; +x_46 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +x_47 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; +x_48 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__4; +lean_inc(x_42); +x_49 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_49, 0, x_42); +lean_ctor_set(x_49, 1, x_47); +lean_ctor_set(x_49, 2, x_48); +x_50 = l_Lean_Elab_Term_elabShow___closed__11; +lean_inc(x_42); +x_51 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_51, 0, x_42); +lean_ctor_set(x_51, 1, x_50); +x_52 = l_Lean_Elab_Term_elabShow___closed__10; +lean_inc(x_42); +x_53 = l_Lean_Syntax_node2(x_42, x_52, x_51, x_23); +lean_inc(x_42); +x_54 = l_Lean_Syntax_node1(x_42, x_47, x_53); +x_55 = l_Lean_Elab_Term_elabShow___closed__12; +lean_inc(x_42); +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_Elab_Term_expandHave___closed__5; +lean_inc(x_42); +x_58 = l_Lean_Syntax_node5(x_42, x_57, x_16, x_49, x_54, x_56, x_39); +x_59 = l_Lean_Elab_Term_expandHave___closed__3; +lean_inc(x_42); +x_60 = l_Lean_Syntax_node1(x_42, x_59, x_58); +x_61 = l_Lean_Elab_Term_elabShow___closed__13; +lean_inc(x_42); +x_62 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_62, 0, x_42); +lean_ctor_set(x_62, 1, x_61); +x_63 = l_Lean_SourceInfo_fromRef(x_32, x_43); +x_64 = l_Lean_Elab_Term_expandShow___closed__10; +x_65 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_65, 0, x_63); +lean_ctor_set(x_65, 1, x_64); +x_66 = l_Lean_Elab_Term_expandShow___closed__9; +lean_inc(x_42); +x_67 = l_Lean_Syntax_node2(x_42, x_66, x_65, x_33); x_68 = l_Lean_Elab_Term_expandHave___closed__1; -x_69 = l_Lean_Syntax_node4(x_29, x_68, x_33, x_67, x_43, x_48); +x_69 = l_Lean_Syntax_node4(x_42, x_68, x_46, x_60, x_62, x_67); x_70 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_70, 0, x_69); -lean_ctor_set(x_70, 1, x_7); +lean_ctor_set(x_70, 1, x_3); return x_70; } } } -} else { -lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; uint8_t x_75; lean_object* x_76; 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; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; -x_71 = l_Lean_Syntax_getArg(x_11, x_8); -lean_dec(x_11); +lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; uint8_t x_75; lean_object* x_76; 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; 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_71 = l_Lean_Syntax_getArg(x_25, x_10); +lean_dec(x_25); x_72 = lean_unsigned_to_nat(3u); -x_73 = l_Lean_Syntax_getArg(x_2, x_72); -x_74 = lean_ctor_get(x_6, 5); +x_73 = l_Lean_Syntax_getArg(x_1, x_72); +lean_dec(x_1); +x_74 = lean_ctor_get(x_2, 5); lean_inc(x_74); -lean_dec(x_6); +lean_dec(x_2); x_75 = 0; x_76 = l_Lean_SourceInfo_fromRef(x_74, x_75); x_77 = 1; -x_78 = l_Lean_SourceInfo_fromRef(x_3, x_77); -x_79 = l_Lean_Elab_Term_expandHave___lambda__2___closed__3; +x_78 = l_Lean_SourceInfo_fromRef(x_9, x_77); +x_79 = l_Lean_Elab_Term_expandHave___lambda__4___closed__1; 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_Elab_Term_elabShow___closed__11; +x_81 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; +x_82 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__4; lean_inc(x_76); -x_82 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_82, 0, x_76); -lean_ctor_set(x_82, 1, x_81); -x_83 = l_Lean_Elab_Term_elabShow___closed__10; +x_83 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_83, 0, x_76); +lean_ctor_set(x_83, 1, x_81); +lean_ctor_set(x_83, 2, x_82); +x_84 = l_Lean_Elab_Term_elabShow___closed__11; lean_inc(x_76); -x_84 = l_Lean_Syntax_node2(x_76, x_83, x_82, x_9); -x_85 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; +x_85 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_85, 0, x_76); +lean_ctor_set(x_85, 1, x_84); +x_86 = l_Lean_Elab_Term_elabShow___closed__10; lean_inc(x_76); -x_86 = l_Lean_Syntax_node1(x_76, x_85, x_84); -x_87 = l_Lean_Elab_Term_elabShow___closed__12; +x_87 = l_Lean_Syntax_node2(x_76, x_86, x_85, x_23); lean_inc(x_76); -x_88 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_88, 0, x_76); -lean_ctor_set(x_88, 1, x_87); -x_89 = l_Lean_Elab_Term_elabShow___closed__13; +x_88 = l_Lean_Syntax_node1(x_76, x_81, x_87); +x_89 = l_Lean_Elab_Term_elabShow___closed__12; 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); -if (lean_obj_tag(x_5) == 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; lean_object* x_99; -x_91 = l_Lean_Elab_Term_expandHave___lambda__1___closed__1; +x_91 = l_Lean_Elab_Term_expandHave___closed__5; lean_inc(x_76); -x_92 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_92, 0, x_76); -lean_ctor_set(x_92, 1, x_85); -lean_ctor_set(x_92, 2, x_91); -x_93 = l_Lean_Elab_Term_expandHave___closed__5; +x_92 = l_Lean_Syntax_node5(x_76, x_91, x_16, x_83, x_88, x_90, x_73); +x_93 = l_Lean_Elab_Term_expandHave___closed__3; lean_inc(x_76); -x_94 = l_Lean_Syntax_node4(x_76, x_93, x_92, x_86, x_88, x_73); -x_95 = l_Lean_Elab_Term_expandHave___closed__3; +x_94 = l_Lean_Syntax_node1(x_76, x_93, x_92); +x_95 = l_Lean_Elab_Term_elabShow___closed__13; lean_inc(x_76); -x_96 = l_Lean_Syntax_node1(x_76, x_95, x_94); +x_96 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_96, 0, x_76); +lean_ctor_set(x_96, 1, x_95); x_97 = l_Lean_Elab_Term_expandHave___closed__1; -x_98 = l_Lean_Syntax_node4(x_76, x_97, x_80, x_96, x_90, x_71); +x_98 = l_Lean_Syntax_node4(x_76, x_97, x_80, x_94, x_96, x_71); x_99 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_99, 0, x_98); -lean_ctor_set(x_99, 1, x_7); +lean_ctor_set(x_99, 1, x_3); return x_99; } -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; -x_100 = lean_ctor_get(x_5, 0); -lean_inc(x_100); -lean_dec(x_5); -x_101 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__4; -lean_inc(x_76); -x_102 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_102, 0, x_76); -lean_ctor_set(x_102, 1, x_85); -lean_ctor_set(x_102, 2, x_101); -x_103 = l_Array_mkArray2___rarg(x_100, x_102); -x_104 = l_Array_append___rarg(x_101, x_103); -lean_inc(x_76); -x_105 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_105, 0, x_76); -lean_ctor_set(x_105, 1, x_85); -lean_ctor_set(x_105, 2, x_104); -x_106 = l_Lean_Elab_Term_expandHave___closed__5; -lean_inc(x_76); -x_107 = l_Lean_Syntax_node4(x_76, x_106, x_105, x_86, x_88, x_73); -x_108 = l_Lean_Elab_Term_expandHave___closed__3; -lean_inc(x_76); -x_109 = l_Lean_Syntax_node1(x_76, x_108, x_107); -x_110 = l_Lean_Elab_Term_expandHave___closed__1; -x_111 = l_Lean_Syntax_node4(x_76, x_110, x_80, x_109, x_90, x_71); -x_112 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_112, 0, x_111); -lean_ctor_set(x_112, 1, x_7); -return x_112; -} -} -} -} -static lean_object* _init_l_Lean_Elab_Term_expandSuffices___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("suffices", 8); -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_Term_expandSuffices___closed__2() { -_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_elabCoe___closed__3; -x_2 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__8; -x_3 = l___regBuiltin_Lean_Elab_Term_elabCoe___closed__5; -x_4 = l_Lean_Elab_Term_expandSuffices___closed__1; -x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); -return x_5; -} -} -static lean_object* _init_l_Lean_Elab_Term_expandSuffices___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("sufficesDecl", 12); -return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_expandSuffices___closed__4() { -_start: +else { -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_elabCoe___closed__3; -x_2 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__8; -x_3 = l___regBuiltin_Lean_Elab_Term_elabCoe___closed__5; -x_4 = l_Lean_Elab_Term_expandSuffices___closed__3; -x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); -return x_5; -} -} -LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandSuffices(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +lean_object* x_100; lean_object* x_101; uint8_t x_102; +x_100 = l_Lean_Syntax_getArg(x_16, x_8); +lean_dec(x_16); +x_101 = l_Lean_Elab_Term_expandHave___closed__13; +lean_inc(x_100); +x_102 = l_Lean_Syntax_isOfKind(x_100, x_101); +if (x_102 == 0) { -lean_object* x_4; uint8_t x_5; -x_4 = l_Lean_Elab_Term_expandSuffices___closed__2; -lean_inc(x_1); -x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); -if (x_5 == 0) +lean_object* x_103; uint8_t x_104; +x_103 = l_Lean_Elab_Term_expandHave___closed__15; +lean_inc(x_100); +x_104 = l_Lean_Syntax_isOfKind(x_100, x_103); +if (x_104 == 0) { -lean_object* x_6; lean_object* x_7; +lean_object* x_105; lean_object* x_106; +lean_dec(x_100); +lean_dec(x_11); +lean_dec(x_9); 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; +x_105 = lean_box(1); +x_106 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_106, 0, x_105); +lean_ctor_set(x_106, 1, x_3); +return x_106; } 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(1u); -x_11 = l_Lean_Syntax_getArg(x_1, x_10); -x_12 = l_Lean_Elab_Term_expandSuffices___closed__4; -lean_inc(x_11); -x_13 = l_Lean_Syntax_isOfKind(x_11, x_12); -if (x_13 == 0) -{ -lean_object* x_14; lean_object* x_15; +lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; uint8_t x_112; +x_107 = l_Lean_Syntax_getArg(x_100, x_8); +lean_dec(x_100); +x_108 = l_Lean_Syntax_getArg(x_11, x_10); +x_109 = lean_unsigned_to_nat(2u); +x_110 = l_Lean_Syntax_getArg(x_11, x_109); lean_dec(x_11); +x_111 = l_Lean_Elab_Term_expandShow___closed__6; +lean_inc(x_110); +x_112 = l_Lean_Syntax_isOfKind(x_110, x_111); +if (x_112 == 0) +{ +lean_object* x_113; uint8_t x_114; +x_113 = l_Lean_Elab_Term_expandShow___closed__4; +lean_inc(x_110); +x_114 = l_Lean_Syntax_isOfKind(x_110, x_113); +if (x_114 == 0) +{ +lean_object* x_115; lean_object* x_116; +lean_dec(x_110); +lean_dec(x_108); +lean_dec(x_107); 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_115 = lean_box(1); +x_116 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_116, 0, x_115); +lean_ctor_set(x_116, 1, x_3); +return x_116; } else { -lean_object* x_16; uint8_t x_17; -x_16 = l_Lean_Syntax_getArg(x_11, x_8); -x_17 = l_Lean_Syntax_isNone(x_16); -if (x_17 == 0) -{ -lean_object* x_18; uint8_t x_19; -x_18 = lean_unsigned_to_nat(2u); -lean_inc(x_16); -x_19 = l_Lean_Syntax_matchesNull(x_16, x_18); -if (x_19 == 0) +lean_object* x_117; lean_object* x_118; lean_object* x_119; uint8_t x_120; +x_117 = l_Lean_Syntax_getArg(x_110, x_8); +x_118 = l_Lean_Syntax_getArg(x_110, x_10); +lean_dec(x_110); +x_119 = l_Lean_Elab_Term_expandSuffices___closed__9; +lean_inc(x_118); +x_120 = l_Lean_Syntax_isOfKind(x_118, x_119); +if (x_120 == 0) { -lean_object* x_20; lean_object* x_21; -lean_dec(x_16); -lean_dec(x_11); +lean_object* x_121; lean_object* x_122; +lean_dec(x_118); +lean_dec(x_117); +lean_dec(x_108); +lean_dec(x_107); 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_121 = lean_box(1); +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; } else { -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_Term_expandSuffices___lambda__1(x_11, x_1, x_9, x_24, x_23, x_2, x_3); +lean_object* x_123; lean_object* x_124; lean_object* x_125; uint8_t x_126; lean_object* x_127; uint8_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; 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_123 = lean_unsigned_to_nat(3u); +x_124 = l_Lean_Syntax_getArg(x_1, x_123); lean_dec(x_1); -lean_dec(x_11); -return x_25; +x_125 = lean_ctor_get(x_2, 5); +lean_inc(x_125); +lean_dec(x_2); +x_126 = 0; +x_127 = l_Lean_SourceInfo_fromRef(x_125, x_126); +x_128 = 1; +x_129 = l_Lean_SourceInfo_fromRef(x_9, x_128); +x_130 = l_Lean_Elab_Term_expandHave___lambda__4___closed__1; +x_131 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_131, 0, x_129); +lean_ctor_set(x_131, 1, x_130); +x_132 = l_Lean_SourceInfo_fromRef(x_107, x_128); +x_133 = l_Lean_Elab_Term_expandHave___lambda__2___closed__3; +x_134 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_134, 0, x_132); +lean_ctor_set(x_134, 1, x_133); +lean_inc(x_127); +x_135 = l_Lean_Syntax_node1(x_127, x_103, x_134); +x_136 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; +x_137 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__4; +lean_inc(x_127); +x_138 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_138, 0, x_127); +lean_ctor_set(x_138, 1, x_136); +lean_ctor_set(x_138, 2, x_137); +x_139 = l_Lean_Elab_Term_elabShow___closed__11; +lean_inc(x_127); +x_140 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_140, 0, x_127); +lean_ctor_set(x_140, 1, x_139); +x_141 = l_Lean_Elab_Term_elabShow___closed__10; +lean_inc(x_127); +x_142 = l_Lean_Syntax_node2(x_127, x_141, x_140, x_108); +lean_inc(x_127); +x_143 = l_Lean_Syntax_node1(x_127, x_136, x_142); +x_144 = l_Lean_Elab_Term_elabShow___closed__12; +lean_inc(x_127); +x_145 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_145, 0, x_127); +lean_ctor_set(x_145, 1, x_144); +x_146 = l_Lean_Elab_Term_expandHave___closed__5; +lean_inc(x_127); +x_147 = l_Lean_Syntax_node5(x_127, x_146, x_135, x_138, x_143, x_145, x_124); +x_148 = l_Lean_Elab_Term_expandHave___closed__3; +lean_inc(x_127); +x_149 = l_Lean_Syntax_node1(x_127, x_148, x_147); +x_150 = l_Lean_Elab_Term_elabShow___closed__13; +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_Lean_SourceInfo_fromRef(x_117, x_128); +x_153 = l_Lean_Elab_Term_expandShow___closed__10; +x_154 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_154, 0, x_152); +lean_ctor_set(x_154, 1, x_153); +x_155 = l_Lean_Elab_Term_expandShow___closed__9; +lean_inc(x_127); +x_156 = l_Lean_Syntax_node2(x_127, x_155, x_154, x_118); +x_157 = l_Lean_Elab_Term_expandHave___closed__1; +x_158 = l_Lean_Syntax_node4(x_127, x_157, x_131, x_149, x_151, 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_3); +return x_159; +} } } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; -lean_dec(x_16); -x_26 = lean_box(0); -x_27 = lean_box(0); -x_28 = l_Lean_Elab_Term_expandSuffices___lambda__1(x_11, x_1, x_9, x_27, x_26, x_2, x_3); +lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; uint8_t x_164; lean_object* x_165; uint8_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_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_160 = l_Lean_Syntax_getArg(x_110, x_10); +lean_dec(x_110); +x_161 = lean_unsigned_to_nat(3u); +x_162 = l_Lean_Syntax_getArg(x_1, x_161); lean_dec(x_1); -lean_dec(x_11); -return x_28; -} +x_163 = lean_ctor_get(x_2, 5); +lean_inc(x_163); +lean_dec(x_2); +x_164 = 0; +x_165 = l_Lean_SourceInfo_fromRef(x_163, x_164); +x_166 = 1; +x_167 = l_Lean_SourceInfo_fromRef(x_9, x_166); +x_168 = l_Lean_Elab_Term_expandHave___lambda__4___closed__1; +x_169 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_169, 0, x_167); +lean_ctor_set(x_169, 1, x_168); +x_170 = l_Lean_SourceInfo_fromRef(x_107, x_166); +x_171 = l_Lean_Elab_Term_expandHave___lambda__2___closed__3; +x_172 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_172, 0, x_170); +lean_ctor_set(x_172, 1, x_171); +lean_inc(x_165); +x_173 = l_Lean_Syntax_node1(x_165, x_103, x_172); +x_174 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; +x_175 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__4; +lean_inc(x_165); +x_176 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_176, 0, x_165); +lean_ctor_set(x_176, 1, x_174); +lean_ctor_set(x_176, 2, x_175); +x_177 = l_Lean_Elab_Term_elabShow___closed__11; +lean_inc(x_165); +x_178 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_178, 0, x_165); +lean_ctor_set(x_178, 1, x_177); +x_179 = l_Lean_Elab_Term_elabShow___closed__10; +lean_inc(x_165); +x_180 = l_Lean_Syntax_node2(x_165, x_179, x_178, x_108); +lean_inc(x_165); +x_181 = l_Lean_Syntax_node1(x_165, x_174, x_180); +x_182 = l_Lean_Elab_Term_elabShow___closed__12; +lean_inc(x_165); +x_183 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_183, 0, x_165); +lean_ctor_set(x_183, 1, x_182); +x_184 = l_Lean_Elab_Term_expandHave___closed__5; +lean_inc(x_165); +x_185 = l_Lean_Syntax_node5(x_165, x_184, x_173, x_176, x_181, x_183, x_162); +x_186 = l_Lean_Elab_Term_expandHave___closed__3; +lean_inc(x_165); +x_187 = l_Lean_Syntax_node1(x_165, x_186, x_185); +x_188 = l_Lean_Elab_Term_elabShow___closed__13; +lean_inc(x_165); +x_189 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_189, 0, x_165); +lean_ctor_set(x_189, 1, x_188); +x_190 = l_Lean_Elab_Term_expandHave___closed__1; +x_191 = l_Lean_Syntax_node4(x_165, x_190, x_169, x_187, x_189, x_160); +x_192 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_192, 0, x_191); +lean_ctor_set(x_192, 1, x_3); +return x_192; } } } +else +{ +lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; uint8_t x_197; +x_193 = l_Lean_Syntax_getArg(x_11, x_10); +x_194 = lean_unsigned_to_nat(2u); +x_195 = l_Lean_Syntax_getArg(x_11, x_194); +lean_dec(x_11); +x_196 = l_Lean_Elab_Term_expandShow___closed__6; +lean_inc(x_195); +x_197 = l_Lean_Syntax_isOfKind(x_195, x_196); +if (x_197 == 0) +{ +lean_object* x_198; uint8_t x_199; +x_198 = l_Lean_Elab_Term_expandShow___closed__4; +lean_inc(x_195); +x_199 = l_Lean_Syntax_isOfKind(x_195, x_198); +if (x_199 == 0) +{ +lean_object* x_200; lean_object* x_201; +lean_dec(x_195); +lean_dec(x_193); +lean_dec(x_100); +lean_dec(x_9); +lean_dec(x_2); +lean_dec(x_1); +x_200 = lean_box(1); +x_201 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_201, 0, x_200); +lean_ctor_set(x_201, 1, x_3); +return x_201; +} +else +{ +lean_object* x_202; lean_object* x_203; lean_object* x_204; uint8_t x_205; +x_202 = l_Lean_Syntax_getArg(x_195, x_8); +x_203 = l_Lean_Syntax_getArg(x_195, x_10); +lean_dec(x_195); +x_204 = l_Lean_Elab_Term_expandSuffices___closed__9; +lean_inc(x_203); +x_205 = l_Lean_Syntax_isOfKind(x_203, x_204); +if (x_205 == 0) +{ +lean_object* x_206; lean_object* x_207; +lean_dec(x_203); +lean_dec(x_202); +lean_dec(x_193); +lean_dec(x_100); +lean_dec(x_9); +lean_dec(x_2); +lean_dec(x_1); +x_206 = lean_box(1); +x_207 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_207, 0, x_206); +lean_ctor_set(x_207, 1, x_3); +return x_207; } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandSuffices___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: +else { -lean_object* x_8; -x_8 = l_Lean_Elab_Term_expandSuffices___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_4); +lean_object* x_208; lean_object* x_209; lean_object* x_210; uint8_t x_211; lean_object* x_212; uint8_t 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_208 = lean_unsigned_to_nat(3u); +x_209 = l_Lean_Syntax_getArg(x_1, x_208); +lean_dec(x_1); +x_210 = lean_ctor_get(x_2, 5); +lean_inc(x_210); lean_dec(x_2); +x_211 = 0; +x_212 = l_Lean_SourceInfo_fromRef(x_210, x_211); +x_213 = 1; +x_214 = l_Lean_SourceInfo_fromRef(x_9, x_213); +x_215 = l_Lean_Elab_Term_expandHave___lambda__4___closed__1; +x_216 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_216, 0, x_214); +lean_ctor_set(x_216, 1, x_215); +x_217 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; +x_218 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__4; +lean_inc(x_212); +x_219 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_219, 0, x_212); +lean_ctor_set(x_219, 1, x_217); +lean_ctor_set(x_219, 2, x_218); +x_220 = l_Lean_Elab_Term_elabShow___closed__11; +lean_inc(x_212); +x_221 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_221, 0, x_212); +lean_ctor_set(x_221, 1, x_220); +x_222 = l_Lean_Elab_Term_elabShow___closed__10; +lean_inc(x_212); +x_223 = l_Lean_Syntax_node2(x_212, x_222, x_221, x_193); +lean_inc(x_212); +x_224 = l_Lean_Syntax_node1(x_212, x_217, x_223); +x_225 = l_Lean_Elab_Term_elabShow___closed__12; +lean_inc(x_212); +x_226 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_226, 0, x_212); +lean_ctor_set(x_226, 1, x_225); +x_227 = l_Lean_Elab_Term_expandHave___closed__5; +lean_inc(x_212); +x_228 = l_Lean_Syntax_node5(x_212, x_227, x_100, x_219, x_224, x_226, x_209); +x_229 = l_Lean_Elab_Term_expandHave___closed__3; +lean_inc(x_212); +x_230 = l_Lean_Syntax_node1(x_212, x_229, x_228); +x_231 = l_Lean_Elab_Term_elabShow___closed__13; +lean_inc(x_212); +x_232 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_232, 0, x_212); +lean_ctor_set(x_232, 1, x_231); +x_233 = l_Lean_SourceInfo_fromRef(x_202, x_213); +x_234 = l_Lean_Elab_Term_expandShow___closed__10; +x_235 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_235, 0, x_233); +lean_ctor_set(x_235, 1, x_234); +x_236 = l_Lean_Elab_Term_expandShow___closed__9; +lean_inc(x_212); +x_237 = l_Lean_Syntax_node2(x_212, x_236, x_235, x_203); +x_238 = l_Lean_Elab_Term_expandHave___closed__1; +x_239 = l_Lean_Syntax_node4(x_212, x_238, x_216, x_230, x_232, x_237); +x_240 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_240, 0, x_239); +lean_ctor_set(x_240, 1, x_3); +return x_240; +} +} +} +else +{ +lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; uint8_t x_245; lean_object* x_246; uint8_t 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; +x_241 = l_Lean_Syntax_getArg(x_195, x_10); +lean_dec(x_195); +x_242 = lean_unsigned_to_nat(3u); +x_243 = l_Lean_Syntax_getArg(x_1, x_242); lean_dec(x_1); -return x_8; +x_244 = lean_ctor_get(x_2, 5); +lean_inc(x_244); +lean_dec(x_2); +x_245 = 0; +x_246 = l_Lean_SourceInfo_fromRef(x_244, x_245); +x_247 = 1; +x_248 = l_Lean_SourceInfo_fromRef(x_9, x_247); +x_249 = l_Lean_Elab_Term_expandHave___lambda__4___closed__1; +x_250 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_250, 0, x_248); +lean_ctor_set(x_250, 1, x_249); +x_251 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; +x_252 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__4; +lean_inc(x_246); +x_253 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_253, 0, x_246); +lean_ctor_set(x_253, 1, x_251); +lean_ctor_set(x_253, 2, x_252); +x_254 = l_Lean_Elab_Term_elabShow___closed__11; +lean_inc(x_246); +x_255 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_255, 0, x_246); +lean_ctor_set(x_255, 1, x_254); +x_256 = l_Lean_Elab_Term_elabShow___closed__10; +lean_inc(x_246); +x_257 = l_Lean_Syntax_node2(x_246, x_256, x_255, x_193); +lean_inc(x_246); +x_258 = l_Lean_Syntax_node1(x_246, x_251, x_257); +x_259 = l_Lean_Elab_Term_elabShow___closed__12; +lean_inc(x_246); +x_260 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_260, 0, x_246); +lean_ctor_set(x_260, 1, x_259); +x_261 = l_Lean_Elab_Term_expandHave___closed__5; +lean_inc(x_246); +x_262 = l_Lean_Syntax_node5(x_246, x_261, x_100, x_253, x_258, x_260, x_243); +x_263 = l_Lean_Elab_Term_expandHave___closed__3; +lean_inc(x_246); +x_264 = l_Lean_Syntax_node1(x_246, x_263, x_262); +x_265 = l_Lean_Elab_Term_elabShow___closed__13; +lean_inc(x_246); +x_266 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_266, 0, x_246); +lean_ctor_set(x_266, 1, x_265); +x_267 = l_Lean_Elab_Term_expandHave___closed__1; +x_268 = l_Lean_Syntax_node4(x_246, x_267, x_250, x_264, x_266, x_241); +x_269 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_269, 0, x_268); +lean_ctor_set(x_269, 1, x_3); +return x_269; +} +} +} +} +} } } static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandSuffices___closed__1() { @@ -6175,7 +6945,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandSuffices_declRange _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(118u); x_2 = lean_unsigned_to_nat(43u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -6187,8 +6957,8 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandSuffices_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(117u); -x_2 = lean_unsigned_to_nat(88u); +x_1 = lean_unsigned_to_nat(125u); +x_2 = lean_unsigned_to_nat(95u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -6202,7 +6972,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_expandSuffices_declRange___closed__1; x_2 = lean_unsigned_to_nat(43u); x_3 = l___regBuiltin_Lean_Elab_Term_expandSuffices_declRange___closed__2; -x_4 = lean_unsigned_to_nat(88u); +x_4 = lean_unsigned_to_nat(95u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -6215,7 +6985,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandSuffices_declRange _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(118u); x_2 = lean_unsigned_to_nat(47u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -6227,7 +6997,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandSuffices_declRange _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(118u); x_2 = lean_unsigned_to_nat(61u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7760,7 +8530,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLeadingParserMacro_d _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(135u); +x_1 = lean_unsigned_to_nat(143u); x_2 = lean_unsigned_to_nat(38u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7772,7 +8542,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLeadingParserMacro_d _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(147u); x_2 = lean_unsigned_to_nat(33u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7800,7 +8570,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLeadingParserMacro_d _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(135u); +x_1 = lean_unsigned_to_nat(143u); x_2 = lean_unsigned_to_nat(42u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7812,7 +8582,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLeadingParserMacro_d _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(135u); +x_1 = lean_unsigned_to_nat(143u); x_2 = lean_unsigned_to_nat(64u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -8413,7 +9183,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabTrailingParserMacro_ _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(155u); x_2 = lean_unsigned_to_nat(39u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -8425,7 +9195,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabTrailingParserMacro_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(151u); +x_1 = lean_unsigned_to_nat(159u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -8453,7 +9223,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabTrailingParserMacro_ _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(155u); x_2 = lean_unsigned_to_nat(43u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -8465,7 +9235,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabTrailingParserMacro_ _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(155u); x_2 = lean_unsigned_to_nat(66u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -8990,7 +9760,7 @@ 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(153u); +x_1 = lean_unsigned_to_nat(161u); x_2 = lean_unsigned_to_nat(44u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9002,7 +9772,7 @@ 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(162u); +x_1 = lean_unsigned_to_nat(170u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9030,7 +9800,7 @@ 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(153u); +x_1 = lean_unsigned_to_nat(161u); x_2 = lean_unsigned_to_nat(48u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9042,7 +9812,7 @@ 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(153u); +x_1 = lean_unsigned_to_nat(161u); x_2 = lean_unsigned_to_nat(57u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9233,7 +10003,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(164u); +x_1 = lean_unsigned_to_nat(172u); x_2 = lean_unsigned_to_nat(47u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9245,7 +10015,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(165u); +x_1 = lean_unsigned_to_nat(173u); x_2 = lean_unsigned_to_nat(47u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9272,7 +10042,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(164u); +x_1 = lean_unsigned_to_nat(172u); x_2 = lean_unsigned_to_nat(51u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9284,7 +10054,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(164u); +x_1 = lean_unsigned_to_nat(172u); x_2 = lean_unsigned_to_nat(68u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9641,7 +10411,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(167u); +x_1 = lean_unsigned_to_nat(175u); x_2 = lean_unsigned_to_nat(42u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9653,7 +10423,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(173u); +x_1 = lean_unsigned_to_nat(181u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9681,7 +10451,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(167u); +x_1 = lean_unsigned_to_nat(175u); x_2 = lean_unsigned_to_nat(46u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9693,7 +10463,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(167u); +x_1 = lean_unsigned_to_nat(175u); x_2 = lean_unsigned_to_nat(58u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9933,39 +10703,11 @@ static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__20() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("hole", 4); -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__21() { -_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_elabCoe___closed__3; -x_2 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__8; -x_3 = l___regBuiltin_Lean_Elab_Term_elabCoe___closed__5; -x_4 = l_Lean_Elab_Term_expandDbgTrace___closed__20; -x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); -return x_5; -} -} -static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__22() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("_", 1); -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_Term_expandDbgTrace___closed__23() { -_start: -{ -lean_object* x_1; 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__21() { _start: { lean_object* x_1; @@ -9973,17 +10715,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__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_expandDbgTrace___closed__24; +x_2 = l_Lean_Elab_Term_expandDbgTrace___closed__21; x_3 = l_Lean_Name_str___override(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__23() { _start: { lean_object* x_1; @@ -10078,12 +10820,12 @@ lean_inc(x_16); x_40 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_40, 0, x_16); lean_ctor_set(x_40, 1, x_39); -x_41 = l_Lean_Elab_Term_expandDbgTrace___closed__22; +x_41 = l_Lean_Elab_Term_expandHave___lambda__2___closed__3; lean_inc(x_16); x_42 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_42, 0, x_16); lean_ctor_set(x_42, 1, x_41); -x_43 = l_Lean_Elab_Term_expandDbgTrace___closed__21; +x_43 = l_Lean_Elab_Term_expandHave___closed__15; lean_inc(x_16); x_44 = l_Lean_Syntax_node1(x_16, x_43, x_42); lean_inc(x_16); @@ -10094,7 +10836,7 @@ x_47 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_47, 0, x_16); lean_ctor_set(x_47, 1, x_31); lean_ctor_set(x_47, 2, x_46); -x_48 = l_Lean_Elab_Term_expandDbgTrace___closed__23; +x_48 = l_Lean_Elab_Term_expandDbgTrace___closed__20; lean_inc(x_16); x_49 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_49, 0, x_16); @@ -10143,12 +10885,12 @@ 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_Lean_Elab_Term_expandDbgTrace___closed__26; +x_71 = l_Lean_Elab_Term_expandDbgTrace___closed__23; lean_inc(x_61); x_72 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_72, 0, x_61); lean_ctor_set(x_72, 1, x_71); -x_73 = l_Lean_Elab_Term_expandDbgTrace___closed__25; +x_73 = l_Lean_Elab_Term_expandDbgTrace___closed__22; lean_inc(x_61); x_74 = l_Lean_Syntax_node2(x_61, x_73, x_72, x_9); x_75 = l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__32; @@ -10164,12 +10906,12 @@ lean_inc(x_61); x_80 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_80, 0, x_61); lean_ctor_set(x_80, 1, x_79); -x_81 = l_Lean_Elab_Term_expandDbgTrace___closed__22; +x_81 = l_Lean_Elab_Term_expandHave___lambda__2___closed__3; lean_inc(x_61); 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 = l_Lean_Elab_Term_expandDbgTrace___closed__21; +x_83 = l_Lean_Elab_Term_expandHave___closed__15; lean_inc(x_61); x_84 = l_Lean_Syntax_node1(x_61, x_83, x_82); x_85 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; @@ -10181,7 +10923,7 @@ x_88 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_88, 0, x_61); lean_ctor_set(x_88, 1, x_85); lean_ctor_set(x_88, 2, x_87); -x_89 = l_Lean_Elab_Term_expandDbgTrace___closed__23; +x_89 = l_Lean_Elab_Term_expandDbgTrace___closed__20; lean_inc(x_61); x_90 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_90, 0, x_61); @@ -10248,7 +10990,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(175u); +x_1 = lean_unsigned_to_nat(183u); x_2 = lean_unsigned_to_nat(44u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -10260,7 +11002,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(178u); +x_1 = lean_unsigned_to_nat(186u); x_2 = lean_unsigned_to_nat(70u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -10288,7 +11030,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(175u); +x_1 = lean_unsigned_to_nat(183u); x_2 = lean_unsigned_to_nat(48u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -10300,7 +11042,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(175u); +x_1 = lean_unsigned_to_nat(183u); x_2 = lean_unsigned_to_nat(62u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -10472,12 +11214,12 @@ lean_ctor_set(x_23, 0, x_12); 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_expandHave___lambda__2___closed__3; lean_inc(x_12); 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_Elab_Term_expandDbgTrace___closed__21; +x_26 = l_Lean_Elab_Term_expandHave___closed__15; lean_inc(x_12); x_27 = l_Lean_Syntax_node1(x_12, x_26, x_25); x_28 = l_Lean_Elab_Term_elabSorry___closed__7; @@ -10572,7 +11314,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(180u); +x_1 = lean_unsigned_to_nat(188u); x_2 = lean_unsigned_to_nat(29u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -10584,7 +11326,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(182u); +x_1 = lean_unsigned_to_nat(190u); x_2 = lean_unsigned_to_nat(64u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -10612,7 +11354,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(180u); +x_1 = lean_unsigned_to_nat(188u); x_2 = lean_unsigned_to_nat(33u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -10624,7 +11366,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(180u); +x_1 = lean_unsigned_to_nat(188u); x_2 = lean_unsigned_to_nat(42u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11904,7 +12646,7 @@ x_30 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_30, 0, x_15); lean_ctor_set(x_30, 1, x_28); lean_ctor_set(x_30, 2, x_7); -x_31 = l_Lean_Elab_Term_expandDbgTrace___closed__23; +x_31 = l_Lean_Elab_Term_expandDbgTrace___closed__20; lean_inc(x_15); x_32 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_32, 0, x_15); @@ -11965,7 +12707,7 @@ x_59 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_59, 0, x_44); lean_ctor_set(x_59, 1, x_57); lean_ctor_set(x_59, 2, x_7); -x_60 = l_Lean_Elab_Term_expandDbgTrace___closed__23; +x_60 = l_Lean_Elab_Term_expandDbgTrace___closed__20; lean_inc(x_44); x_61 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_61, 0, x_44); @@ -13959,7 +14701,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(259u); +x_1 = lean_unsigned_to_nat(267u); x_2 = lean_unsigned_to_nat(40u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13971,7 +14713,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(261u); +x_1 = lean_unsigned_to_nat(269u); x_2 = lean_unsigned_to_nat(37u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13999,7 +14741,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(259u); +x_1 = lean_unsigned_to_nat(267u); x_2 = lean_unsigned_to_nat(44u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -14011,7 +14753,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(259u); +x_1 = lean_unsigned_to_nat(267u); x_2 = lean_unsigned_to_nat(55u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -14347,7 +15089,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandTuple_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(263u); +x_1 = lean_unsigned_to_nat(271u); x_2 = lean_unsigned_to_nat(40u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -14359,7 +15101,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandTuple_declRange___ _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(276u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -14387,7 +15129,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandTuple_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(263u); +x_1 = lean_unsigned_to_nat(271u); x_2 = lean_unsigned_to_nat(44u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -14399,7 +15141,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandTuple_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(263u); +x_1 = lean_unsigned_to_nat(271u); x_2 = lean_unsigned_to_nat(55u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -14719,7 +15461,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandTypeAscription_dec _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(270u); +x_1 = lean_unsigned_to_nat(278u); x_2 = lean_unsigned_to_nat(49u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -14731,7 +15473,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandTypeAscription_dec _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(283u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -14759,7 +15501,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandTypeAscription_dec _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(270u); +x_1 = lean_unsigned_to_nat(278u); x_2 = lean_unsigned_to_nat(53u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -14771,7 +15513,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandTypeAscription_dec _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(270u); +x_1 = lean_unsigned_to_nat(278u); x_2 = lean_unsigned_to_nat(73u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -15083,7 +15825,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabTypeAscription_declR _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(285u); x_2 = lean_unsigned_to_nat(36u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -15095,7 +15837,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabTypeAscription_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(285u); +x_1 = lean_unsigned_to_nat(293u); x_2 = lean_unsigned_to_nat(34u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -15123,7 +15865,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabTypeAscription_declR _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(285u); x_2 = lean_unsigned_to_nat(40u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -15135,7 +15877,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabTypeAscription_declR _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(285u); x_2 = lean_unsigned_to_nat(58u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -17529,7 +18271,7 @@ static lean_object* _init_l_Lean_Elab_Term_elabSubst___lambda__4___closed__2() { 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_elabCoe___closed__3; x_2 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__8; -x_3 = l_Lean_Elab_Term_expandSuffices___lambda__1___closed__1; +x_3 = l_Lean_Elab_Term_expandSuffices___closed__7; x_4 = l_Lean_Elab_Term_elabSubst___lambda__4___closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; @@ -17549,7 +18291,7 @@ static lean_object* _init_l_Lean_Elab_Term_elabSubst___lambda__4___closed__4() { 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_elabCoe___closed__3; x_2 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__8; -x_3 = l_Lean_Elab_Term_expandSuffices___lambda__1___closed__1; +x_3 = l_Lean_Elab_Term_expandSuffices___closed__7; x_4 = l_Lean_Elab_Term_elabSubst___lambda__4___closed__3; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; @@ -17569,7 +18311,7 @@ static lean_object* _init_l_Lean_Elab_Term_elabSubst___lambda__4___closed__6() { 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_elabCoe___closed__3; x_2 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__8; -x_3 = l_Lean_Elab_Term_expandSuffices___lambda__1___closed__1; +x_3 = l_Lean_Elab_Term_expandSuffices___closed__7; x_4 = l_Lean_Elab_Term_elabSubst___lambda__4___closed__5; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; @@ -17621,7 +18363,7 @@ x_31 = l_Lean_Syntax_node3(x_14, x_21, x_24, x_26, x_30); x_32 = l_Lean_Elab_Term_elabSubst___lambda__4___closed__2; lean_inc(x_14); x_33 = l_Lean_Syntax_node1(x_14, x_32, x_31); -x_34 = l_Lean_Elab_Term_expandSuffices___lambda__1___closed__3; +x_34 = l_Lean_Elab_Term_expandSuffices___closed__9; lean_inc(x_14); x_35 = l_Lean_Syntax_node1(x_14, x_34, x_33); x_36 = l_Lean_Elab_Term_expandShow___closed__9; @@ -17708,7 +18450,7 @@ x_38 = l_Lean_Syntax_node3(x_21, x_28, x_31, x_33, x_37); x_39 = l_Lean_Elab_Term_elabSubst___lambda__4___closed__2; lean_inc(x_21); x_40 = l_Lean_Syntax_node1(x_21, x_39, x_38); -x_41 = l_Lean_Elab_Term_expandSuffices___lambda__1___closed__3; +x_41 = l_Lean_Elab_Term_expandSuffices___closed__9; lean_inc(x_21); x_42 = l_Lean_Syntax_node1(x_21, x_41, x_40); x_43 = l_Lean_Elab_Term_expandShow___closed__9; @@ -19890,7 +20632,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(306u); +x_1 = lean_unsigned_to_nat(314u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -19902,7 +20644,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(378u); +x_1 = lean_unsigned_to_nat(386u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -19930,7 +20672,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(306u); +x_1 = lean_unsigned_to_nat(314u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -19942,7 +20684,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(306u); +x_1 = lean_unsigned_to_nat(314u); x_2 = lean_unsigned_to_nat(40u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -20393,7 +21135,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(380u); +x_1 = lean_unsigned_to_nat(388u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -20405,7 +21147,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(389u); +x_1 = lean_unsigned_to_nat(397u); x_2 = lean_unsigned_to_nat(32u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -20433,7 +21175,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(380u); +x_1 = lean_unsigned_to_nat(388u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -20445,7 +21187,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(380u); +x_1 = lean_unsigned_to_nat(388u); x_2 = lean_unsigned_to_nat(48u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -20623,7 +21365,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(391u); +x_1 = lean_unsigned_to_nat(399u); x_2 = lean_unsigned_to_nat(29u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -20635,7 +21377,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(393u); +x_1 = lean_unsigned_to_nat(401u); x_2 = lean_unsigned_to_nat(40u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -20663,7 +21405,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(391u); +x_1 = lean_unsigned_to_nat(399u); x_2 = lean_unsigned_to_nat(33u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -20675,7 +21417,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(391u); +x_1 = lean_unsigned_to_nat(399u); x_2 = lean_unsigned_to_nat(44u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -21032,10 +21774,8 @@ l_Lean_Elab_Term_expandHave___lambda__2___closed__2 = _init_l_Lean_Elab_Term_exp lean_mark_persistent(l_Lean_Elab_Term_expandHave___lambda__2___closed__2); l_Lean_Elab_Term_expandHave___lambda__2___closed__3 = _init_l_Lean_Elab_Term_expandHave___lambda__2___closed__3(); lean_mark_persistent(l_Lean_Elab_Term_expandHave___lambda__2___closed__3); -l_Lean_Elab_Term_expandHave___lambda__3___closed__1 = _init_l_Lean_Elab_Term_expandHave___lambda__3___closed__1(); -lean_mark_persistent(l_Lean_Elab_Term_expandHave___lambda__3___closed__1); -l_Lean_Elab_Term_expandHave___lambda__3___closed__2 = _init_l_Lean_Elab_Term_expandHave___lambda__3___closed__2(); -lean_mark_persistent(l_Lean_Elab_Term_expandHave___lambda__3___closed__2); +l_Lean_Elab_Term_expandHave___lambda__4___closed__1 = _init_l_Lean_Elab_Term_expandHave___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_Elab_Term_expandHave___lambda__4___closed__1); l_Lean_Elab_Term_expandHave___closed__1 = _init_l_Lean_Elab_Term_expandHave___closed__1(); lean_mark_persistent(l_Lean_Elab_Term_expandHave___closed__1); l_Lean_Elab_Term_expandHave___closed__2 = _init_l_Lean_Elab_Term_expandHave___closed__2(); @@ -21054,6 +21794,18 @@ l_Lean_Elab_Term_expandHave___closed__8 = _init_l_Lean_Elab_Term_expandHave___cl lean_mark_persistent(l_Lean_Elab_Term_expandHave___closed__8); l_Lean_Elab_Term_expandHave___closed__9 = _init_l_Lean_Elab_Term_expandHave___closed__9(); lean_mark_persistent(l_Lean_Elab_Term_expandHave___closed__9); +l_Lean_Elab_Term_expandHave___closed__10 = _init_l_Lean_Elab_Term_expandHave___closed__10(); +lean_mark_persistent(l_Lean_Elab_Term_expandHave___closed__10); +l_Lean_Elab_Term_expandHave___closed__11 = _init_l_Lean_Elab_Term_expandHave___closed__11(); +lean_mark_persistent(l_Lean_Elab_Term_expandHave___closed__11); +l_Lean_Elab_Term_expandHave___closed__12 = _init_l_Lean_Elab_Term_expandHave___closed__12(); +lean_mark_persistent(l_Lean_Elab_Term_expandHave___closed__12); +l_Lean_Elab_Term_expandHave___closed__13 = _init_l_Lean_Elab_Term_expandHave___closed__13(); +lean_mark_persistent(l_Lean_Elab_Term_expandHave___closed__13); +l_Lean_Elab_Term_expandHave___closed__14 = _init_l_Lean_Elab_Term_expandHave___closed__14(); +lean_mark_persistent(l_Lean_Elab_Term_expandHave___closed__14); +l_Lean_Elab_Term_expandHave___closed__15 = _init_l_Lean_Elab_Term_expandHave___closed__15(); +lean_mark_persistent(l_Lean_Elab_Term_expandHave___closed__15); l___regBuiltin_Lean_Elab_Term_expandHave___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_expandHave___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_expandHave___closed__1); l___regBuiltin_Lean_Elab_Term_expandHave___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_expandHave___closed__2(); @@ -21080,12 +21832,6 @@ lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_expandHave_declRange___closed res = l___regBuiltin_Lean_Elab_Term_expandHave_declRange(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l_Lean_Elab_Term_expandSuffices___lambda__1___closed__1 = _init_l_Lean_Elab_Term_expandSuffices___lambda__1___closed__1(); -lean_mark_persistent(l_Lean_Elab_Term_expandSuffices___lambda__1___closed__1); -l_Lean_Elab_Term_expandSuffices___lambda__1___closed__2 = _init_l_Lean_Elab_Term_expandSuffices___lambda__1___closed__2(); -lean_mark_persistent(l_Lean_Elab_Term_expandSuffices___lambda__1___closed__2); -l_Lean_Elab_Term_expandSuffices___lambda__1___closed__3 = _init_l_Lean_Elab_Term_expandSuffices___lambda__1___closed__3(); -lean_mark_persistent(l_Lean_Elab_Term_expandSuffices___lambda__1___closed__3); l_Lean_Elab_Term_expandSuffices___closed__1 = _init_l_Lean_Elab_Term_expandSuffices___closed__1(); lean_mark_persistent(l_Lean_Elab_Term_expandSuffices___closed__1); l_Lean_Elab_Term_expandSuffices___closed__2 = _init_l_Lean_Elab_Term_expandSuffices___closed__2(); @@ -21094,6 +21840,16 @@ l_Lean_Elab_Term_expandSuffices___closed__3 = _init_l_Lean_Elab_Term_expandSuffi lean_mark_persistent(l_Lean_Elab_Term_expandSuffices___closed__3); l_Lean_Elab_Term_expandSuffices___closed__4 = _init_l_Lean_Elab_Term_expandSuffices___closed__4(); lean_mark_persistent(l_Lean_Elab_Term_expandSuffices___closed__4); +l_Lean_Elab_Term_expandSuffices___closed__5 = _init_l_Lean_Elab_Term_expandSuffices___closed__5(); +lean_mark_persistent(l_Lean_Elab_Term_expandSuffices___closed__5); +l_Lean_Elab_Term_expandSuffices___closed__6 = _init_l_Lean_Elab_Term_expandSuffices___closed__6(); +lean_mark_persistent(l_Lean_Elab_Term_expandSuffices___closed__6); +l_Lean_Elab_Term_expandSuffices___closed__7 = _init_l_Lean_Elab_Term_expandSuffices___closed__7(); +lean_mark_persistent(l_Lean_Elab_Term_expandSuffices___closed__7); +l_Lean_Elab_Term_expandSuffices___closed__8 = _init_l_Lean_Elab_Term_expandSuffices___closed__8(); +lean_mark_persistent(l_Lean_Elab_Term_expandSuffices___closed__8); +l_Lean_Elab_Term_expandSuffices___closed__9 = _init_l_Lean_Elab_Term_expandSuffices___closed__9(); +lean_mark_persistent(l_Lean_Elab_Term_expandSuffices___closed__9); l___regBuiltin_Lean_Elab_Term_expandSuffices___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_expandSuffices___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_expandSuffices___closed__1); l___regBuiltin_Lean_Elab_Term_expandSuffices___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_expandSuffices___closed__2(); @@ -21516,12 +22272,6 @@ l_Lean_Elab_Term_expandDbgTrace___closed__22 = _init_l_Lean_Elab_Term_expandDbgT lean_mark_persistent(l_Lean_Elab_Term_expandDbgTrace___closed__22); l_Lean_Elab_Term_expandDbgTrace___closed__23 = _init_l_Lean_Elab_Term_expandDbgTrace___closed__23(); lean_mark_persistent(l_Lean_Elab_Term_expandDbgTrace___closed__23); -l_Lean_Elab_Term_expandDbgTrace___closed__24 = _init_l_Lean_Elab_Term_expandDbgTrace___closed__24(); -lean_mark_persistent(l_Lean_Elab_Term_expandDbgTrace___closed__24); -l_Lean_Elab_Term_expandDbgTrace___closed__25 = _init_l_Lean_Elab_Term_expandDbgTrace___closed__25(); -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___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(); diff --git a/stage0/stdlib/Lean/Elab/BuiltinTerm.c b/stage0/stdlib/Lean/Elab/BuiltinTerm.c index 2042d08048e4..f64b983e2316 100644 --- a/stage0/stdlib/Lean/Elab/BuiltinTerm.c +++ b/stage0/stdlib/Lean/Elab/BuiltinTerm.c @@ -184,7 +184,6 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabWaitIfTypeContainsMVar_dec LEAN_EXPORT lean_object* l_Lean_resolveNamespaceCore___at_Lean_Elab_Term_elabOpen___spec__9___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_Term_elabDoubleQuotedName_declRange___closed__7; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabStrLit(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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*); lean_object* l_Lean_replaceRef(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabStrLit_declRange___closed__7; static lean_object* l___regBuiltin_Lean_Elab_Term_elabNoImplicitLambda___closed__2; @@ -394,6 +393,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabTypeOf(lean_object*, lean_object*, LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabEnsureTypeOf(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ScopedEnvExtension_popScope___rarg(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabPipeCompletion_declRange___closed__3; +lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabRawNatLit_declRange___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Term_elabPipeCompletion_declRange___closed__1; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Term_elabPipeCompletion___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -475,7 +475,6 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabOpen__ 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*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabWithDeclName_declRange___closed__7; static lean_object* l_Lean_resolveGlobalConst___at_Lean_Elab_Term_elabDoubleQuotedName___spec__3___closed__2; -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabNumLit_declRange(lean_object*); extern lean_object* l_Lean_Elab_Term_termElabAttribute; lean_object* l_Lean_Syntax_isStrLit_x3f(lean_object*); @@ -729,6 +728,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabByTactic___closed__5; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabOpen___spec__22(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_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabClear___spec__4___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Term_elabByTactic_declRange___closed__1; +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_OpenDecl_elabOpenDecl___at_Lean_Elab_Term_elabOpen___spec__3___lambda__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_resolveGlobalConstNoOverloadCore___at_Lean_Elab_Term_elabOpen___spec__13(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_elabScientificLit_declRange___closed__5; @@ -5163,7 +5163,7 @@ x_19 = lean_ctor_get(x_12, 0); lean_inc(x_19); lean_dec(x_12); x_20 = l_Lean_Expr_mvar___override(x_19); -x_21 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_20, x_2, x_3, x_4, x_5, x_9); +x_21 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_20, x_2, x_3, x_4, x_5, x_9); return x_21; } } @@ -21313,7 +21313,7 @@ lean_dec(x_20); x_21 = lean_ctor_get(x_7, 2); lean_dec(x_21); x_22 = l_Lean_Elab_Term_elabSetOption___closed__1; -x_23 = l_Lean_Option_get___at_Std_Format_pretty_x27___spec__1(x_15, x_22); +x_23 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_15, x_22); lean_ctor_set(x_7, 4, x_23); lean_ctor_set(x_7, 2, x_15); x_24 = 1; @@ -21343,7 +21343,7 @@ lean_inc(x_27); lean_inc(x_26); lean_dec(x_7); x_35 = l_Lean_Elab_Term_elabSetOption___closed__1; -x_36 = l_Lean_Option_get___at_Std_Format_pretty_x27___spec__1(x_15, x_35); +x_36 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_15, x_35); x_37 = lean_alloc_ctor(0, 11, 0); lean_ctor_set(x_37, 0, x_26); lean_ctor_set(x_37, 1, x_27); diff --git a/stage0/stdlib/Lean/Elab/Calc.c b/stage0/stdlib/Lean/Elab/Calc.c index 9e253a0f5e62..e646622bf50e 100644 --- a/stage0/stdlib/Lean/Elab/Calc.c +++ b/stage0/stdlib/Lean/Elab/Calc.c @@ -108,7 +108,6 @@ static lean_object* l_Lean_Elab_Term_getCalcFirstStep___closed__9; static lean_object* l_Lean_Elab_Term_getCalcSteps___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabCalc_docString(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_getCalcSteps(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Elab_Term_termElabAttribute; static lean_object* l___regBuiltin_Lean_Elab_Term_elabCalc_declRange___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_annotateFirstHoleWithType_go___spec__12___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*); @@ -168,6 +167,7 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_annotateFirs uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_getCalcRelation_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_annotateFirstHoleWithType_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*, lean_object*); static lean_object* l_Lean_Elab_Term_annotateFirstHoleWithType_go___lambda__1___closed__11; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); @@ -730,7 +730,7 @@ lean_inc(x_19); x_20 = lean_ctor_get(x_16, 1); lean_inc(x_20); lean_dec(x_16); -x_21 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_4, x_5, x_6, x_7, x_8, x_17); +x_21 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_4, x_5, x_6, x_7, x_8, x_17); x_22 = lean_ctor_get(x_21, 0); lean_inc(x_22); x_23 = lean_ctor_get(x_21, 1); @@ -1006,7 +1006,7 @@ lean_inc(x_113); x_114 = lean_ctor_get(x_112, 1); lean_inc(x_114); lean_dec(x_112); -x_115 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_113, x_5, x_6, x_7, x_8, x_114); +x_115 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_113, x_5, x_6, x_7, x_8, x_114); x_116 = lean_ctor_get(x_115, 0); lean_inc(x_116); x_117 = lean_ctor_get(x_115, 1); @@ -1734,7 +1734,7 @@ lean_inc(x_266); x_267 = lean_ctor_get(x_265, 1); lean_inc(x_267); lean_dec(x_265); -x_268 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_266, x_5, x_6, x_7, x_8, x_267); +x_268 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_266, x_5, x_6, x_7, x_8, x_267); x_269 = lean_ctor_get(x_268, 0); lean_inc(x_269); x_270 = lean_ctor_get(x_268, 1); diff --git a/stage0/stdlib/Lean/Elab/Command.c b/stage0/stdlib/Lean/Elab/Command.c index 550c194c6fed..91017bc0488b 100644 --- a/stage0/stdlib/Lean/Elab/Command.c +++ b/stage0/stdlib/Lean/Elab/Command.c @@ -13,29 +13,37 @@ #ifdef __cplusplus extern "C" { #endif +lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___spec__1(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftTermElabM___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_EXPORT lean_object* l_Lean_Elab_Command_Context_cmdPos___default; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftTermElabM___spec__14(lean_object*, size_t, size_t, lean_object*); lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_getBracketedBinderIds___closed__9; -uint8_t l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__12___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__4(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadCommandElabM___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabCommand___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__16; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_foldl___at_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___spec__2___closed__3; static lean_object* l_Lean_Elab_Command_liftTermElabM___rarg___closed__11; +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_448____closed__15; LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_liftAttrM___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__11; LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__12___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addMessageContextPartial___at_Lean_Elab_Command_instAddMessageContextCommandElabM___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapM___at_Lean_Elab_Command_liftTermElabM___spec__12___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_instMonadMacroAdapterCommandElabM___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Command_expandDeclId___spec__7___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadQuotationCommandElabM___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoTree___at_Lean_Elab_Command_expandDeclId___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__17(lean_object*, lean_object*); 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_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__13___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkMessageAux___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_instInhabitedScope___closed__2; lean_object* l_Lean_MapDeclarationExtension_insert___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_lintersRef; static lean_object* l_Lean_Elab_Command_getBracketedBinderIds___closed__5; +extern lean_object* l_Lean_profiler; lean_object* l_Lean_Core_getMaxHeartbeats(lean_object*); uint8_t l_Lean_Elab_isAbortExceptionId(lean_object*); static lean_object* l_Lean_Elab_Command_instMonadLogCommandElabM___closed__1; @@ -43,20 +51,23 @@ LEAN_EXPORT lean_object* l_Lean_log___at_Lean_Elab_Command_elabCommand___spec__5 LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Command_0__Lean_Elab_Command_mkTermContext___spec__1(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at_Lean_Elab_Command_elabCommandTopLevel___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__18___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_getConstInfo___at_Lean_Elab_Command_expandDeclId___spec__9___closed__2; static lean_object* l_Lean_Elab_Command_liftTermElabM___rarg___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Command_liftTermElabM(lean_object*); size_t lean_usize_shift_right(size_t, size_t); +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_448____closed__23; static lean_object* l_Lean_Elab_Command_instMonadRecDepthCommandElabM___closed__1; static lean_object* l_Lean_Elab_Command_instMonadEnvCommandElabM___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_applyVisibility___at_Lean_Elab_Command_expandDeclId___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4380____closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_getBracketedBinderIds___spec__1___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadCommandElabM___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addDocString_x27___at_Lean_Elab_Command_expandDeclId___spec__13(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_448____closed__17; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Command_0__Lean_Elab_Command_mkTermContext___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___closed__1; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__3; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__6; static lean_object* l_Lean_Elab_Command_getBracketedBinderIds___closed__6; static lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Command_expandDeclId___spec__7___lambda__1___closed__2; lean_object* lean_private_to_user_name(lean_object*); @@ -66,31 +77,33 @@ static lean_object* l_Lean_Elab_Command_State_nextMacroScope___default___closed_ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftTermElabM___spec__10(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_PersistentArray_forIn___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_Lean_throwError___at___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_448____closed__18; static lean_object* l_Lean_Elab_Command_liftTermElabM___rarg___closed__15; static lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___lambda__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkCommandElabAttributeUnsafe(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_instMonadCommandElabM___closed__3; +lean_object* l_Lean_PersistentArray_toArray___rarg(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_expandDeclId___spec__20(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_profileitIOUnsafe___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommandTopLevel___spec__12___lambda__1___closed__7; +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_Scope_opts___default; LEAN_EXPORT lean_object* l_Lean_Elab_checkIfShadowingStructureField___at_Lean_Elab_Command_expandDeclId___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_expandDeclId___at_Lean_Elab_Command_expandDeclId___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_State_infoState___default___closed__1; static lean_object* l_Lean_Elab_Command_instMonadInfoTreeCommandElabM___closed__1; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__14; -static lean_object* l_Lean_Elab_Command_elabCommand___lambda__3___closed__1; +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_448____closed__22; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadTraceCommandElabM; static lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___closed__3; extern lean_object* l_Lean_maxRecDepthErrorMessage; static lean_object* l_Lean_Elab_Command_instInhabitedCommandElabM___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_getLevelNames___rarg(lean_object*, lean_object*); lean_object* l_Lean_ConstantInfo_levelParams(lean_object*); +static lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3___closed__4; lean_object* l_Lean_indentD(lean_object*); +double lean_float_div(double, double); LEAN_EXPORT lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Command_expandDeclId___spec__7(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_liftTermElabM___rarg___closed__7; lean_object* l_Lean_Elab_mkMessageCore(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabCommand___spec__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_getMainModule___rarg(lean_object*, lean_object*); 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*); uint64_t lean_uint64_of_nat(lean_object*); @@ -109,15 +122,20 @@ lean_object* l_Lean_MessageData_joinSep(lean_object*, lean_object*); static lean_object* l_Lean_Elab_expandDeclId___at_Lean_Elab_Command_expandDeclId___spec__1___closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__20(lean_object*, lean_object*, size_t, size_t, lean_object*); 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*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_withSaveInfoContext___at_Lean_Elab_Command_liftTermElabM___spec__1(lean_object*); +LEAN_EXPORT lean_object* l_Lean_withSeconds___at_Lean_Elab_Command_runLinters___spec__15(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__10; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadLiftTIOCommandElabM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCommandTopLevel(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at_Lean_Elab_Command_elabCommandTopLevel___spec__9___closed__1; +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__11___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_liftTermElabM___rarg___closed__12; LEAN_EXPORT lean_object* l_Lean_Elab_logException___at_Lean_Elab_Command_elabCommand___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addDocString___at_Lean_Elab_Command_expandDeclId___spec__14(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__18(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); +double l_Lean_trace_profiler_threshold_getSecs(lean_object*); +static lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3___closed__1; lean_object* l_Lean_Syntax_getId(lean_object*); LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Elab_Command_elabCommandTopLevel___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_expandDeclId___spec__17___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -128,6 +146,7 @@ lean_object* l_Lean_PersistentArray_push___rarg(lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); lean_object* l_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_State_scopes___default___closed__1; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__12___lambda__1(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_qpartition___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMap_toArray___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__16(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapMAux___at_Lean_Elab_Command_liftTermElabM___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -137,23 +156,29 @@ static size_t l_Lean_Elab_Command_expandDeclId___closed__2; static lean_object* l_Lean_Elab_Command_runLinters___closed__1; lean_object* l_Lean_instBEqLocalInstance___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_expandDeclId___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__11___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_ReaderT_instFunctorReaderT___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Elab_Command_elabCommand___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_getScopes___boxed(lean_object*); lean_object* lean_mk_array(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_State_messages___default; uint8_t lean_usize_dec_eq(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadInfoTreeCommandElabM___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapMAux___at_Lean_Elab_Command_liftCoreM___spec__2___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftTermElabM___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_Lean_Elab_applyVisibility___at_Lean_Elab_Command_expandDeclId___spec__6___closed__1; 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*, 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_mkState(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_elabCommand___lambda__4___closed__3; lean_object* l_Lean_Syntax_getArgs(lean_object*); 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_PersistentArray_mapMAux___at_Lean_Elab_Command_liftCoreM___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_liftIO___rarg(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*); lean_object* l_Lean_replaceRef(lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_448____closed__10; LEAN_EXPORT lean_object* l_Lean_Elab_Command_Context_currMacroScope___default; +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4678_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_Elab_Command_elabCommandTopLevel___spec__4___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getPos_x3f(lean_object*, uint8_t); @@ -165,19 +190,24 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_liftCoreM(lean_object*); lean_object* l_Lean_Syntax_getTailPos_x3f(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_mkConstWithLevelParams___at_Lean_Elab_Command_expandDeclId___spec__8(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_getScopes___rarg(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_elabCommandTopLevel___lambda__2___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Command_runLinters___spec__6(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Elab_Command_runLinters___spec__17(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_withInfoTreeContext___at_Lean_Elab_Command_withMacroExpansion___spec__1___rarg(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*, size_t, size_t, lean_object*); static lean_object* l_Lean_Elab_Command_instMonadQuotationCommandElabM___closed__3; +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_448____closed__9; +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_448____closed__21; static lean_object* l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__6; lean_object* l_ReaderT_instMonadReaderT___rarg(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessages___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabCommand___spec__6___lambda__3___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_elabCommand___spec__12___rarg(lean_object*); lean_object* l_Lean_Elab_InfoTree_format(lean_object*, lean_object*, lean_object*); lean_object* lean_environment_find(lean_object*, lean_object*); +uint8_t lean_float_decLt(double, double); LEAN_EXPORT lean_object* l_Lean_Elab_Command_addUnivLevel___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4641____closed__1; LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabCommand___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabCommand___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_removeLeadingSpaces(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_liftCoreM___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__8(lean_object*, lean_object*, lean_object*); @@ -186,51 +216,58 @@ lean_object* l_Lean_MessageData_hasSyntheticSorry(lean_object*); LEAN_EXPORT lean_object* l_Lean_mkHashMap___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__1(lean_object*); lean_object* l_EStateM_instMonadEStateM(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_getLevelNames(lean_object*); -static lean_object* l_Lean_Elab_Command_elabCommand___closed__5; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCommandTopLevel___spec__6___closed__3; +LEAN_EXPORT lean_object* l_Lean_Elab_Command_runLinters___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_expandDeclId___at_Lean_Elab_Command_expandDeclId___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_logAt___at_Lean_Elab_Command_elabCommand___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftTermElabM___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___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Elab_Command_runLinters___spec__9___boxed(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_Command_elabCommand___spec__11(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Elab_Command_runLinters___spec__17___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_LocalContext_empty; LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__13(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadLogCommandElabM; lean_object* lean_io_get_num_heartbeats(lean_object*); uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Command_expandDeclId___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_elabCommand___lambda__3___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Command_getCurrMacroScope___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkTermContext___closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCommandTopLevel___spec__7(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabCommand___spec__12(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCommand___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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*); uint8_t l___private_Lean_Util_Trace_0__Lean_checkTraceOption(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_734_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadLogCommandElabM___lambda__2___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_instInhabitedPersistentArrayNode(lean_object*); extern lean_object* l_Lean_maxRecDepth; lean_object* l_List_head_x21___rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_expandMacroImpl_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_applyVisibility___at_Lean_Elab_Command_expandDeclId___spec__6(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_string_dec_eq(lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__11___lambda__1___closed__2; static lean_object* l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__5; LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Expr_instBEqExpr; +LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabCommand___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_Elab_Command_elabCommandTopLevel___spec__4___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Name_isPrefixOf(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCommandTopLevel___spec__5___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_Command_initFn____x40_Lean_Elab_Command___hyg_4641____closed__3; +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_448____closed__26; lean_object* l_Lean_Option_register___at_Lean_Elab_initFn____x40_Lean_Elab_AutoBound___hyg_7____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_instBEqProd___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_qsort_sort___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__19___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_liftTermElabM___rarg___closed__6; LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkTermContext(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_677_(lean_object*); +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_448____closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Command_getRef___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_logAt___at_Lean_Elab_Command_runLinters___spec__2___closed__1; LEAN_EXPORT 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_instMonadRecDepthCommandElabM___lambda__2(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__8; LEAN_EXPORT uint8_t l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommandTopLevel___spec__12___lambda__1(lean_object*); +LEAN_EXPORT lean_object* l_Lean_withSeconds___at_Lean_Elab_Command_runLinters___spec__7(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Command_runLinters___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Command_expandDeclId___spec__2___lambda__2___closed__1; lean_object* l_List_mapTR_loop___at_Lean_mkConstWithLevelParams___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadResolveNameCommandElabM; @@ -243,6 +280,7 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Comm LEAN_EXPORT lean_object* l_Lean_Elab_Command_liftEIO___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__11___lambda__1(lean_object*, lean_object*); static lean_object* l_List_foldl___at_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___spec__2___closed__1; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___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_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommandTopLevel___spec__12___lambda__1___closed__1; lean_object* l_instHashableArray___rarg___boxed(lean_object*, lean_object*); lean_object* l_Lean_InternalExceptionId_getName(lean_object*, lean_object*); @@ -250,91 +288,115 @@ LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabComma LEAN_EXPORT lean_object* l_Lean_Elab_withLogging___at_Lean_Elab_Command_elabCommand___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabCommand___closed__2; -static lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_elabCommand___spec__7___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_throwAlreadyDeclaredUniverseLevel___at_Lean_Elab_Command_addUnivLevel___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___spec__2(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_log___at_Lean_Elab_Command_runLinters___spec__3(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_elabCommand___lambda__3___closed__2; +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabCommand___spec__10(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadInfoTreeCommandElabM___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_expandDeclId___spec__5___closed__8; LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_Elab_Command_elabCommandTopLevel___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_ReaderT_instApplicativeReaderT___rarg(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_771____closed__5; +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabCommand___spec__14(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_addMessageContextPartial___at_Lean_Elab_Command_instAddMessageContextCommandElabM___spec__1___closed__3; size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadLogCommandElabM___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessages(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCommandTopLevel___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommandTopLevel___spec__12___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommandTopLevel___spec__11___boxed(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_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__2(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_Elab_Command_elabCommandTopLevel___spec__3___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__12___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_Elab_Command_instMonadRecDepthCommandElabM___lambda__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__17___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_liftTermElabM___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadEnvCommandElabM; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__4; LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_Elab_Command_expandDeclId___spec__9(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCommandTopLevel___spec__6___closed__1; +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__17___rarg(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_getVarDecls___boxed(lean_object*); static lean_object* l_Lean_Elab_Command_instMonadLogCommandElabM___closed__2; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runTermElabM___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_Lean_Elab_Command_instMonadMacroAdapterCommandElabM; lean_object* lean_st_ref_take(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__3(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Elab_Command_runLinters___spec__5___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkHashMap___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__1___boxed(lean_object*); lean_object* l_liftExcept___at_Lean_Elab_liftMacroM___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_commandElabAttribute; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__8; lean_object* l_Lean_Elab_expandMacroImpl_x3f(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_State_ngen___default___closed__2; -LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Elab_Command_elabCommand___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_liftIO___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_expandDeclId___spec__5___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Command_getBracketedBinderIds(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCommand___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ResolveName_resolveNamespace(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkPrivateName(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabCommand___spec__6(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Command_runLinters___spec__6___rarg(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__12___rarg___closed__1; +LEAN_EXPORT lean_object* l___auto____x40_Lean_Elab_Command___hyg_448_; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_expandDeclId___spec__10(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_elabCommand___lambda__4___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Command_State_nextMacroScope___default; +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__17___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__5___closed__1; LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__8___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__9(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___closed__4; static lean_object* l_Lean_Elab_Command_State_ngen___default___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadOptionsCommandElabM___rarg(lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_448____closed__20; lean_object* l_Lean_MessageData_ofSyntax(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabCommand___spec__6___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadInfoTreeCommandElabM; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadTraceCommandElabM___lambda__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Command_expandDeclId___spec__7___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCommand___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4678____closed__1; +static lean_object* l_Lean_Elab_Command_elabCommand___lambda__4___closed__2; +LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCommand___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_liftIO(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*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadResolveNameCommandElabM___lambda__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__20___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadTraceCommandElabM___lambda__2___boxed(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_liftTermElabM___rarg___closed__8; +static lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3___closed__3; static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabCommand___spec__1___closed__1; lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadLiftTIOCommandElabM(lean_object*); static lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___lambda__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabCommand___spec__13(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Command_runLinters___spec__14___boxed(lean_object*); static lean_object* l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__7; static lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___spec__1___closed__4; static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommandTopLevel___spec__12___lambda__1___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Command_getRef(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapM___at_Lean_Elab_Command_liftTermElabM___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_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadLiftTIOCommandElabM___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommandTopLevel___spec__12___lambda__1___closed__4; static lean_object* l_Lean_Elab_Command_instMonadQuotationCommandElabM___closed__1; +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_5179_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadResolveNameCommandElabM___lambda__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_expandDeclId___spec__5___closed__3; static lean_object* l_Lean_getConstInfo___at_Lean_Elab_Command_expandDeclId___spec__9___closed__1; +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Command_runLinters___spec__18(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftCoreM___spec__3(lean_object*, size_t, size_t, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__14___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadMacroAdapterCommandElabM___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Command_expandDeclId___spec__2___lambda__2___closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Command_runLinters___spec__14___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadRefCommandElabM___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___lambda__1___closed__2; lean_object* l_Lean_PersistentArray_append___rarg(lean_object*, lean_object*); @@ -345,28 +407,31 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_Context_macroStack___default; lean_object* l_Lean_Elab_getBetterRef(lean_object*, lean_object*); extern lean_object* l___private_Lean_DocString_0__Lean_docStringExt; LEAN_EXPORT lean_object* l_Lean_Elab_Command_runLinters(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__11___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftTermElabM___spec__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4641_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapM___at_Lean_Elab_Command_liftCoreM___spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4641____closed__2; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__12___lambda__3(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_getResetInfoTrees___at___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___spec__3___rarg(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__12; -LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCommand___lambda__4___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_expandDeclId___spec__19___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_expandDeclId___closed__1; -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3670_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__18___rarg(lean_object*); static lean_object* l_Lean_Elab_Command_liftTermElabM___rarg___closed__10; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadEnvCommandElabM___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__12; LEAN_EXPORT lean_object* l_Lean_Elab_Command_liftTermElabM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_replacePrefix(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_getBracketedBinderIds___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Command_withScope(lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabCommand___spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_expandDeclId___spec__16(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_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*); +static lean_object* l_Lean_Elab_Command_runLinters___closed__2; lean_object* lean_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCommandTopLevel___spec__6___boxed(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_Command___hyg_4417____closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_Command_State_scopes___default; static lean_object* l_Lean_Elab_Command_instMonadEnvCommandElabM___closed__2; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__14___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_elabCommand___lambda__5___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getOptional_x3f(lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_State_ngen___default___closed__3; @@ -374,37 +439,42 @@ static lean_object* l_Lean_Elab_Command_instMonadCommandElabM___closed__4; lean_object* l_liftExcept___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__8___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMapImp_expand___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__6(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkMetaContext___closed__2; +static lean_object* l_Lean_Elab_Command_runLinters___lambda__1___closed__2; lean_object* lean_array_to_list(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Elab_Command_elabCommand___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_MessageData_hasTag(lean_object*, lean_object*); lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__12(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_append(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_instMonadTraceCommandElabM___closed__1; static lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Command_expandDeclId___spec__7___closed__2; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__11; LEAN_EXPORT lean_object* l_Lean_Elab_Command_State_ngen___default; LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_Elab_Command_elabCommandTopLevel___spec__3___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Expr_instHashableExpr; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, double, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_getLevelNames___rarg___boxed(lean_object*, lean_object*); extern lean_object* l_Lean_levelZero; static lean_object* l_Lean_Elab_Command_elabCommand___closed__4; LEAN_EXPORT lean_object* l_Lean_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*); -LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCommandTopLevel___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_io_mono_nanos_now(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_modifyScope(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__1; lean_object* l_Lean_Elab_InfoTree_substitute(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*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftTermElabM___spec__6(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_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__20___closed__2; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_expandDeclId___spec__5(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_448____closed__13; static lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Command_expandDeclId___spec__7___lambda__2___closed__1; lean_object* l_Lean_RBNode_insert___at_Lean_NameMap_insert___spec__1___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_getCurrMacroScope(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_instMonadRefCommandElabM___closed__1; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__15; static lean_object* l_Lean_addMessageContextPartial___at_Lean_Elab_Command_instAddMessageContextCommandElabM___spec__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Command_getMainModule___rarg___boxed(lean_object*, lean_object*); lean_object* l_Lean_Elab_expandDeclIdCore(lean_object*); static lean_object* l_Lean_Elab_Command_getBracketedBinderIds___closed__7; 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*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4380____closed__5; LEAN_EXPORT lean_object* l_Lean_AssocList_replace___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__10(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Command_expandDeclId___spec__7___lambda__2___closed__2; static lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___lambda__1___closed__5; @@ -417,51 +487,57 @@ lean_object* l_Lean_extractMacroScopes(lean_object*); static lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___spec__1___closed__3; static lean_object* l_Lean_addMessageContextPartial___at_Lean_Elab_Command_instAddMessageContextCommandElabM___spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_withMacroExpansion___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_448____closed__5; lean_object* l_instHashableProd___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadCommandElabM; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadMacroAdapterCommandElabM___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_log___at_Lean_Elab_Command_elabCommand___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMapImp_moveEntries___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__7(lean_object*, lean_object*, lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); +static lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__4___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_getBracketedBinderIds___spec__1(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadEnvCommandElabM___lambda__1___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__4___closed__1; static lean_object* l_Lean_Elab_Command_State_messages___default___closed__2; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_expandDeclId___spec__18___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCommand___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__18___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__8; extern lean_object* l_Lean_warningAsError; extern lean_object* l_Lean_Elab_pp_macroStack; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommandTopLevel___spec__12(lean_object*, size_t, size_t, lean_object*); +uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabCommand___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_expandDeclId___spec__5___closed__6; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_5179____closed__1; static lean_object* l_Lean_Elab_Command_Scope_varDecls___default___closed__1; static lean_object* l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__1; LEAN_EXPORT lean_object* l_Lean_addMessageContextPartial___at_Lean_Elab_Command_instAddMessageContextCommandElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3670____closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_Scope_levelNames___default; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadQuotationCommandElabM; lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_instMonadQuotationCommandElabM___closed__4; static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommandTopLevel___spec__12___lambda__1___closed__3; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__12___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, double, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_mkState___closed__1; uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__4(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_getMainModule___boxed(lean_object*); static lean_object* l_Lean_Elab_Command_modifyScope___closed__3; static lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Command_expandDeclId___spec__2___closed__3; static lean_object* l_Lean_Elab_Command_modifyScope___closed__4; static lean_object* l_Lean_Elab_Command_liftTermElabM___rarg___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Command_Scope_currNamespace___default; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4380____closed__1; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4678____closed__2; +static lean_object* l_Lean_Elab_Command_elabCommandTopLevel___closed__1; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__20___closed__1; +static lean_object* l_Lean_Elab_Command_runLinters___lambda__1___closed__1; static lean_object* l_Lean_Elab_Command_instMonadResolveNameCommandElabM___closed__2; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCommandTopLevel___spec__6___closed__2; LEAN_EXPORT lean_object* l_Lean_HashMapImp_find_x3f___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_contains___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__5___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_getScope___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__11(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_forM___at_Lean_Elab_Command_elabCommand___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_ioErrorToMessage(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Command_expandDeclId___spec__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadLogCommandElabM___lambda__2(lean_object*, lean_object*, lean_object*); @@ -469,118 +545,132 @@ LEAN_EXPORT lean_object* l_Lean_addDocString_x27___at_Lean_Elab_Command_expandDe LEAN_EXPORT lean_object* l_Lean_Elab_throwAlreadyDeclaredUniverseLevel___at_Lean_Elab_Command_expandDeclId___spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMap_insert___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__4(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MacroScopesView_review(lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__17; LEAN_EXPORT lean_object* l_Lean_Elab_expandDeclId___at_Lean_Elab_Command_expandDeclId___spec__1___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_State_messages___default___closed__3; lean_object* l_Lean_Elab_mkElabAttribute___rarg(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_getResetTraces___at_Lean_Elab_Command_runLinters___spec__14___rarg(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_instMonadLogCommandElabM___closed__3; -LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabCommand___spec__17___boxed(lean_object*, lean_object*, lean_object*, lean_object*); 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_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabCommand___spec__6___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadLogCommandElabM___lambda__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_instMonadMacroAdapterCommandElabM___closed__3; +LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapMAux___at_Lean_Elab_Command_liftTermElabM___spec__13___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_addLinter(lean_object*, lean_object*); lean_object* lean_usize_to_nat(size_t); size_t lean_data_hashmap_mk_idx(lean_object*, uint64_t); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instInhabitedCommandElabM___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__7; LEAN_EXPORT lean_object* l_Lean_Elab_Command_runTermElabM___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_Lean_isTracingEnabledFor___at_Lean_Elab_Command_elabCommand___spec__7(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Command_runLinters___spec__16(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_448____closed__16; lean_object* l_Array_append___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_expandDeclId___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Environment_contains(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__13; extern lean_object* l_Lean_inheritedTraceOptions; +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_448____closed__3; lean_object* l_Lean_MessageData_ofExpr(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftTermElabM___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___auto____x40_Lean_Elab_Command___hyg_448____closed__19; LEAN_EXPORT lean_object* l_Lean_Elab_getResetInfoTrees___at___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___spec__3___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instInhabitedCommandElabM___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftTermElabM___spec__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3931____closed__2; +LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabCommand___spec__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_runTermElabM(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadOptionsCommandElabM___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_runLinters___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_instMonadCommandElabM___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Command_withMacroExpansion___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_448____closed__24; static lean_object* l_Lean_addMessageContextPartial___at_Lean_Elab_Command_instAddMessageContextCommandElabM___spec__1___closed__4; static lean_object* l_Lean_Elab_Command_elabCommand___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_Scope_varDecls___default; static lean_object* l_Lean_Elab_Command_instMonadCommandElabM___closed__6; static lean_object* l_Lean_Elab_Command_modifyScope___closed__2; -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__19(lean_object*, size_t, size_t, 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* 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_PersistentArray_mapMAux___at_Lean_Elab_Command_liftTermElabM___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_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_PersistentArray_anyM___at_Lean_MessageLog_hasErrors___spec__1(lean_object*); +double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftCoreM___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__14; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftTermElabM___spec__9(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_qsort_sort___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__19___lambda__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabCommand___spec__17(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__5(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabCommand___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_elabCommand___lambda__4___closed__5; static lean_object* l_Lean_addMessageContextPartial___at_Lean_Elab_Command_instAddMessageContextCommandElabM___spec__1___closed__5; +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabCommand___spec__15(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_expandDeclId___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_withoutModifyingElabMetaStateWithInfo___spec__2___rarg(lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_448____closed__27; LEAN_EXPORT lean_object* l_Lean_Elab_Command_withFreshMacroScope___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Elab_Command_elabCommand___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabCommand___spec__6___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); 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_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__13___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_logException___at_Lean_Elab_Command_runLinters___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadTraceCommandElabM___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__9; static lean_object* l_Lean_Elab_Command_instMonadEnvCommandElabM___closed__1; static lean_object* l_Lean_Elab_Command_State_messages___default___closed__1; -LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Elab_Command_runLinters___spec__5(lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftCoreM___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoTree___at_Lean_Elab_Command_expandDeclId___spec__12(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabCommand___closed__3; static lean_object* l_Lean_Elab_Command_instMonadInfoTreeCommandElabM___closed__3; -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabCommand___spec__16(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3___closed__2; +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__13(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_getScopes(lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__17(lean_object*, lean_object*); +static double l_Lean_withSeconds___at_Lean_Elab_Command_runLinters___spec__7___closed__1; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftCoreM___spec__5(lean_object*, size_t, size_t, lean_object*); +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_448____closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadOptionsCommandElabM___rarg___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCommand___lambda__4(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__9; +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Command_runLinters___spec__6___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCommand___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Command_runLinters___spec__14(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_Scope_openDecls___default; 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*); uint8_t l_Lean_Name_isAtomic(lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftCoreM___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); 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_Lean_Elab_mkDeclName___at_Lean_Elab_Command_expandDeclId___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_logException___at_Lean_Elab_Command_runLinters___spec__1___closed__3; static lean_object* l_Lean_Elab_Command_instAddMessageContextCommandElabM___closed__1; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_instMonadLogCommandElabM___closed__5; -LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabCommand___spec__14(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_throwAlreadyDeclaredUniverseLevel___at_Lean_Elab_Command_addUnivLevel___spec__1___closed__4; +LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Elab_Command_runLinters___spec__19(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_getScope(lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4918____closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadLogCommandElabM___lambda__3(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4380____closed__4; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__13; lean_object* lean_environment_main_module(lean_object*); static lean_object* l_Lean_Elab_Command_liftTermElabM___rarg___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_withInfoTreeContext___at_Lean_Elab_Command_withMacroExpansion___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommandTopLevel___spec__12___lambda__1___boxed(lean_object*); +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_448____closed__7; lean_object* l_Lean_Elab_Term_TermElabM_run___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadResolveNameCommandElabM___lambda__2___boxed(lean_object*, lean_object*, lean_object*); +lean_object* lean_float_to_string(double); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftTermElabM___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_profileitM___at_Lean_Elab_Command_runLinters___spec__5___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__10; LEAN_EXPORT lean_object* l_Lean_PersistentArray_foldlM___at_Lean_Elab_Command_elabCommandTopLevel___spec__8(lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__11___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkMessageAux(lean_object*, lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Elab_Command_getLevelNames___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabCommand___spec__11(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_adaptExpander(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_addLinter___closed__1; lean_object* l_Lean_Name_mkStr2(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*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_5179____closed__2; +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_448____closed__11; LEAN_EXPORT lean_object* l_Lean_Elab_Command_withMacroExpansion___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___spec__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Command_withFreshMacroScope(lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__4; -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__13___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_getMainModule(lean_object*); +LEAN_EXPORT lean_object* l_List_forM___at_Lean_Elab_Command_elabCommand___spec__8(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Command_expandDeclId___spec__7___closed__1; -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_elabCommand___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_addDocString___at_Lean_Elab_Command_expandDeclId___spec__14___closed__1; static lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___closed__2; -LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabCommand___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_getBracketedBinderIds___closed__8; LEAN_EXPORT lean_object* l_Lean_Elab_getResetInfoTrees___at___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___spec__3(lean_object*); static lean_object* l_Lean_Elab_Command_instMonadRefCommandElabM___closed__3; @@ -588,22 +678,20 @@ LEAN_EXPORT lean_object* l_Array_qsort_sort___at___private_Lean_Elab_Command_0__ static lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___lambda__1___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadTraceCommandElabM___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_TagDeclarationExtension_tag(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__19___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCommandTopLevel___lambda__1(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_runLinters___spec__11___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_Command_expandDeclId___spec__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadLogCommandElabM___lambda__3___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3670____closed__2; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__10; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__2; +static lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3___closed__5; LEAN_EXPORT lean_object* l_Lean_addDocString___at_Lean_Elab_Command_expandDeclId___spec__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabCommand___spec__15(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_448____closed__8; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftTermElabM___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*); static lean_object* l_Lean_Elab_Command_instMonadInfoTreeCommandElabM___closed__2; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__14(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Command_expandDeclId___spec__7___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_getScopes___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Elab_Command_runLinters___spec__19___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_instMonadCommandElabM___closed__1; -static lean_object* l_Lean_Elab_Command_elabCommand___closed__6; uint8_t l_Lean_Syntax_isNone(lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_Elab_Command_modifyScope___spec__1(lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); @@ -618,37 +706,44 @@ static lean_object* l_Lean_Elab_Command_instMonadResolveNameCommandElabM___close lean_object* l_Lean_ResolveName_resolveGlobalName(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadMacroAdapterCommandElabM___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Command_runLinters___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__18(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_getInfoTrees___at_Lean_Elab_Command_elabCommandTopLevel___spec__1___rarg(lean_object*, lean_object*); uint8_t l_Lean_isStructure(lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_448____closed__25; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_5179____closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Command_liftCoreM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlMAux___at_Lean_Elab_Command_elabCommandTopLevel___spec__10(lean_object*, lean_object*); lean_object* lean_nat_mul(lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_expandDeclId___spec__19___closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__13(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); lean_object* l_Lean_Core_resetMessageLog(lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_448____closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabCommand___spec__6___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_logException___at_Lean_Elab_Command_runLinters___spec__1___closed__2; 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_object* l_Lean_Meta_InfoCacheKey_instHashableInfoCacheKey___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftTermElabM___spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwAlreadyDeclaredUniverseLevel___at_Lean_Elab_Command_expandDeclId___spec__15(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__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_Elab_Command_instInhabitedCommandElabM(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_getInfoTrees___at_Lean_Elab_Command_elabCommandTopLevel___spec__1(lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__12___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_Lean_Elab_Command_liftTermElabM___rarg___closed__14; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_expandDeclId___spec__19___closed__2; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCommandTopLevel___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___spec__1___closed__5; lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Command_runLinters___spec__6___boxed(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___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Elab_Command_runLinters___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_protectedExt; LEAN_EXPORT lean_object* l_Lean_mkConstWithLevelParams___at_Lean_Elab_Command_expandDeclId___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_expandDeclId___spec__18(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Nat_decEq___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4380____closed__3; lean_object* l_List_reverse___rarg(lean_object*); static lean_object* l_Lean_Elab_Command_instMonadRecDepthCommandElabM___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadOptionsCommandElabM(lean_object*); -LEAN_EXPORT lean_object* l_List_forM___at_Lean_Elab_Command_elabCommand___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_foldl___at_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___spec__2___closed__2; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__20___closed__3; static lean_object* l_Lean_Elab_Command_instMonadCommandElabM___closed__7; @@ -657,31 +752,37 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadLogCommandElabM___lambda__ LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCommand___boxed__const__1; size_t lean_usize_sub(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instInhabitedScope; +LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapM___at_Lean_Elab_Command_liftCoreM___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_liftTermElabM___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkHashMapImp___rarg(lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabCommand___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_448____closed__2; +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_448____closed__14; LEAN_EXPORT lean_object* l_Lean_Elab_Command_runTermElabM___rarg___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_Lean_Elab_Command_instMonadRecDepthCommandElabM; +static lean_object* l_Lean_Elab_Command_elabCommand___lambda__4___closed__1; static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommandTopLevel___spec__12___closed__1; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_expandDeclId___spec__5___closed__7; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_expandDeclId___spec__21(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_Scope_varUIds___default; lean_object* l_Lean_Syntax_hasMissing(lean_object*); +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapM___at_Lean_Elab_Command_liftTermElabM___spec__12(lean_object*, lean_object*); 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*); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Command_runLinters___spec__8___boxed(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_EXPORT lean_object* l_Lean_Elab_Command_instMonadLogCommandElabM___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkTermContext___elambda__1(lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__7; LEAN_EXPORT uint8_t l_Lean_Elab_Command_Scope_isNoncomputable___default; LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandDeclId(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__12(lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_expandDeclId___spec__5___closed__5; -static lean_object* l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__11; LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_ioErrorToMessage___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__1; static lean_object* l_Lean_Elab_Command_liftTermElabM___rarg___closed__13; LEAN_EXPORT lean_object* l_Lean_Elab_Command_getScope___rarg___boxed(lean_object*, lean_object*); +extern lean_object* l_Lean_trace_profiler; uint8_t l_Lean_PersistentArray_isEmpty___rarg(lean_object*); -LEAN_EXPORT lean_object* l_List_forM___at_Lean_Elab_Command_elabCommand___spec__9(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_expandDeclId___spec__21___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_instMonadCommandElabM___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Command_Context_currRecDepth___default; @@ -689,27 +790,29 @@ lean_object* lean_io_error_to_string(lean_object*); lean_object* l_Lean_KeyedDeclsAttribute_getEntries___rarg(lean_object*, lean_object*, lean_object*); 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*); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__9; LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkCoreContext(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_logException___at_Lean_Elab_Command_runLinters___spec__1___closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_expandDeclId___spec__20___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_Command_0__Lean_Elab_Command_mkMetaContext; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__12; static lean_object* l_Lean_Elab_Command_getBracketedBinderIds___closed__3; -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__18(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_instMonadRefCommandElabM___closed__2; -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4380_(lean_object*); size_t lean_usize_shift_left(size_t, size_t); lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__13___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_addMacroStack___at_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_instBEq___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_instMonadRecDepthCommandElabM___closed__3; 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_Lean_Elab_getInfoTrees___at_Lean_Elab_Command_elabCommandTopLevel___spec__1___boxed(lean_object*); -static lean_object* l_Lean_Elab_Command_elabCommand___lambda__3___closed__5; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___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*); uint8_t l_Lean_Name_isSuffixOf(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_instMonadTraceCommandElabM___closed__3; +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3931_(lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftCoreM___spec__4(lean_object*, size_t, size_t, 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_Command_initFn____x40_Lean_Elab_Command___hyg_4918____closed__1; lean_object* lean_array_get_size(lean_object*); static lean_object* l_Lean_Elab_Command_liftCoreM___rarg___closed__1; LEAN_EXPORT uint8_t l_Lean_AssocList_contains___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__5(lean_object*, lean_object*); @@ -731,31 +834,37 @@ extern lean_object* l_Lean_Elab_unsupportedSyntaxExceptionId; LEAN_EXPORT lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Command_expandDeclId___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_withMacroExpansion(lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Command_runLinters___spec__18___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_instMonadTraceCommandElabM___closed__2; static lean_object* l_Lean_Elab_throwAlreadyDeclaredUniverseLevel___at_Lean_Elab_Command_addUnivLevel___spec__1___closed__2; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4641____closed__5; static lean_object* l_Array_qsort_sort___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__19___closed__1; static lean_object* l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at_Lean_Elab_Command_elabCommandTopLevel___spec__9(lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Command_runLinters___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__12___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_nat_add(lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_expandDeclId___spec__5___closed__4; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4918____closed__3; static lean_object* l_Lean_Elab_Command_modifyScope___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapMAux___at_Lean_Elab_Command_liftTermElabM___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_Command_instMonadResolveNameCommandElabM___closed__1; LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Elab_Command_elabCommandTopLevel___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_448____closed__12; LEAN_EXPORT lean_object* l_Lean_Elab_getResetInfoTrees___at___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___spec__3___boxed(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCommandTopLevel___spec__6(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_instBEqArray___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_expandDeclId___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_Command_initFn____x40_Lean_Elab_Command___hyg_4918_(lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommandTopLevel___spec__12___lambda__1___closed__6; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftTermElabM___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_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftTermElabM___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_Command_instMonadMacroAdapterCommandElabM___closed__2; +LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabCommand___spec__16(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l___private_Lean_Message_0__Lean_beqMessageSeverity____x40_Lean_Message___hyg_120_(uint8_t, uint8_t); LEAN_EXPORT lean_object* l_Lean_Elab_Command_State_infoState___default; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadInfoTreeCommandElabM___lambda__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftTermElabM___spec__16(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_log___at_Lean_Elab_Command_runLinters___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Elab_Command___hyg_448____closed__28; LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkTermContext___elambda__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_liftEIO___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___spec__1___closed__1; @@ -763,6 +872,7 @@ lean_object* l_Lean_getStructureFieldsFlattened(lean_object*, lean_object*, uint LEAN_EXPORT lean_object* l_Lean_Elab_Command_instAddMessageContextCommandElabM; LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Command_expandDeclId___spec__11(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_modifyScope___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Elab_Command_runLinters___spec__19___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwAlreadyDeclaredUniverseLevel___at_Lean_Elab_Command_addUnivLevel___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_State_infoState___default___closed__4; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommandTopLevel___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -771,7 +881,9 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_runTermElabM___rarg___lambda__1(lea LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkTermContext___boxed(lean_object*, lean_object*); 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_Array_qsort_sort___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__19___lambda__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCommand___lambda__5(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkMetaContext___closed__1; +LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapMAux___at_Lean_Elab_Command_liftTermElabM___spec__13(lean_object*, lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadResolveNameCommandElabM___lambda__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_throwAlreadyDeclaredUniverseLevel___at_Lean_Elab_Command_addUnivLevel___spec__1___closed__1; @@ -779,12 +891,14 @@ lean_object* l_Lean_MessageData_ofName(lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_expandDeclId___spec__5___closed__1; lean_object* lean_array_get(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_4641____closed__4; static lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Command_expandDeclId___spec__2___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Command_expandDeclId___spec__2___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_EXPORT lean_object* l_Lean_Elab_Command_catchExceptions(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Command_runLinters___spec__16___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_instMonadRecDepthCommandElabM___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_runTermElabM___rarg___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_elabCommand___lambda__3___closed__3; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftTermElabM___spec__15(lean_object*, size_t, size_t, lean_object*); size_t lean_usize_land(size_t, size_t); static lean_object* l_Lean_Elab_Command_State_infoState___default___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadRecDepthCommandElabM___lambda__3___boxed(lean_object*, lean_object*, lean_object*); @@ -793,6 +907,8 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_liftTermElabM___rarg___boxed(lean_o lean_object* l_ReaderT_pure___at_Lean_Elab_liftMacroM___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapM___at_Lean_Elab_Command_liftTermElabM___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Data_HashMap_0__Lean_numBucketsForCapacity(lean_object*); +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Command_runLinters___spec__10(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3931____closed__1; LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_Elab_Command_expandDeclId___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_liftTermElabM___rarg___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_getVarDecls(lean_object*); @@ -1127,6 +1243,306 @@ x_1 = lean_box(0); return x_1; } } +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_448____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Lean", 4); +return x_1; +} +} +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_448____closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Parser", 6); +return x_1; +} +} +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_448____closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Tactic", 6); +return x_1; +} +} +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_448____closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("tacticSeq", 9); +return x_1; +} +} +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_448____closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___auto____x40_Lean_Elab_Command___hyg_448____closed__1; +x_2 = l___auto____x40_Lean_Elab_Command___hyg_448____closed__2; +x_3 = l___auto____x40_Lean_Elab_Command___hyg_448____closed__3; +x_4 = l___auto____x40_Lean_Elab_Command___hyg_448____closed__4; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_448____closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("tacticSeq1Indented", 18); +return x_1; +} +} +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_448____closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___auto____x40_Lean_Elab_Command___hyg_448____closed__1; +x_2 = l___auto____x40_Lean_Elab_Command___hyg_448____closed__2; +x_3 = l___auto____x40_Lean_Elab_Command___hyg_448____closed__3; +x_4 = l___auto____x40_Lean_Elab_Command___hyg_448____closed__6; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_448____closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("null", 4); +return x_1; +} +} +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_448____closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___auto____x40_Lean_Elab_Command___hyg_448____closed__8; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_448____closed__10() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("exact", 5); +return x_1; +} +} +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_448____closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___auto____x40_Lean_Elab_Command___hyg_448____closed__1; +x_2 = l___auto____x40_Lean_Elab_Command___hyg_448____closed__2; +x_3 = l___auto____x40_Lean_Elab_Command___hyg_448____closed__3; +x_4 = l___auto____x40_Lean_Elab_Command___hyg_448____closed__10; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_448____closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(2); +x_2 = l___auto____x40_Lean_Elab_Command___hyg_448____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___auto____x40_Lean_Elab_Command___hyg_448____closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Command_Scope_varDecls___default___closed__1; +x_2 = l___auto____x40_Lean_Elab_Command___hyg_448____closed__12; +x_3 = lean_array_push(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_448____closed__14() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Term", 4); +return x_1; +} +} +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_448____closed__15() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("declName", 8); +return x_1; +} +} +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_448____closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___auto____x40_Lean_Elab_Command___hyg_448____closed__1; +x_2 = l___auto____x40_Lean_Elab_Command___hyg_448____closed__2; +x_3 = l___auto____x40_Lean_Elab_Command___hyg_448____closed__14; +x_4 = l___auto____x40_Lean_Elab_Command___hyg_448____closed__15; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_448____closed__17() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("decl_name%", 10); +return x_1; +} +} +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_448____closed__18() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(2); +x_2 = l___auto____x40_Lean_Elab_Command___hyg_448____closed__17; +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___auto____x40_Lean_Elab_Command___hyg_448____closed__19() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Command_Scope_varDecls___default___closed__1; +x_2 = l___auto____x40_Lean_Elab_Command___hyg_448____closed__18; +x_3 = lean_array_push(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_448____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___auto____x40_Lean_Elab_Command___hyg_448____closed__16; +x_3 = l___auto____x40_Lean_Elab_Command___hyg_448____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; +} +} +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_448____closed__21() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___auto____x40_Lean_Elab_Command___hyg_448____closed__13; +x_2 = l___auto____x40_Lean_Elab_Command___hyg_448____closed__20; +x_3 = lean_array_push(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_448____closed__22() { +_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___auto____x40_Lean_Elab_Command___hyg_448____closed__11; +x_3 = l___auto____x40_Lean_Elab_Command___hyg_448____closed__21; +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___auto____x40_Lean_Elab_Command___hyg_448____closed__23() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Command_Scope_varDecls___default___closed__1; +x_2 = l___auto____x40_Lean_Elab_Command___hyg_448____closed__22; +x_3 = lean_array_push(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_448____closed__24() { +_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___auto____x40_Lean_Elab_Command___hyg_448____closed__9; +x_3 = l___auto____x40_Lean_Elab_Command___hyg_448____closed__23; +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___auto____x40_Lean_Elab_Command___hyg_448____closed__25() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Command_Scope_varDecls___default___closed__1; +x_2 = l___auto____x40_Lean_Elab_Command___hyg_448____closed__24; +x_3 = lean_array_push(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_448____closed__26() { +_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___auto____x40_Lean_Elab_Command___hyg_448____closed__7; +x_3 = l___auto____x40_Lean_Elab_Command___hyg_448____closed__25; +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___auto____x40_Lean_Elab_Command___hyg_448____closed__27() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Command_Scope_varDecls___default___closed__1; +x_2 = l___auto____x40_Lean_Elab_Command___hyg_448____closed__26; +x_3 = lean_array_push(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___auto____x40_Lean_Elab_Command___hyg_448____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___auto____x40_Lean_Elab_Command___hyg_448____closed__5; +x_3 = l___auto____x40_Lean_Elab_Command___hyg_448____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___auto____x40_Lean_Elab_Command___hyg_448_() { +_start: +{ +lean_object* x_1; +x_1 = l___auto____x40_Lean_Elab_Command___hyg_448____closed__28; +return x_1; +} +} LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadCommandElabM___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { @@ -1331,7 +1747,7 @@ x_10 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_10, 0, x_9); lean_ctor_set(x_10, 1, x_4); x_11 = l_Lean_Elab_Command_mkState___closed__1; -x_12 = l_Lean_Option_get___at_Std_Format_pretty_x27___spec__1(x_3, x_11); +x_12 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_3, x_11); lean_dec(x_3); x_13 = l_Lean_Elab_Command_State_nextMacroScope___default___closed__1; x_14 = lean_unsigned_to_nat(1u); @@ -1351,7 +1767,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_677_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_734_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; uint8_t x_4; @@ -1377,53 +1793,222 @@ return x_7; } } } -static lean_object* _init_l_Lean_Elab_Command_addLinter___closed__1() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__1() { _start: { lean_object* x_1; -x_1 = l_Lean_Elab_Command_lintersRef; +x_1 = lean_mk_string_from_bytes("Elab", 4); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_addLinter(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__2() { _start: { -lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; -x_3 = l_Lean_Elab_Command_addLinter___closed__1; -x_4 = lean_st_ref_get(x_3, x_2); -x_5 = lean_ctor_get(x_4, 0); -lean_inc(x_5); -x_6 = lean_ctor_get(x_4, 1); -lean_inc(x_6); -lean_dec(x_4); -x_7 = lean_array_push(x_5, x_1); -x_8 = lean_st_ref_set(x_3, x_7, x_6); -x_9 = !lean_is_exclusive(x_8); -if (x_9 == 0) -{ -return x_8; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("lint", 4); +return x_1; } -else -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = lean_ctor_get(x_8, 0); -x_11 = lean_ctor_get(x_8, 1); -lean_inc(x_11); -lean_inc(x_10); -lean_dec(x_8); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_10); -lean_ctor_set(x_12, 1, x_11); -return x_12; } +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__3() { +_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_771____closed__1; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__2; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadInfoTreeCommandElabM___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__4() { _start: { -lean_object* x_4; uint8_t x_5; -x_4 = lean_st_ref_get(x_2, x_3); -x_5 = !lean_is_exclusive(x_4); +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___auto____x40_Lean_Elab_Command___hyg_448____closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__5() { +_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_771____closed__4; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Command", 7); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__7() { +_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_771____closed__5; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__6; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("initFn", 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__9() { +_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_771____closed__7; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__8; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__10() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("_@", 2); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__11() { +_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_771____closed__9; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__10; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__12() { +_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_771____closed__11; +x_2 = l___auto____x40_Lean_Elab_Command___hyg_448____closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__13() { +_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_771____closed__12; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__14() { +_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_771____closed__13; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__6; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__15() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("_hyg", 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__16() { +_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_771____closed__14; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__15; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__17() { +_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_771____closed__16; +x_2 = lean_unsigned_to_nat(771u); +x_3 = l_Lean_Name_num___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771_(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__3; +x_3 = 0; +x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__17; +x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); +return x_5; +} +} +static lean_object* _init_l_Lean_Elab_Command_addLinter___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Elab_Command_lintersRef; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Command_addLinter(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_3 = l_Lean_Elab_Command_addLinter___closed__1; +x_4 = lean_st_ref_get(x_3, x_2); +x_5 = lean_ctor_get(x_4, 0); +lean_inc(x_5); +x_6 = lean_ctor_get(x_4, 1); +lean_inc(x_6); +lean_dec(x_4); +x_7 = lean_array_push(x_5, x_1); +x_8 = lean_st_ref_set(x_3, x_7, x_6); +x_9 = !lean_is_exclusive(x_8); +if (x_9 == 0) +{ +return x_8; +} +else +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = lean_ctor_get(x_8, 0); +x_11 = lean_ctor_get(x_8, 1); +lean_inc(x_11); +lean_inc(x_10); +lean_dec(x_8); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_10); +lean_ctor_set(x_12, 1, x_11); +return x_12; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadInfoTreeCommandElabM___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; +x_4 = lean_st_ref_get(x_2, x_3); +x_5 = !lean_is_exclusive(x_4); if (x_5 == 0) { lean_object* x_6; lean_object* x_7; @@ -2457,7 +3042,7 @@ x_12 = lean_ctor_get(x_11, 1); lean_inc(x_12); lean_dec(x_11); x_13 = l_Lean_Elab_addMacroStack___at_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___spec__1___closed__1; -x_14 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_12, x_13); +x_14 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_12, x_13); lean_dec(x_12); if (x_14 == 0) { @@ -2517,7 +3102,7 @@ x_30 = lean_ctor_get(x_29, 1); lean_inc(x_30); lean_dec(x_29); x_31 = l_Lean_Elab_addMacroStack___at_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___spec__1___closed__1; -x_32 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_30, x_31); +x_32 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_30, x_31); lean_dec(x_30); if (x_32 == 0) { @@ -4632,133 +5217,386 @@ lean_dec(x_2); return x_4; } } -static lean_object* _init_l_Lean_Elab_Command_liftCoreM___rarg___closed__1() { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftCoreM___spec__3(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_addMessageContextPartial___at_Lean_Elab_Command_instAddMessageContextCommandElabM___spec__1___closed__2; -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; +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 = l_Lean_PersistentArray_mapMAux___at_Lean_Elab_Command_liftCoreM___spec__2(x_1, 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_Lean_Elab_Command_liftCoreM___rarg(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_Lean_Elab_Command_liftCoreM___spec__4(lean_object* x_1, size_t x_2, size_t 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; uint8_t x_16; -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_io_get_num_heartbeats(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_dec(x_8); -x_11 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkCoreContext(x_2, x_6, x_9); -x_12 = lean_ctor_get(x_6, 7); -lean_inc(x_12); -x_13 = lean_ctor_get(x_6, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_6, 6); -lean_inc(x_14); -x_15 = lean_ctor_get(x_6, 8); -lean_inc(x_15); -lean_dec(x_6); -x_16 = !lean_is_exclusive(x_12); -if (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_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; -x_17 = lean_ctor_get(x_12, 1); -lean_dec(x_17); -x_18 = lean_ctor_get(x_12, 0); -lean_dec(x_18); -x_19 = l_Lean_Elab_Command_State_infoState___default___closed__3; -x_20 = l_Lean_Elab_Command_State_messages___default___closed__3; -lean_ctor_set(x_12, 1, x_20); -lean_ctor_set(x_12, 0, x_19); -x_21 = l_Lean_Elab_Command_State_nextMacroScope___default___closed__1; -x_22 = l_Lean_Elab_Command_liftCoreM___rarg___closed__1; -x_23 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_23, 0, x_13); -lean_ctor_set(x_23, 1, x_21); -lean_ctor_set(x_23, 2, x_14); -lean_ctor_set(x_23, 3, x_15); -lean_ctor_set(x_23, 4, x_22); -lean_ctor_set(x_23, 5, x_20); -lean_ctor_set(x_23, 6, x_12); -x_102 = lean_st_mk_ref(x_23, x_10); -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); -lean_inc(x_103); -x_105 = lean_apply_3(x_1, x_11, x_103, x_104); -if (lean_obj_tag(x_105) == 0) +uint8_t x_5; +x_5 = lean_usize_dec_lt(x_3, x_2); +if (x_5 == 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; -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_alloc_ctor(1, 1, 0); -lean_ctor_set(x_108, 0, x_106); -x_109 = lean_st_ref_get(x_103, x_107); -lean_dec(x_103); -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_alloc_ctor(0, 2, 0); -lean_ctor_set(x_112, 0, x_108); -lean_ctor_set(x_112, 1, x_110); -x_24 = x_112; -x_25 = x_111; -goto block_101; +return x_4; } 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; -x_113 = lean_ctor_get(x_105, 0); -lean_inc(x_113); -x_114 = lean_ctor_get(x_105, 1); -lean_inc(x_114); -lean_dec(x_105); -x_115 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_115, 0, x_113); -x_116 = lean_st_ref_get(x_103, x_114); -lean_dec(x_103); -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_alloc_ctor(0, 2, 0); -lean_ctor_set(x_119, 0, x_115); -lean_ctor_set(x_119, 1, x_117); -x_24 = x_119; -x_25 = x_118; -goto block_101; +lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; +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_is_exclusive(x_6); +if (x_9 == 0) +{ +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_10 = lean_ctor_get(x_6, 0); +x_11 = lean_ctor_get(x_1, 6); +x_12 = l_Lean_replaceRef(x_10, x_11); +lean_dec(x_10); +lean_ctor_set(x_6, 0, x_12); +x_13 = 1; +x_14 = lean_usize_add(x_3, x_13); +x_15 = lean_array_uset(x_8, x_3, x_6); +x_3 = x_14; +x_4 = x_15; +goto _start; } -block_101: +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; uint8_t x_36; -x_26 = lean_ctor_get(x_24, 0); -lean_inc(x_26); -x_27 = lean_ctor_get(x_24, 1); -lean_inc(x_27); -lean_dec(x_24); -x_28 = lean_st_ref_take(x_3, x_25); -x_29 = lean_ctor_get(x_28, 0); -lean_inc(x_29); -x_30 = lean_ctor_get(x_28, 1); +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; +x_17 = lean_ctor_get(x_6, 0); +x_18 = lean_ctor_get(x_6, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_6); +x_19 = lean_ctor_get(x_1, 6); +x_20 = l_Lean_replaceRef(x_17, x_19); +lean_dec(x_17); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_18); +x_22 = 1; +x_23 = lean_usize_add(x_3, x_22); +x_24 = lean_array_uset(x_8, x_3, x_21); +x_3 = x_23; +x_4 = x_24; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapMAux___at_Lean_Elab_Command_liftCoreM___spec__2(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_2); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; size_t x_6; size_t x_7; lean_object* x_8; +x_4 = lean_ctor_get(x_2, 0); +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_Lean_Elab_Command_liftCoreM___spec__3(x_1, x_6, x_7, x_4); +lean_ctor_set(x_2, 0, x_8); +return x_2; +} +else +{ +lean_object* x_9; lean_object* x_10; size_t x_11; size_t x_12; lean_object* x_13; lean_object* x_14; +x_9 = lean_ctor_get(x_2, 0); +lean_inc(x_9); +lean_dec(x_2); +x_10 = lean_array_get_size(x_9); +x_11 = lean_usize_of_nat(x_10); +lean_dec(x_10); +x_12 = 0; +x_13 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftCoreM___spec__3(x_1, x_11, x_12, x_9); +x_14 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_14, 0, x_13); +return x_14; +} +} +else +{ +uint8_t x_15; +x_15 = !lean_is_exclusive(x_2); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; size_t x_18; size_t x_19; lean_object* x_20; +x_16 = lean_ctor_get(x_2, 0); +x_17 = lean_array_get_size(x_16); +x_18 = lean_usize_of_nat(x_17); +lean_dec(x_17); +x_19 = 0; +x_20 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftCoreM___spec__4(x_1, x_18, x_19, x_16); +lean_ctor_set(x_2, 0, x_20); +return x_2; +} +else +{ +lean_object* x_21; lean_object* x_22; size_t x_23; size_t x_24; lean_object* x_25; lean_object* x_26; +x_21 = lean_ctor_get(x_2, 0); +lean_inc(x_21); +lean_dec(x_2); +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_mapMUnsafe_map___at_Lean_Elab_Command_liftCoreM___spec__4(x_1, x_23, x_24, x_21); +x_26 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_26, 0, x_25); +return x_26; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftCoreM___spec__5(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; uint8_t x_9; +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_is_exclusive(x_6); +if (x_9 == 0) +{ +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_10 = lean_ctor_get(x_6, 0); +x_11 = lean_ctor_get(x_1, 6); +x_12 = l_Lean_replaceRef(x_10, x_11); +lean_dec(x_10); +lean_ctor_set(x_6, 0, x_12); +x_13 = 1; +x_14 = lean_usize_add(x_3, x_13); +x_15 = lean_array_uset(x_8, x_3, x_6); +x_3 = x_14; +x_4 = x_15; +goto _start; +} +else +{ +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; +x_17 = lean_ctor_get(x_6, 0); +x_18 = lean_ctor_get(x_6, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_6); +x_19 = lean_ctor_get(x_1, 6); +x_20 = l_Lean_replaceRef(x_17, x_19); +lean_dec(x_17); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_18); +x_22 = 1; +x_23 = lean_usize_add(x_3, x_22); +x_24 = lean_array_uset(x_8, x_3, x_21); +x_3 = x_23; +x_4 = x_24; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapM___at_Lean_Elab_Command_liftCoreM___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_2); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; size_t x_8; size_t x_9; lean_object* x_10; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_2, 1); +x_6 = l_Lean_PersistentArray_mapMAux___at_Lean_Elab_Command_liftCoreM___spec__2(x_1, x_4); +x_7 = lean_array_get_size(x_5); +x_8 = lean_usize_of_nat(x_7); +lean_dec(x_7); +x_9 = 0; +x_10 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftCoreM___spec__5(x_1, x_8, x_9, x_5); +lean_ctor_set(x_2, 1, x_10); +lean_ctor_set(x_2, 0, x_6); +return x_2; +} +else +{ +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; size_t x_18; size_t x_19; lean_object* x_20; lean_object* x_21; +x_11 = lean_ctor_get(x_2, 0); +x_12 = lean_ctor_get(x_2, 1); +x_13 = lean_ctor_get(x_2, 2); +x_14 = lean_ctor_get_usize(x_2, 4); +x_15 = lean_ctor_get(x_2, 3); +lean_inc(x_15); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_dec(x_2); +x_16 = l_Lean_PersistentArray_mapMAux___at_Lean_Elab_Command_liftCoreM___spec__2(x_1, x_11); +x_17 = lean_array_get_size(x_12); +x_18 = lean_usize_of_nat(x_17); +lean_dec(x_17); +x_19 = 0; +x_20 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftCoreM___spec__5(x_1, x_18, x_19, x_12); +x_21 = lean_alloc_ctor(0, 4, sizeof(size_t)*1); +lean_ctor_set(x_21, 0, x_16); +lean_ctor_set(x_21, 1, x_20); +lean_ctor_set(x_21, 2, x_13); +lean_ctor_set(x_21, 3, x_15); +lean_ctor_set_usize(x_21, 4, x_14); +return x_21; +} +} +} +static lean_object* _init_l_Lean_Elab_Command_liftCoreM___rarg___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_addMessageContextPartial___at_Lean_Elab_Command_instAddMessageContextCommandElabM___spec__1___closed__2; +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_Lean_Elab_Command_liftCoreM___rarg(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; uint8_t x_16; +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_io_get_num_heartbeats(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_dec(x_8); +x_11 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkCoreContext(x_2, x_6, x_9); +x_12 = lean_ctor_get(x_6, 7); +lean_inc(x_12); +x_13 = lean_ctor_get(x_6, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_6, 6); +lean_inc(x_14); +x_15 = lean_ctor_get(x_6, 8); +lean_inc(x_15); +lean_dec(x_6); +x_16 = !lean_is_exclusive(x_12); +if (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_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; +x_17 = lean_ctor_get(x_12, 1); +lean_dec(x_17); +x_18 = lean_ctor_get(x_12, 0); +lean_dec(x_18); +x_19 = l_Lean_Elab_Command_State_infoState___default___closed__3; +x_20 = l_Lean_Elab_Command_State_messages___default___closed__3; +lean_ctor_set(x_12, 1, x_20); +lean_ctor_set(x_12, 0, x_19); +x_21 = l_Lean_Elab_Command_State_nextMacroScope___default___closed__1; +x_22 = l_Lean_Elab_Command_liftCoreM___rarg___closed__1; +x_23 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_23, 0, x_13); +lean_ctor_set(x_23, 1, x_21); +lean_ctor_set(x_23, 2, x_14); +lean_ctor_set(x_23, 3, x_15); +lean_ctor_set(x_23, 4, x_22); +lean_ctor_set(x_23, 5, x_20); +lean_ctor_set(x_23, 6, x_12); +x_103 = lean_st_mk_ref(x_23, x_10); +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); +lean_inc(x_104); +x_106 = lean_apply_3(x_1, x_11, x_104, x_105); +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; +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_alloc_ctor(1, 1, 0); +lean_ctor_set(x_109, 0, x_107); +x_110 = lean_st_ref_get(x_104, x_108); +lean_dec(x_104); +x_111 = lean_ctor_get(x_110, 0); +lean_inc(x_111); +x_112 = lean_ctor_get(x_110, 1); +lean_inc(x_112); +lean_dec(x_110); +x_113 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_113, 0, x_109); +lean_ctor_set(x_113, 1, x_111); +x_24 = x_113; +x_25 = x_112; +goto block_102; +} +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; +x_114 = lean_ctor_get(x_106, 0); +lean_inc(x_114); +x_115 = lean_ctor_get(x_106, 1); +lean_inc(x_115); +lean_dec(x_106); +x_116 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_116, 0, x_114); +x_117 = lean_st_ref_get(x_104, x_115); +lean_dec(x_104); +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); +x_120 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_120, 0, x_116); +lean_ctor_set(x_120, 1, x_118); +x_24 = x_120; +x_25 = x_119; +goto block_102; +} +block_102: +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; 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; +x_26 = lean_ctor_get(x_24, 0); +lean_inc(x_26); +x_27 = lean_ctor_get(x_24, 1); +lean_inc(x_27); +lean_dec(x_24); +x_28 = lean_st_ref_take(x_3, 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); lean_dec(x_28); x_31 = lean_ctor_get(x_27, 0); @@ -4775,7 +5613,7 @@ lean_dec(x_27); x_36 = !lean_is_exclusive(x_29); if (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; uint8_t x_44; +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; x_37 = lean_ctor_get(x_29, 1); x_38 = lean_ctor_get(x_29, 7); x_39 = lean_ctor_get(x_29, 8); @@ -4785,21 +5623,20 @@ lean_dec(x_40); x_41 = lean_ctor_get(x_29, 0); lean_dec(x_41); x_42 = l_Lean_PersistentArray_append___rarg(x_37, x_34); -lean_inc(x_33); -x_43 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore(x_2, x_42, x_33); -x_44 = !lean_is_exclusive(x_38); -if (x_44 == 0) +x_43 = !lean_is_exclusive(x_38); +if (x_43 == 0) { -lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_45 = lean_ctor_get(x_38, 1); -x_46 = lean_ctor_get(x_35, 1); -lean_inc(x_46); +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_44 = lean_ctor_get(x_38, 1); +x_45 = lean_ctor_get(x_35, 1); +lean_inc(x_45); lean_dec(x_35); -x_47 = l_Lean_PersistentArray_append___rarg(x_45, x_46); -lean_ctor_set(x_38, 1, x_47); -lean_ctor_set(x_29, 8, x_33); +x_46 = l_Lean_PersistentArray_append___rarg(x_44, x_45); +lean_ctor_set(x_38, 1, x_46); +x_47 = l_Lean_PersistentArray_mapM___at_Lean_Elab_Command_liftCoreM___spec__1(x_2, x_33); +lean_ctor_set(x_29, 8, x_47); lean_ctor_set(x_29, 6, x_32); -lean_ctor_set(x_29, 1, x_43); +lean_ctor_set(x_29, 1, x_42); lean_ctor_set(x_29, 0, x_31); x_48 = lean_st_ref_set(x_3, x_29, x_30); if (lean_obj_tag(x_26) == 0) @@ -4866,7 +5703,7 @@ return x_60; } else { -uint8_t 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_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_ctor_get_uint8(x_38, sizeof(void*)*2); x_62 = lean_ctor_get(x_38, 0); x_63 = lean_ctor_get(x_38, 1); @@ -4881,95 +5718,94 @@ x_66 = lean_alloc_ctor(0, 2, 1); lean_ctor_set(x_66, 0, x_62); lean_ctor_set(x_66, 1, x_65); lean_ctor_set_uint8(x_66, sizeof(void*)*2, x_61); -lean_ctor_set(x_29, 8, x_33); +x_67 = l_Lean_PersistentArray_mapM___at_Lean_Elab_Command_liftCoreM___spec__1(x_2, x_33); +lean_ctor_set(x_29, 8, x_67); lean_ctor_set(x_29, 7, x_66); lean_ctor_set(x_29, 6, x_32); -lean_ctor_set(x_29, 1, x_43); +lean_ctor_set(x_29, 1, x_42); lean_ctor_set(x_29, 0, x_31); -x_67 = lean_st_ref_set(x_3, x_29, x_30); +x_68 = lean_st_ref_set(x_3, x_29, x_30); if (lean_obj_tag(x_26) == 0) { -lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; -x_68 = lean_ctor_get(x_67, 1); -lean_inc(x_68); -if (lean_is_exclusive(x_67)) { - lean_ctor_release(x_67, 0); - lean_ctor_release(x_67, 1); - x_69 = x_67; +lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_69 = lean_ctor_get(x_68, 1); +lean_inc(x_69); +if (lean_is_exclusive(x_68)) { + lean_ctor_release(x_68, 0); + lean_ctor_release(x_68, 1); + x_70 = x_68; } else { - lean_dec_ref(x_67); - x_69 = lean_box(0); + lean_dec_ref(x_68); + x_70 = lean_box(0); } -x_70 = lean_ctor_get(x_26, 0); -lean_inc(x_70); +x_71 = lean_ctor_get(x_26, 0); +lean_inc(x_71); lean_dec(x_26); -if (lean_is_scalar(x_69)) { - x_71 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_70)) { + x_72 = lean_alloc_ctor(1, 2, 0); } else { - x_71 = x_69; - lean_ctor_set_tag(x_71, 1); + x_72 = x_70; + lean_ctor_set_tag(x_72, 1); } -lean_ctor_set(x_71, 0, x_70); -lean_ctor_set(x_71, 1, x_68); -return x_71; +lean_ctor_set(x_72, 0, x_71); +lean_ctor_set(x_72, 1, x_69); +return x_72; } else { -lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; -x_72 = lean_ctor_get(x_67, 1); -lean_inc(x_72); -if (lean_is_exclusive(x_67)) { - lean_ctor_release(x_67, 0); - lean_ctor_release(x_67, 1); - x_73 = x_67; +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_73 = lean_ctor_get(x_68, 1); +lean_inc(x_73); +if (lean_is_exclusive(x_68)) { + lean_ctor_release(x_68, 0); + lean_ctor_release(x_68, 1); + x_74 = x_68; } else { - lean_dec_ref(x_67); - x_73 = lean_box(0); + lean_dec_ref(x_68); + x_74 = lean_box(0); } -x_74 = lean_ctor_get(x_26, 0); -lean_inc(x_74); +x_75 = lean_ctor_get(x_26, 0); +lean_inc(x_75); lean_dec(x_26); -if (lean_is_scalar(x_73)) { - x_75 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_74)) { + x_76 = lean_alloc_ctor(0, 2, 0); } else { - x_75 = x_73; + x_76 = x_74; } -lean_ctor_set(x_75, 0, x_74); -lean_ctor_set(x_75, 1, x_72); -return x_75; +lean_ctor_set(x_76, 0, x_75); +lean_ctor_set(x_76, 1, x_73); +return x_76; } } } 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; 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; -x_76 = lean_ctor_get(x_29, 1); -x_77 = lean_ctor_get(x_29, 2); -x_78 = lean_ctor_get(x_29, 3); -x_79 = lean_ctor_get(x_29, 4); -x_80 = lean_ctor_get(x_29, 5); -x_81 = lean_ctor_get(x_29, 7); +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; 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; +x_77 = lean_ctor_get(x_29, 1); +x_78 = lean_ctor_get(x_29, 2); +x_79 = lean_ctor_get(x_29, 3); +x_80 = lean_ctor_get(x_29, 4); +x_81 = lean_ctor_get(x_29, 5); +x_82 = lean_ctor_get(x_29, 7); +lean_inc(x_82); lean_inc(x_81); lean_inc(x_80); lean_inc(x_79); lean_inc(x_78); lean_inc(x_77); -lean_inc(x_76); lean_dec(x_29); -x_82 = l_Lean_PersistentArray_append___rarg(x_76, x_34); -lean_inc(x_33); -x_83 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore(x_2, x_82, x_33); -x_84 = lean_ctor_get_uint8(x_81, sizeof(void*)*2); -x_85 = lean_ctor_get(x_81, 0); +x_83 = l_Lean_PersistentArray_append___rarg(x_77, x_34); +x_84 = lean_ctor_get_uint8(x_82, sizeof(void*)*2); +x_85 = lean_ctor_get(x_82, 0); lean_inc(x_85); -x_86 = lean_ctor_get(x_81, 1); +x_86 = lean_ctor_get(x_82, 1); lean_inc(x_86); -if (lean_is_exclusive(x_81)) { - lean_ctor_release(x_81, 0); - lean_ctor_release(x_81, 1); - x_87 = x_81; +if (lean_is_exclusive(x_82)) { + lean_ctor_release(x_82, 0); + lean_ctor_release(x_82, 1); + x_87 = x_82; } else { - lean_dec_ref(x_81); + lean_dec_ref(x_82); x_87 = lean_box(0); } x_88 = lean_ctor_get(x_35, 1); @@ -4984,219 +5820,218 @@ if (lean_is_scalar(x_87)) { lean_ctor_set(x_90, 0, x_85); lean_ctor_set(x_90, 1, x_89); lean_ctor_set_uint8(x_90, sizeof(void*)*2, x_84); -x_91 = lean_alloc_ctor(0, 9, 0); -lean_ctor_set(x_91, 0, x_31); -lean_ctor_set(x_91, 1, x_83); -lean_ctor_set(x_91, 2, x_77); -lean_ctor_set(x_91, 3, x_78); -lean_ctor_set(x_91, 4, x_79); -lean_ctor_set(x_91, 5, x_80); -lean_ctor_set(x_91, 6, x_32); -lean_ctor_set(x_91, 7, x_90); -lean_ctor_set(x_91, 8, x_33); -x_92 = lean_st_ref_set(x_3, x_91, x_30); +x_91 = l_Lean_PersistentArray_mapM___at_Lean_Elab_Command_liftCoreM___spec__1(x_2, x_33); +x_92 = lean_alloc_ctor(0, 9, 0); +lean_ctor_set(x_92, 0, x_31); +lean_ctor_set(x_92, 1, x_83); +lean_ctor_set(x_92, 2, x_78); +lean_ctor_set(x_92, 3, x_79); +lean_ctor_set(x_92, 4, x_80); +lean_ctor_set(x_92, 5, x_81); +lean_ctor_set(x_92, 6, x_32); +lean_ctor_set(x_92, 7, x_90); +lean_ctor_set(x_92, 8, x_91); +x_93 = lean_st_ref_set(x_3, x_92, x_30); if (lean_obj_tag(x_26) == 0) { -lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; -x_93 = lean_ctor_get(x_92, 1); -lean_inc(x_93); -if (lean_is_exclusive(x_92)) { - lean_ctor_release(x_92, 0); - lean_ctor_release(x_92, 1); - x_94 = x_92; +lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; +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_92); - x_94 = lean_box(0); + lean_dec_ref(x_93); + x_95 = lean_box(0); } -x_95 = lean_ctor_get(x_26, 0); -lean_inc(x_95); +x_96 = lean_ctor_get(x_26, 0); +lean_inc(x_96); lean_dec(x_26); -if (lean_is_scalar(x_94)) { - x_96 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_95)) { + x_97 = lean_alloc_ctor(1, 2, 0); } else { - x_96 = x_94; - lean_ctor_set_tag(x_96, 1); + x_97 = x_95; + lean_ctor_set_tag(x_97, 1); } -lean_ctor_set(x_96, 0, x_95); -lean_ctor_set(x_96, 1, x_93); -return x_96; +lean_ctor_set(x_97, 0, x_96); +lean_ctor_set(x_97, 1, x_94); +return x_97; } else { -lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; -x_97 = lean_ctor_get(x_92, 1); -lean_inc(x_97); -if (lean_is_exclusive(x_92)) { - lean_ctor_release(x_92, 0); - lean_ctor_release(x_92, 1); - x_98 = x_92; +lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; +x_98 = lean_ctor_get(x_93, 1); +lean_inc(x_98); +if (lean_is_exclusive(x_93)) { + lean_ctor_release(x_93, 0); + lean_ctor_release(x_93, 1); + x_99 = x_93; } else { - lean_dec_ref(x_92); - x_98 = lean_box(0); + lean_dec_ref(x_93); + x_99 = lean_box(0); } -x_99 = lean_ctor_get(x_26, 0); -lean_inc(x_99); +x_100 = lean_ctor_get(x_26, 0); +lean_inc(x_100); lean_dec(x_26); -if (lean_is_scalar(x_98)) { - x_100 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_99)) { + x_101 = lean_alloc_ctor(0, 2, 0); } else { - x_100 = x_98; + x_101 = x_99; } -lean_ctor_set(x_100, 0, x_99); -lean_ctor_set(x_100, 1, x_97); -return x_100; +lean_ctor_set(x_101, 0, x_100); +lean_ctor_set(x_101, 1, x_98); +return x_101; } } } } else { -uint8_t 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_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; -x_120 = lean_ctor_get_uint8(x_12, sizeof(void*)*2); +uint8_t 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_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; +x_121 = lean_ctor_get_uint8(x_12, sizeof(void*)*2); lean_dec(x_12); -x_121 = l_Lean_Elab_Command_State_infoState___default___closed__3; -x_122 = l_Lean_Elab_Command_State_messages___default___closed__3; -x_123 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_123, 0, x_121); -lean_ctor_set(x_123, 1, x_122); -lean_ctor_set_uint8(x_123, sizeof(void*)*2, x_120); -x_124 = l_Lean_Elab_Command_State_nextMacroScope___default___closed__1; -x_125 = l_Lean_Elab_Command_liftCoreM___rarg___closed__1; -x_126 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_126, 0, x_13); -lean_ctor_set(x_126, 1, x_124); -lean_ctor_set(x_126, 2, x_14); -lean_ctor_set(x_126, 3, x_15); -lean_ctor_set(x_126, 4, x_125); -lean_ctor_set(x_126, 5, x_122); -lean_ctor_set(x_126, 6, x_123); -x_166 = lean_st_mk_ref(x_126, x_10); -x_167 = lean_ctor_get(x_166, 0); -lean_inc(x_167); -x_168 = lean_ctor_get(x_166, 1); +x_122 = l_Lean_Elab_Command_State_infoState___default___closed__3; +x_123 = l_Lean_Elab_Command_State_messages___default___closed__3; +x_124 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_124, 0, x_122); +lean_ctor_set(x_124, 1, x_123); +lean_ctor_set_uint8(x_124, sizeof(void*)*2, x_121); +x_125 = l_Lean_Elab_Command_State_nextMacroScope___default___closed__1; +x_126 = l_Lean_Elab_Command_liftCoreM___rarg___closed__1; +x_127 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_127, 0, x_13); +lean_ctor_set(x_127, 1, x_125); +lean_ctor_set(x_127, 2, x_14); +lean_ctor_set(x_127, 3, x_15); +lean_ctor_set(x_127, 4, x_126); +lean_ctor_set(x_127, 5, x_123); +lean_ctor_set(x_127, 6, x_124); +x_167 = lean_st_mk_ref(x_127, x_10); +x_168 = lean_ctor_get(x_167, 0); lean_inc(x_168); -lean_dec(x_166); -lean_inc(x_167); -x_169 = lean_apply_3(x_1, x_11, x_167, x_168); -if (lean_obj_tag(x_169) == 0) +x_169 = lean_ctor_get(x_167, 1); +lean_inc(x_169); +lean_dec(x_167); +lean_inc(x_168); +x_170 = lean_apply_3(x_1, x_11, x_168, x_169); +if (lean_obj_tag(x_170) == 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; -x_170 = lean_ctor_get(x_169, 0); -lean_inc(x_170); -x_171 = lean_ctor_get(x_169, 1); +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; +x_171 = lean_ctor_get(x_170, 0); lean_inc(x_171); -lean_dec(x_169); -x_172 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_172, 0, x_170); -x_173 = lean_st_ref_get(x_167, x_171); -lean_dec(x_167); -x_174 = lean_ctor_get(x_173, 0); -lean_inc(x_174); -x_175 = lean_ctor_get(x_173, 1); +x_172 = lean_ctor_get(x_170, 1); +lean_inc(x_172); +lean_dec(x_170); +x_173 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_173, 0, x_171); +x_174 = lean_st_ref_get(x_168, x_172); +lean_dec(x_168); +x_175 = lean_ctor_get(x_174, 0); lean_inc(x_175); -lean_dec(x_173); -x_176 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_176, 0, x_172); -lean_ctor_set(x_176, 1, x_174); -x_127 = x_176; -x_128 = x_175; -goto block_165; +x_176 = lean_ctor_get(x_174, 1); +lean_inc(x_176); +lean_dec(x_174); +x_177 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_177, 0, x_173); +lean_ctor_set(x_177, 1, x_175); +x_128 = x_177; +x_129 = x_176; +goto block_166; } else { -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_177 = lean_ctor_get(x_169, 0); -lean_inc(x_177); -x_178 = lean_ctor_get(x_169, 1); +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_178 = lean_ctor_get(x_170, 0); lean_inc(x_178); -lean_dec(x_169); -x_179 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_179, 0, x_177); -x_180 = lean_st_ref_get(x_167, x_178); -lean_dec(x_167); -x_181 = lean_ctor_get(x_180, 0); -lean_inc(x_181); -x_182 = lean_ctor_get(x_180, 1); +x_179 = lean_ctor_get(x_170, 1); +lean_inc(x_179); +lean_dec(x_170); +x_180 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_180, 0, x_178); +x_181 = lean_st_ref_get(x_168, x_179); +lean_dec(x_168); +x_182 = lean_ctor_get(x_181, 0); lean_inc(x_182); -lean_dec(x_180); -x_183 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_183, 0, x_179); -lean_ctor_set(x_183, 1, x_181); -x_127 = x_183; -x_128 = x_182; -goto block_165; -} -block_165: -{ -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; uint8_t 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_129 = lean_ctor_get(x_127, 0); -lean_inc(x_129); -x_130 = lean_ctor_get(x_127, 1); +x_183 = lean_ctor_get(x_181, 1); +lean_inc(x_183); +lean_dec(x_181); +x_184 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_184, 0, x_180); +lean_ctor_set(x_184, 1, x_182); +x_128 = x_184; +x_129 = x_183; +goto block_166; +} +block_166: +{ +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; uint8_t 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_130 = lean_ctor_get(x_128, 0); lean_inc(x_130); -lean_dec(x_127); -x_131 = lean_st_ref_take(x_3, x_128); -x_132 = lean_ctor_get(x_131, 0); -lean_inc(x_132); -x_133 = lean_ctor_get(x_131, 1); +x_131 = lean_ctor_get(x_128, 1); +lean_inc(x_131); +lean_dec(x_128); +x_132 = lean_st_ref_take(x_3, x_129); +x_133 = lean_ctor_get(x_132, 0); lean_inc(x_133); -lean_dec(x_131); -x_134 = lean_ctor_get(x_130, 0); +x_134 = lean_ctor_get(x_132, 1); lean_inc(x_134); -x_135 = lean_ctor_get(x_130, 2); +lean_dec(x_132); +x_135 = lean_ctor_get(x_131, 0); lean_inc(x_135); -x_136 = lean_ctor_get(x_130, 3); +x_136 = lean_ctor_get(x_131, 2); lean_inc(x_136); -x_137 = lean_ctor_get(x_130, 5); +x_137 = lean_ctor_get(x_131, 3); lean_inc(x_137); -x_138 = lean_ctor_get(x_130, 6); +x_138 = lean_ctor_get(x_131, 5); lean_inc(x_138); -lean_dec(x_130); -x_139 = lean_ctor_get(x_132, 1); +x_139 = lean_ctor_get(x_131, 6); lean_inc(x_139); -x_140 = lean_ctor_get(x_132, 2); +lean_dec(x_131); +x_140 = lean_ctor_get(x_133, 1); lean_inc(x_140); -x_141 = lean_ctor_get(x_132, 3); +x_141 = lean_ctor_get(x_133, 2); lean_inc(x_141); -x_142 = lean_ctor_get(x_132, 4); +x_142 = lean_ctor_get(x_133, 3); lean_inc(x_142); -x_143 = lean_ctor_get(x_132, 5); +x_143 = lean_ctor_get(x_133, 4); lean_inc(x_143); -x_144 = lean_ctor_get(x_132, 7); +x_144 = lean_ctor_get(x_133, 5); lean_inc(x_144); -if (lean_is_exclusive(x_132)) { - lean_ctor_release(x_132, 0); - lean_ctor_release(x_132, 1); - lean_ctor_release(x_132, 2); - lean_ctor_release(x_132, 3); - lean_ctor_release(x_132, 4); - lean_ctor_release(x_132, 5); - lean_ctor_release(x_132, 6); - lean_ctor_release(x_132, 7); - lean_ctor_release(x_132, 8); - x_145 = x_132; -} else { - lean_dec_ref(x_132); - x_145 = lean_box(0); -} -x_146 = l_Lean_PersistentArray_append___rarg(x_139, x_137); -lean_inc(x_136); -x_147 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore(x_2, x_146, x_136); -x_148 = lean_ctor_get_uint8(x_144, sizeof(void*)*2); -x_149 = lean_ctor_get(x_144, 0); +x_145 = lean_ctor_get(x_133, 7); +lean_inc(x_145); +if (lean_is_exclusive(x_133)) { + lean_ctor_release(x_133, 0); + lean_ctor_release(x_133, 1); + lean_ctor_release(x_133, 2); + lean_ctor_release(x_133, 3); + lean_ctor_release(x_133, 4); + lean_ctor_release(x_133, 5); + lean_ctor_release(x_133, 6); + lean_ctor_release(x_133, 7); + lean_ctor_release(x_133, 8); + x_146 = x_133; +} else { + lean_dec_ref(x_133); + x_146 = lean_box(0); +} +x_147 = l_Lean_PersistentArray_append___rarg(x_140, x_138); +x_148 = lean_ctor_get_uint8(x_145, sizeof(void*)*2); +x_149 = lean_ctor_get(x_145, 0); lean_inc(x_149); -x_150 = lean_ctor_get(x_144, 1); +x_150 = lean_ctor_get(x_145, 1); lean_inc(x_150); -if (lean_is_exclusive(x_144)) { - lean_ctor_release(x_144, 0); - lean_ctor_release(x_144, 1); - x_151 = x_144; +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_144); + lean_dec_ref(x_145); x_151 = lean_box(0); } -x_152 = lean_ctor_get(x_138, 1); +x_152 = lean_ctor_get(x_139, 1); lean_inc(x_152); -lean_dec(x_138); +lean_dec(x_139); x_153 = l_Lean_PersistentArray_append___rarg(x_150, x_152); if (lean_is_scalar(x_151)) { x_154 = lean_alloc_ctor(0, 2, 1); @@ -5206,71 +6041,72 @@ if (lean_is_scalar(x_151)) { lean_ctor_set(x_154, 0, x_149); lean_ctor_set(x_154, 1, x_153); lean_ctor_set_uint8(x_154, sizeof(void*)*2, x_148); -if (lean_is_scalar(x_145)) { - x_155 = lean_alloc_ctor(0, 9, 0); -} else { - x_155 = x_145; -} -lean_ctor_set(x_155, 0, x_134); -lean_ctor_set(x_155, 1, x_147); -lean_ctor_set(x_155, 2, x_140); -lean_ctor_set(x_155, 3, x_141); -lean_ctor_set(x_155, 4, x_142); -lean_ctor_set(x_155, 5, x_143); -lean_ctor_set(x_155, 6, x_135); -lean_ctor_set(x_155, 7, x_154); -lean_ctor_set(x_155, 8, x_136); -x_156 = lean_st_ref_set(x_3, x_155, x_133); -if (lean_obj_tag(x_129) == 0) -{ -lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; -x_157 = lean_ctor_get(x_156, 1); -lean_inc(x_157); -if (lean_is_exclusive(x_156)) { - lean_ctor_release(x_156, 0); - lean_ctor_release(x_156, 1); - x_158 = x_156; +x_155 = l_Lean_PersistentArray_mapM___at_Lean_Elab_Command_liftCoreM___spec__1(x_2, x_137); +if (lean_is_scalar(x_146)) { + x_156 = lean_alloc_ctor(0, 9, 0); +} else { + x_156 = x_146; +} +lean_ctor_set(x_156, 0, x_135); +lean_ctor_set(x_156, 1, x_147); +lean_ctor_set(x_156, 2, x_141); +lean_ctor_set(x_156, 3, x_142); +lean_ctor_set(x_156, 4, x_143); +lean_ctor_set(x_156, 5, x_144); +lean_ctor_set(x_156, 6, x_136); +lean_ctor_set(x_156, 7, x_154); +lean_ctor_set(x_156, 8, x_155); +x_157 = lean_st_ref_set(x_3, x_156, x_134); +if (lean_obj_tag(x_130) == 0) +{ +lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; +x_158 = lean_ctor_get(x_157, 1); +lean_inc(x_158); +if (lean_is_exclusive(x_157)) { + lean_ctor_release(x_157, 0); + lean_ctor_release(x_157, 1); + x_159 = x_157; } else { - lean_dec_ref(x_156); - x_158 = lean_box(0); + lean_dec_ref(x_157); + x_159 = lean_box(0); } -x_159 = lean_ctor_get(x_129, 0); -lean_inc(x_159); -lean_dec(x_129); -if (lean_is_scalar(x_158)) { - x_160 = lean_alloc_ctor(1, 2, 0); +x_160 = lean_ctor_get(x_130, 0); +lean_inc(x_160); +lean_dec(x_130); +if (lean_is_scalar(x_159)) { + x_161 = lean_alloc_ctor(1, 2, 0); } else { - x_160 = x_158; - lean_ctor_set_tag(x_160, 1); + x_161 = x_159; + lean_ctor_set_tag(x_161, 1); } -lean_ctor_set(x_160, 0, x_159); -lean_ctor_set(x_160, 1, x_157); -return x_160; +lean_ctor_set(x_161, 0, x_160); +lean_ctor_set(x_161, 1, x_158); +return x_161; } else { -lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; -x_161 = lean_ctor_get(x_156, 1); -lean_inc(x_161); -if (lean_is_exclusive(x_156)) { - lean_ctor_release(x_156, 0); - lean_ctor_release(x_156, 1); - x_162 = x_156; +lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; +x_162 = lean_ctor_get(x_157, 1); +lean_inc(x_162); +if (lean_is_exclusive(x_157)) { + lean_ctor_release(x_157, 0); + lean_ctor_release(x_157, 1); + x_163 = x_157; } else { - lean_dec_ref(x_156); - x_162 = lean_box(0); + lean_dec_ref(x_157); + x_163 = lean_box(0); } -x_163 = lean_ctor_get(x_129, 0); -lean_inc(x_163); -lean_dec(x_129); -if (lean_is_scalar(x_162)) { - x_164 = lean_alloc_ctor(0, 2, 0); +x_164 = lean_ctor_get(x_130, 0); +lean_inc(x_164); +lean_dec(x_130); +if (lean_is_scalar(x_163)) { + x_165 = lean_alloc_ctor(0, 2, 0); } else { - x_164 = x_162; + x_165 = x_163; } -lean_ctor_set(x_164, 0, x_163); -lean_ctor_set(x_164, 1, x_161); -return x_164; +lean_ctor_set(x_165, 0, x_164); +lean_ctor_set(x_165, 1, x_162); +return x_165; } } } @@ -5284,12 +6120,70 @@ x_2 = lean_alloc_closure((void*)(l_Lean_Elab_Command_liftCoreM___rarg___boxed), return x_2; } } +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftCoreM___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_mapMUnsafe_map___at_Lean_Elab_Command_liftCoreM___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_Command_liftCoreM___spec__4___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_Lean_Elab_Command_liftCoreM___spec__4(x_1, x_5, x_6, x_4); +lean_dec(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapMAux___at_Lean_Elab_Command_liftCoreM___spec__2___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_PersistentArray_mapMAux___at_Lean_Elab_Command_liftCoreM___spec__2(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftCoreM___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; +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_Lean_Elab_Command_liftCoreM___spec__5(x_1, x_5, x_6, x_4); +lean_dec(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapM___at_Lean_Elab_Command_liftCoreM___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_PersistentArray_mapM___at_Lean_Elab_Command_liftCoreM___spec__1(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_Elab_Command_liftCoreM___rarg___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_Command_liftCoreM___rarg(x_1, x_2, x_3, x_4); lean_dec(x_3); +lean_dec(x_2); return x_5; } } @@ -6187,7 +7081,7 @@ else { lean_object* x_349; uint8_t x_350; x_349 = l_Lean_logAt___at_Lean_Elab_Command_runLinters___spec__2___closed__1; -x_350 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_14, x_349); +x_350 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_14, x_349); lean_dec(x_14); if (x_350 == 0) { @@ -7441,476 +8335,2648 @@ return x_60; } } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__4___closed__1() { +static lean_object* _init_l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__5___closed__1() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = lean_box(0); -x_2 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; +lean_object* x_1; +x_1 = l_Lean_inheritedTraceOptions; +return x_1; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___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_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -uint8_t x_9; -x_9 = lean_usize_dec_lt(x_4, x_3); -if (x_9 == 0) +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_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__5___closed__1; +x_6 = lean_st_ref_get(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_9 = lean_st_ref_get(x_3, x_8); +x_10 = !lean_is_exclusive(x_9); +if (x_10 == 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; uint8_t x_16; lean_object* x_17; +x_11 = lean_ctor_get(x_9, 0); +x_12 = lean_ctor_get(x_11, 2); +lean_inc(x_12); +lean_dec(x_11); +x_13 = l_Lean_Elab_Command_instInhabitedScope; +x_14 = l_List_head_x21___rarg(x_13, x_12); +lean_dec(x_12); +x_15 = lean_ctor_get(x_14, 1); +lean_inc(x_15); +lean_dec(x_14); +x_16 = l___private_Lean_Util_Trace_0__Lean_checkTraceOption(x_7, x_15, x_1); +lean_dec(x_15); lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_1); -x_10 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_10, 0, x_5); -lean_ctor_set(x_10, 1, x_8); -return x_10; +x_17 = lean_box(x_16); +lean_ctor_set(x_9, 0, x_17); +return x_9; } else { -lean_object* x_11; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; -lean_dec(x_5); -x_23 = lean_st_ref_get(x_7, x_8); -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_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; +x_18 = lean_ctor_get(x_9, 0); +x_19 = lean_ctor_get(x_9, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_9); +x_20 = lean_ctor_get(x_18, 2); +lean_inc(x_20); +lean_dec(x_18); +x_21 = l_Lean_Elab_Command_instInhabitedScope; +x_22 = l_List_head_x21___rarg(x_21, x_20); +lean_dec(x_20); +x_23 = lean_ctor_get(x_22, 1); +lean_inc(x_23); +lean_dec(x_22); +x_24 = l___private_Lean_Util_Trace_0__Lean_checkTraceOption(x_7, x_23, x_1); lean_dec(x_23); -x_26 = lean_array_uget(x_2, x_4); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_1); -x_27 = lean_apply_4(x_26, x_1, x_6, x_7, x_25); -if (lean_obj_tag(x_27) == 0) +lean_dec(x_7); +x_25 = lean_box(x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_19); +return x_26; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Command_runLinters___spec__6___rarg(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_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_28 = lean_ctor_get(x_27, 1); -lean_inc(x_28); -lean_dec(x_27); -x_29 = lean_st_ref_take(x_7, 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); -lean_dec(x_29); -x_32 = lean_ctor_get(x_24, 0); -lean_inc(x_32); -x_33 = lean_ctor_get(x_24, 2); -lean_inc(x_33); -x_34 = lean_ctor_get(x_24, 3); -lean_inc(x_34); -x_35 = lean_ctor_get(x_24, 4); -lean_inc(x_35); -x_36 = lean_ctor_get(x_24, 5); -lean_inc(x_36); -x_37 = lean_ctor_get(x_24, 6); -lean_inc(x_37); -x_38 = lean_ctor_get(x_24, 7); -lean_inc(x_38); -x_39 = lean_ctor_get(x_24, 8); -lean_inc(x_39); -lean_dec(x_24); -x_40 = !lean_is_exclusive(x_30); -if (x_40 == 0) +lean_object* x_3; 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_st_ref_get(x_1, x_2); +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 = lean_ctor_get(x_4, 8); +lean_inc(x_6); +lean_dec(x_4); +x_7 = lean_st_ref_take(x_1, x_5); +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 = !lean_is_exclusive(x_8); +if (x_10 == 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; uint8_t x_50; -x_41 = lean_ctor_get(x_30, 8); -lean_dec(x_41); -x_42 = lean_ctor_get(x_30, 7); -lean_dec(x_42); -x_43 = lean_ctor_get(x_30, 6); -lean_dec(x_43); -x_44 = lean_ctor_get(x_30, 5); -lean_dec(x_44); -x_45 = lean_ctor_get(x_30, 4); -lean_dec(x_45); -x_46 = lean_ctor_get(x_30, 3); -lean_dec(x_46); -x_47 = lean_ctor_get(x_30, 2); -lean_dec(x_47); -x_48 = lean_ctor_get(x_30, 0); -lean_dec(x_48); -lean_ctor_set(x_30, 8, x_39); -lean_ctor_set(x_30, 7, x_38); -lean_ctor_set(x_30, 6, x_37); -lean_ctor_set(x_30, 5, x_36); -lean_ctor_set(x_30, 4, x_35); -lean_ctor_set(x_30, 3, x_34); -lean_ctor_set(x_30, 2, x_33); -lean_ctor_set(x_30, 0, x_32); -x_49 = lean_st_ref_set(x_7, x_30, x_31); -x_50 = !lean_is_exclusive(x_49); -if (x_50 == 0) +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_11 = lean_ctor_get(x_8, 8); +lean_dec(x_11); +x_12 = l_Lean_Elab_Command_State_messages___default___closed__3; +lean_ctor_set(x_8, 8, x_12); +x_13 = lean_st_ref_set(x_1, x_8, x_9); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) { -lean_object* x_51; lean_object* x_52; -x_51 = lean_ctor_get(x_49, 0); -lean_dec(x_51); -x_52 = l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__4___closed__1; -lean_ctor_set(x_49, 0, x_52); -x_11 = x_49; -goto block_22; +lean_object* x_15; +x_15 = lean_ctor_get(x_13, 0); +lean_dec(x_15); +lean_ctor_set(x_13, 0, x_6); +return x_13; } else { -lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_53 = lean_ctor_get(x_49, 1); -lean_inc(x_53); -lean_dec(x_49); -x_54 = l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__4___closed__1; -x_55 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_55, 0, x_54); -lean_ctor_set(x_55, 1, x_53); -x_11 = x_55; -goto block_22; +lean_object* x_16; lean_object* x_17; +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); +lean_dec(x_13); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_6); +lean_ctor_set(x_17, 1, x_16); +return x_17; } } 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; -x_56 = lean_ctor_get(x_30, 1); -lean_inc(x_56); -lean_dec(x_30); -x_57 = lean_alloc_ctor(0, 9, 0); -lean_ctor_set(x_57, 0, x_32); -lean_ctor_set(x_57, 1, x_56); -lean_ctor_set(x_57, 2, x_33); -lean_ctor_set(x_57, 3, x_34); -lean_ctor_set(x_57, 4, x_35); -lean_ctor_set(x_57, 5, x_36); -lean_ctor_set(x_57, 6, x_37); -lean_ctor_set(x_57, 7, x_38); -lean_ctor_set(x_57, 8, x_39); -x_58 = lean_st_ref_set(x_7, x_57, x_31); -x_59 = lean_ctor_get(x_58, 1); -lean_inc(x_59); -if (lean_is_exclusive(x_58)) { - lean_ctor_release(x_58, 0); - lean_ctor_release(x_58, 1); - x_60 = x_58; +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_18 = lean_ctor_get(x_8, 0); +x_19 = lean_ctor_get(x_8, 1); +x_20 = lean_ctor_get(x_8, 2); +x_21 = lean_ctor_get(x_8, 3); +x_22 = lean_ctor_get(x_8, 4); +x_23 = lean_ctor_get(x_8, 5); +x_24 = lean_ctor_get(x_8, 6); +x_25 = lean_ctor_get(x_8, 7); +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_dec(x_8); +x_26 = l_Lean_Elab_Command_State_messages___default___closed__3; +x_27 = lean_alloc_ctor(0, 9, 0); +lean_ctor_set(x_27, 0, x_18); +lean_ctor_set(x_27, 1, x_19); +lean_ctor_set(x_27, 2, x_20); +lean_ctor_set(x_27, 3, x_21); +lean_ctor_set(x_27, 4, x_22); +lean_ctor_set(x_27, 5, x_23); +lean_ctor_set(x_27, 6, x_24); +lean_ctor_set(x_27, 7, x_25); +lean_ctor_set(x_27, 8, x_26); +x_28 = lean_st_ref_set(x_1, x_27, x_9); +x_29 = lean_ctor_get(x_28, 1); +lean_inc(x_29); +if (lean_is_exclusive(x_28)) { + lean_ctor_release(x_28, 0); + lean_ctor_release(x_28, 1); + x_30 = x_28; } else { - lean_dec_ref(x_58); - x_60 = lean_box(0); + lean_dec_ref(x_28); + x_30 = lean_box(0); } -x_61 = l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__4___closed__1; -if (lean_is_scalar(x_60)) { - x_62 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_30)) { + x_31 = lean_alloc_ctor(0, 2, 0); } else { - x_62 = x_60; + x_31 = x_30; +} +lean_ctor_set(x_31, 0, x_6); +lean_ctor_set(x_31, 1, x_29); +return x_31; } -lean_ctor_set(x_62, 0, x_61); -lean_ctor_set(x_62, 1, x_59); -x_11 = x_62; -goto block_22; } } -else +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Command_runLinters___spec__6(lean_object* x_1) { +_start: { -lean_object* x_63; lean_object* x_64; lean_object* x_65; -x_63 = lean_ctor_get(x_27, 0); -lean_inc(x_63); -x_64 = lean_ctor_get(x_27, 1); -lean_inc(x_64); -lean_dec(x_27); -lean_inc(x_7); +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Command_runLinters___spec__6___rarg___boxed), 2, 0); +return x_2; +} +} +static double _init_l_Lean_withSeconds___at_Lean_Elab_Command_runLinters___spec__7___closed__1() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; double x_4; +x_1 = lean_unsigned_to_nat(1000000000u); +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_Lean_withSeconds___at_Lean_Elab_Command_runLinters___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; lean_object* x_8; +x_5 = lean_io_mono_nanos_now(x_4); +x_6 = lean_ctor_get(x_5, 0); lean_inc(x_6); -x_65 = l_Lean_Elab_logException___at_Lean_Elab_Command_runLinters___spec__1(x_63, x_6, x_7, x_64); -if (lean_obj_tag(x_65) == 0) +x_7 = lean_ctor_get(x_5, 1); +lean_inc(x_7); +lean_dec(x_5); +x_8 = lean_apply_3(x_1, x_2, x_3, x_7); +if (lean_obj_tag(x_8) == 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; uint8_t x_78; -x_66 = lean_ctor_get(x_65, 1); -lean_inc(x_66); -lean_dec(x_65); -x_67 = lean_st_ref_take(x_7, x_66); -x_68 = lean_ctor_get(x_67, 0); -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_24, 0); -lean_inc(x_70); -x_71 = lean_ctor_get(x_24, 2); -lean_inc(x_71); -x_72 = lean_ctor_get(x_24, 3); -lean_inc(x_72); -x_73 = lean_ctor_get(x_24, 4); -lean_inc(x_73); -x_74 = lean_ctor_get(x_24, 5); -lean_inc(x_74); -x_75 = lean_ctor_get(x_24, 6); -lean_inc(x_75); -x_76 = lean_ctor_get(x_24, 7); -lean_inc(x_76); -x_77 = lean_ctor_get(x_24, 8); -lean_inc(x_77); +lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +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_dec(x_8); +x_11 = lean_io_mono_nanos_now(x_10); +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; double x_17; double x_18; double x_19; lean_object* x_20; lean_object* x_21; +x_13 = lean_ctor_get(x_11, 0); +x_14 = lean_nat_sub(x_13, x_6); +lean_dec(x_6); +lean_dec(x_13); +x_15 = 0; +x_16 = lean_unsigned_to_nat(0u); +x_17 = l_Float_ofScientific(x_14, x_15, x_16); +lean_dec(x_14); +x_18 = l_Lean_withSeconds___at_Lean_Elab_Command_runLinters___spec__7___closed__1; +x_19 = lean_float_div(x_17, x_18); +x_20 = lean_box_float(x_19); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_9); +lean_ctor_set(x_21, 1, x_20); +lean_ctor_set(x_11, 0, x_21); +return x_11; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; lean_object* x_26; double x_27; double x_28; double x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_22 = lean_ctor_get(x_11, 0); +x_23 = lean_ctor_get(x_11, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_11); +x_24 = lean_nat_sub(x_22, x_6); +lean_dec(x_6); +lean_dec(x_22); +x_25 = 0; +x_26 = lean_unsigned_to_nat(0u); +x_27 = l_Float_ofScientific(x_24, x_25, x_26); lean_dec(x_24); -x_78 = !lean_is_exclusive(x_68); -if (x_78 == 0) +x_28 = l_Lean_withSeconds___at_Lean_Elab_Command_runLinters___spec__7___closed__1; +x_29 = lean_float_div(x_27, x_28); +x_30 = lean_box_float(x_29); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_9); +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_23); +return x_32; +} +} +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; lean_object* x_85; lean_object* x_86; lean_object* x_87; uint8_t x_88; -x_79 = lean_ctor_get(x_68, 8); -lean_dec(x_79); -x_80 = lean_ctor_get(x_68, 7); -lean_dec(x_80); -x_81 = lean_ctor_get(x_68, 6); -lean_dec(x_81); -x_82 = lean_ctor_get(x_68, 5); -lean_dec(x_82); -x_83 = lean_ctor_get(x_68, 4); -lean_dec(x_83); -x_84 = lean_ctor_get(x_68, 3); -lean_dec(x_84); -x_85 = lean_ctor_get(x_68, 2); -lean_dec(x_85); -x_86 = lean_ctor_get(x_68, 0); -lean_dec(x_86); -lean_ctor_set(x_68, 8, x_77); -lean_ctor_set(x_68, 7, x_76); -lean_ctor_set(x_68, 6, x_75); -lean_ctor_set(x_68, 5, x_74); -lean_ctor_set(x_68, 4, x_73); -lean_ctor_set(x_68, 3, x_72); -lean_ctor_set(x_68, 2, x_71); -lean_ctor_set(x_68, 0, x_70); -x_87 = lean_st_ref_set(x_7, x_68, x_69); -x_88 = !lean_is_exclusive(x_87); -if (x_88 == 0) +uint8_t x_33; +lean_dec(x_6); +x_33 = !lean_is_exclusive(x_8); +if (x_33 == 0) { -lean_object* x_89; lean_object* x_90; -x_89 = lean_ctor_get(x_87, 0); -lean_dec(x_89); -x_90 = l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__4___closed__1; -lean_ctor_set(x_87, 0, x_90); -x_11 = x_87; -goto block_22; +return x_8; } else { -lean_object* x_91; lean_object* x_92; lean_object* x_93; -x_91 = lean_ctor_get(x_87, 1); -lean_inc(x_91); -lean_dec(x_87); -x_92 = l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__4___closed__1; -x_93 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_93, 0, x_92); -lean_ctor_set(x_93, 1, x_91); -x_11 = x_93; -goto block_22; +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_8, 0); +x_35 = lean_ctor_get(x_8, 1); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_8); +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_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Elab_Command_runLinters___spec__9(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) { +_start: +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_9 = lean_st_ref_take(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_is_exclusive(x_10); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; size_t x_16; size_t 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; +x_13 = lean_ctor_get(x_10, 8); +x_14 = l_Lean_PersistentArray_toArray___rarg(x_13); +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___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___spec__1(x_16, x_17, x_14); +x_19 = lean_alloc_ctor(9, 3, 1); +lean_ctor_set(x_19, 0, x_2); +lean_ctor_set(x_19, 1, x_4); +lean_ctor_set(x_19, 2, x_18); +lean_ctor_set_uint8(x_19, sizeof(void*)*3, x_5); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_3); +lean_ctor_set(x_20, 1, x_19); +x_21 = l_Lean_PersistentArray_push___rarg(x_1, x_20); +lean_ctor_set(x_10, 8, x_21); +x_22 = lean_st_ref_set(x_7, x_10, x_11); +x_23 = !lean_is_exclusive(x_22); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; +x_24 = lean_ctor_get(x_22, 0); +lean_dec(x_24); +x_25 = lean_box(0); +lean_ctor_set(x_22, 0, x_25); +return x_22; +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_22, 1); +lean_inc(x_26); +lean_dec(x_22); +x_27 = lean_box(0); +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; } } else { -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_94 = lean_ctor_get(x_68, 1); -lean_inc(x_94); -lean_dec(x_68); -x_95 = lean_alloc_ctor(0, 9, 0); -lean_ctor_set(x_95, 0, x_70); -lean_ctor_set(x_95, 1, x_94); -lean_ctor_set(x_95, 2, x_71); -lean_ctor_set(x_95, 3, x_72); -lean_ctor_set(x_95, 4, x_73); -lean_ctor_set(x_95, 5, x_74); -lean_ctor_set(x_95, 6, x_75); -lean_ctor_set(x_95, 7, x_76); -lean_ctor_set(x_95, 8, x_77); -x_96 = lean_st_ref_set(x_7, x_95, x_69); -x_97 = lean_ctor_get(x_96, 1); -lean_inc(x_97); -if (lean_is_exclusive(x_96)) { - lean_ctor_release(x_96, 0); - lean_ctor_release(x_96, 1); - x_98 = x_96; +lean_object* x_29; lean_object* x_30; lean_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; size_t x_40; size_t x_41; 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_29 = lean_ctor_get(x_10, 0); +x_30 = lean_ctor_get(x_10, 1); +x_31 = lean_ctor_get(x_10, 2); +x_32 = lean_ctor_get(x_10, 3); +x_33 = lean_ctor_get(x_10, 4); +x_34 = lean_ctor_get(x_10, 5); +x_35 = lean_ctor_get(x_10, 6); +x_36 = lean_ctor_get(x_10, 7); +x_37 = lean_ctor_get(x_10, 8); +lean_inc(x_37); +lean_inc(x_36); +lean_inc(x_35); +lean_inc(x_34); +lean_inc(x_33); +lean_inc(x_32); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_29); +lean_dec(x_10); +x_38 = l_Lean_PersistentArray_toArray___rarg(x_37); +x_39 = lean_array_get_size(x_38); +x_40 = lean_usize_of_nat(x_39); +lean_dec(x_39); +x_41 = 0; +x_42 = l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___spec__1(x_40, x_41, x_38); +x_43 = lean_alloc_ctor(9, 3, 1); +lean_ctor_set(x_43, 0, x_2); +lean_ctor_set(x_43, 1, x_4); +lean_ctor_set(x_43, 2, x_42); +lean_ctor_set_uint8(x_43, sizeof(void*)*3, x_5); +x_44 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_44, 0, x_3); +lean_ctor_set(x_44, 1, x_43); +x_45 = l_Lean_PersistentArray_push___rarg(x_1, x_44); +x_46 = lean_alloc_ctor(0, 9, 0); +lean_ctor_set(x_46, 0, x_29); +lean_ctor_set(x_46, 1, x_30); +lean_ctor_set(x_46, 2, x_31); +lean_ctor_set(x_46, 3, x_32); +lean_ctor_set(x_46, 4, x_33); +lean_ctor_set(x_46, 5, x_34); +lean_ctor_set(x_46, 6, x_35); +lean_ctor_set(x_46, 7, x_36); +lean_ctor_set(x_46, 8, x_45); +x_47 = lean_st_ref_set(x_7, x_46, x_11); +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_96); - x_98 = lean_box(0); + 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___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Command_runLinters___spec__8(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) { +_start: +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_9 = l_Lean_Elab_Command_getRef(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_replaceRef(x_3, x_10); +lean_dec(x_10); +x_13 = !lean_is_exclusive(x_6); +if (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_6, 6); +lean_dec(x_14); +lean_ctor_set(x_6, 6, x_12); +x_15 = l_Lean_addMessageContextPartial___at_Lean_Elab_Command_instAddMessageContextCommandElabM___spec__1(x_4, x_6, x_7, x_11); +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___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Elab_Command_runLinters___spec__9(x_1, x_2, x_3, x_16, x_5, x_6, x_7, x_17); +lean_dec(x_6); +return x_18; +} +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_ctor_get(x_6, 0); +x_20 = lean_ctor_get(x_6, 1); +x_21 = lean_ctor_get(x_6, 2); +x_22 = lean_ctor_get(x_6, 3); +x_23 = lean_ctor_get(x_6, 4); +x_24 = lean_ctor_get(x_6, 5); +x_25 = lean_ctor_get(x_6, 7); +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_dec(x_6); +x_26 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_26, 0, x_19); +lean_ctor_set(x_26, 1, x_20); +lean_ctor_set(x_26, 2, x_21); +lean_ctor_set(x_26, 3, x_22); +lean_ctor_set(x_26, 4, x_23); +lean_ctor_set(x_26, 5, x_24); +lean_ctor_set(x_26, 6, x_12); +lean_ctor_set(x_26, 7, x_25); +x_27 = l_Lean_addMessageContextPartial___at_Lean_Elab_Command_instAddMessageContextCommandElabM___spec__1(x_4, x_26, x_7, x_11); +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_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Elab_Command_runLinters___spec__9(x_1, x_2, x_3, x_28, x_5, x_26, x_7, x_29); +lean_dec(x_26); +return x_30; +} +} +} +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Command_runLinters___spec__10(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_5; lean_object* x_6; +x_5 = lean_ctor_get(x_1, 0); +lean_inc(x_5); +x_6 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_6, 0, x_5); +lean_ctor_set(x_6, 1, x_4); +return x_6; +} +else +{ +lean_object* x_7; lean_object* x_8; +x_7 = lean_ctor_get(x_1, 0); +lean_inc(x_7); +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; +} +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = lean_apply_3(x_1, x_2, x_3, x_4); +if (lean_obj_tag(x_5) == 0) +{ +uint8_t x_6; +x_6 = !lean_is_exclusive(x_5); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; +x_7 = lean_ctor_get(x_5, 0); +x_8 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_8, 0, x_7); +lean_ctor_set(x_5, 0, x_8); +return x_5; +} +else +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_9 = lean_ctor_get(x_5, 0); +x_10 = lean_ctor_get(x_5, 1); +lean_inc(x_10); +lean_inc(x_9); +lean_dec(x_5); +x_11 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_11, 0, x_9); +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; +} +} +else +{ +uint8_t x_13; +x_13 = !lean_is_exclusive(x_5); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_ctor_get(x_5, 0); +x_15 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set_tag(x_5, 0); +lean_ctor_set(x_5, 0, x_15); +return x_5; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_16 = lean_ctor_get(x_5, 0); +x_17 = lean_ctor_get(x_5, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_5); +x_18 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_18, 0, x_16); +x_19 = lean_alloc_ctor(0, 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_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___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) { +_start: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; +lean_inc(x_8); +x_11 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Command_runLinters___spec__8(x_1, x_2, x_3, x_6, x_4, x_8, x_9, x_10); +x_12 = lean_ctor_get(x_11, 1); +lean_inc(x_12); +lean_dec(x_11); +x_13 = l_MonadExcept_ofExcept___at_Lean_Elab_Command_runLinters___spec__10(x_5, x_8, x_9, x_12); +lean_dec(x_8); +return x_13; +} +} +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_profiler; +return x_1; +} +} +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("[", 1); +return x_1; +} +} +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3___closed__2; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("s] ", 3); +return x_1; +} +} +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3___closed__4; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, uint8_t x_7, double 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_dec(x_9); +x_13 = l_Lean_Elab_Command_getRef(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); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_2); +x_16 = lean_apply_4(x_1, x_2, x_10, x_11, x_15); +if (lean_obj_tag(x_16) == 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); +x_19 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3___closed__1; +x_20 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_6, x_19); +lean_dec(x_6); +if (x_20 == 0) +{ +if (x_7 == 0) +{ +lean_object* x_21; lean_object* x_22; +x_21 = lean_box(0); +x_22 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__2(x_3, x_4, x_14, x_5, x_2, x_17, x_21, x_10, x_11, x_18); +lean_dec(x_11); +lean_dec(x_2); +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; +x_23 = lean_float_to_string(x_8); +x_24 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_24, 0, x_23); +x_25 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_25, 0, x_24); +x_26 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3___closed__3; +x_27 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_25); +x_28 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3___closed__5; +x_29 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +x_30 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_17); +x_31 = l_Lean_Elab_logException___at_Lean_Elab_Command_runLinters___spec__1___closed__3; +x_32 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +x_33 = lean_box(0); +x_34 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__2(x_3, x_4, x_14, x_5, x_2, x_32, x_33, x_10, x_11, x_18); +lean_dec(x_11); +lean_dec(x_2); +return x_34; +} +} +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; +x_35 = lean_float_to_string(x_8); +x_36 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_36, 0, x_35); +x_37 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_37, 0, x_36); +x_38 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3___closed__3; +x_39 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_37); +x_40 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3___closed__5; +x_41 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_40); +x_42 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_17); +x_43 = l_Lean_Elab_logException___at_Lean_Elab_Command_runLinters___spec__1___closed__3; +x_44 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_44, 0, x_42); +lean_ctor_set(x_44, 1, x_43); +x_45 = lean_box(0); +x_46 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__2(x_3, x_4, x_14, x_5, x_2, x_44, x_45, x_10, x_11, x_18); +lean_dec(x_11); +lean_dec(x_2); +return x_46; +} +} +else +{ +uint8_t x_47; +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_47 = !lean_is_exclusive(x_16); +if (x_47 == 0) +{ +return x_16; +} +else +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_16, 0); +x_49 = lean_ctor_get(x_16, 1); +lean_inc(x_49); +lean_inc(x_48); +lean_dec(x_16); +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; +} +} +} +} +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__4___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_trace_profiler; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t 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) { +_start: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +lean_dec(x_7); +x_11 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Command_runLinters___spec__6___rarg(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_alloc_closure((void*)(l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__1), 4, 1); +lean_closure_set(x_14, 0, x_1); +lean_inc(x_9); +lean_inc(x_8); +x_15 = l_Lean_withSeconds___at_Lean_Elab_Command_runLinters___spec__7(x_14, x_8, x_9, x_13); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; lean_object* x_67; uint8_t x_68; +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); +x_19 = lean_ctor_get(x_16, 1); +lean_inc(x_19); +lean_dec(x_16); +x_67 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__4___closed__1; +x_68 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_5, x_67); +if (x_68 == 0) +{ +uint8_t x_69; +x_69 = 0; +x_20 = x_69; +goto block_66; +} +else +{ +double x_70; double x_71; uint8_t x_72; +x_70 = l_Lean_trace_profiler_threshold_getSecs(x_5); +x_71 = lean_unbox_float(x_19); +x_72 = lean_float_decLt(x_70, x_71); +x_20 = x_72; +goto block_66; +} +block_66: +{ +if (x_6 == 0) +{ +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +lean_dec(x_19); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +x_21 = lean_st_ref_take(x_9, x_17); +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_is_exclusive(x_22); +if (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_22, 8); +x_26 = l_Lean_PersistentArray_append___rarg(x_12, x_25); +lean_ctor_set(x_22, 8, x_26); +x_27 = lean_st_ref_set(x_9, x_22, x_23); +x_28 = lean_ctor_get(x_27, 1); +lean_inc(x_28); +lean_dec(x_27); +x_29 = l_MonadExcept_ofExcept___at_Lean_Elab_Command_runLinters___spec__10(x_18, x_8, x_9, x_28); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_18); +if (lean_obj_tag(x_29) == 0) +{ +uint8_t x_30; +x_30 = !lean_is_exclusive(x_29); +if (x_30 == 0) +{ +return x_29; +} +else +{ +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 +{ +uint8_t x_34; +x_34 = !lean_is_exclusive(x_29); +if (x_34 == 0) +{ +return x_29; +} +else +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_29, 0); +x_36 = lean_ctor_get(x_29, 1); +lean_inc(x_36); +lean_inc(x_35); +lean_dec(x_29); +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_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; 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_38 = lean_ctor_get(x_22, 0); +x_39 = lean_ctor_get(x_22, 1); +x_40 = lean_ctor_get(x_22, 2); +x_41 = lean_ctor_get(x_22, 3); +x_42 = lean_ctor_get(x_22, 4); +x_43 = lean_ctor_get(x_22, 5); +x_44 = lean_ctor_get(x_22, 6); +x_45 = lean_ctor_get(x_22, 7); +x_46 = lean_ctor_get(x_22, 8); +lean_inc(x_46); +lean_inc(x_45); +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_dec(x_22); +x_47 = l_Lean_PersistentArray_append___rarg(x_12, x_46); +x_48 = lean_alloc_ctor(0, 9, 0); +lean_ctor_set(x_48, 0, x_38); +lean_ctor_set(x_48, 1, x_39); +lean_ctor_set(x_48, 2, x_40); +lean_ctor_set(x_48, 3, x_41); +lean_ctor_set(x_48, 4, x_42); +lean_ctor_set(x_48, 5, x_43); +lean_ctor_set(x_48, 6, x_44); +lean_ctor_set(x_48, 7, x_45); +lean_ctor_set(x_48, 8, x_47); +x_49 = lean_st_ref_set(x_9, x_48, x_23); +x_50 = lean_ctor_get(x_49, 1); +lean_inc(x_50); +lean_dec(x_49); +x_51 = l_MonadExcept_ofExcept___at_Lean_Elab_Command_runLinters___spec__10(x_18, x_8, x_9, x_50); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_18); +if (lean_obj_tag(x_51) == 0) +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_52 = lean_ctor_get(x_51, 0); +lean_inc(x_52); +x_53 = lean_ctor_get(x_51, 1); +lean_inc(x_53); +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_51); + x_54 = lean_box(0); +} +if (lean_is_scalar(x_54)) { + x_55 = lean_alloc_ctor(0, 2, 0); +} else { + x_55 = x_54; +} +lean_ctor_set(x_55, 0, x_52); +lean_ctor_set(x_55, 1, x_53); +return x_55; +} +else +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_56 = lean_ctor_get(x_51, 0); +lean_inc(x_56); +x_57 = lean_ctor_get(x_51, 1); +lean_inc(x_57); +if (lean_is_exclusive(x_51)) { + lean_ctor_release(x_51, 0); + lean_ctor_release(x_51, 1); + x_58 = x_51; +} else { + lean_dec_ref(x_51); + x_58 = lean_box(0); +} +if (lean_is_scalar(x_58)) { + x_59 = lean_alloc_ctor(1, 2, 0); +} else { + x_59 = x_58; +} +lean_ctor_set(x_59, 0, x_56); +lean_ctor_set(x_59, 1, x_57); +return x_59; +} +} +} +else +{ +lean_object* x_60; double x_61; lean_object* x_62; +x_60 = lean_box(0); +x_61 = lean_unbox_float(x_19); +lean_dec(x_19); +x_62 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3(x_2, x_18, x_12, x_3, x_4, x_5, x_20, x_61, x_60, x_8, x_9, x_17); +return x_62; +} +} +else +{ +lean_object* x_63; double x_64; lean_object* x_65; +x_63 = lean_box(0); +x_64 = lean_unbox_float(x_19); +lean_dec(x_19); +x_65 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3(x_2, x_18, x_12, x_3, x_4, x_5, x_20, x_64, x_63, x_8, x_9, x_17); +return x_65; +} +} +} +else +{ +uint8_t x_73; +lean_dec(x_12); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +x_73 = !lean_is_exclusive(x_15); +if (x_73 == 0) +{ +return x_15; +} +else +{ +lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_74 = lean_ctor_get(x_15, 0); +x_75 = lean_ctor_get(x_15, 1); +lean_inc(x_75); +lean_inc(x_74); +lean_dec(x_15); +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_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4(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) { +_start: +{ +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_st_ref_get(x_6, 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_dec(x_8); +x_11 = lean_ctor_get(x_9, 2); +lean_inc(x_11); +lean_dec(x_9); +x_12 = l_Lean_Elab_Command_instInhabitedScope; +x_13 = l_List_head_x21___rarg(x_12, x_11); +lean_dec(x_11); +x_14 = lean_ctor_get(x_13, 1); +lean_inc(x_14); +lean_dec(x_13); +lean_inc(x_1); +x_15 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__5(x_1, x_5, x_6, x_10); +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_unbox(x_16); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_18 = lean_ctor_get(x_15, 1); +lean_inc(x_18); +lean_dec(x_15); +x_19 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__4___closed__1; +x_20 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_14, x_19); +if (x_20 == 0) +{ +lean_object* x_21; +lean_dec(x_16); +lean_dec(x_14); +lean_dec(x_2); +lean_dec(x_1); +x_21 = lean_apply_3(x_3, x_5, x_6, x_18); +if (lean_obj_tag(x_21) == 0) +{ +uint8_t x_22; +x_22 = !lean_is_exclusive(x_21); +if (x_22 == 0) +{ +return x_21; +} +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; +} +} +else +{ +uint8_t x_26; +x_26 = !lean_is_exclusive(x_21); +if (x_26 == 0) +{ +return x_21; +} +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_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 +{ +lean_object* x_30; uint8_t x_31; lean_object* x_32; +x_30 = lean_box(0); +x_31 = lean_unbox(x_16); +lean_dec(x_16); +x_32 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__4(x_3, x_2, x_1, x_4, x_14, x_31, x_30, x_5, x_6, x_18); +return x_32; +} +} +else +{ +lean_object* x_33; lean_object* x_34; uint8_t x_35; lean_object* x_36; +x_33 = lean_ctor_get(x_15, 1); +lean_inc(x_33); +lean_dec(x_15); +x_34 = lean_box(0); +x_35 = lean_unbox(x_16); +lean_dec(x_16); +x_36 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__4(x_3, x_2, x_1, x_4, x_14, x_35, x_34, x_5, x_6, x_33); +return x_36; +} +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__11___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("running linter: ", 16); +return x_1; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__11___lambda__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__11___lambda__1___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__11___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; +x_6 = lean_ctor_get(x_1, 1); +lean_inc(x_6); +lean_dec(x_1); +x_7 = l_Lean_MessageData_ofName(x_6); +x_8 = l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__11___lambda__1___closed__2; +x_9 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_9, 0, x_8); +lean_ctor_set(x_9, 1, x_7); +x_10 = l_Lean_Elab_logException___at_Lean_Elab_Command_runLinters___spec__1___closed__3; +x_11 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_11, 0, x_9); +lean_ctor_set(x_11, 1, x_10); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_5); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__11___lambda__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; lean_object* x_10; +x_6 = lean_st_ref_get(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 = lean_ctor_get(x_1, 0); +lean_inc(x_9); +lean_dec(x_1); +lean_inc(x_4); +lean_inc(x_3); +x_10 = lean_apply_4(x_9, x_2, x_3, x_4, x_8); +if (lean_obj_tag(x_10) == 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; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +lean_dec(x_3); +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_st_ref_take(x_4, 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_7, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_7, 2); +lean_inc(x_17); +x_18 = lean_ctor_get(x_7, 3); +lean_inc(x_18); +x_19 = lean_ctor_get(x_7, 4); +lean_inc(x_19); +x_20 = lean_ctor_get(x_7, 5); +lean_inc(x_20); +x_21 = lean_ctor_get(x_7, 6); +lean_inc(x_21); +x_22 = lean_ctor_get(x_7, 7); +lean_inc(x_22); +x_23 = lean_ctor_get(x_7, 8); +lean_inc(x_23); +lean_dec(x_7); +x_24 = !lean_is_exclusive(x_14); +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; uint8_t x_34; +x_25 = lean_ctor_get(x_14, 8); +lean_dec(x_25); +x_26 = lean_ctor_get(x_14, 7); +lean_dec(x_26); +x_27 = lean_ctor_get(x_14, 6); +lean_dec(x_27); +x_28 = lean_ctor_get(x_14, 5); +lean_dec(x_28); +x_29 = lean_ctor_get(x_14, 4); +lean_dec(x_29); +x_30 = lean_ctor_get(x_14, 3); +lean_dec(x_30); +x_31 = lean_ctor_get(x_14, 2); +lean_dec(x_31); +x_32 = lean_ctor_get(x_14, 0); +lean_dec(x_32); +lean_ctor_set(x_14, 8, x_23); +lean_ctor_set(x_14, 7, x_22); +lean_ctor_set(x_14, 6, x_21); +lean_ctor_set(x_14, 5, x_20); +lean_ctor_set(x_14, 4, x_19); +lean_ctor_set(x_14, 3, x_18); +lean_ctor_set(x_14, 2, x_17); +lean_ctor_set(x_14, 0, x_16); +x_33 = lean_st_ref_set(x_4, x_14, x_15); +lean_dec(x_4); +x_34 = !lean_is_exclusive(x_33); +if (x_34 == 0) +{ +lean_object* x_35; +x_35 = lean_ctor_get(x_33, 0); +lean_dec(x_35); +lean_ctor_set(x_33, 0, x_11); +return x_33; +} +else +{ +lean_object* x_36; lean_object* x_37; +x_36 = lean_ctor_get(x_33, 1); +lean_inc(x_36); +lean_dec(x_33); +x_37 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_37, 0, x_11); +lean_ctor_set(x_37, 1, x_36); +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; +x_38 = lean_ctor_get(x_14, 1); +lean_inc(x_38); +lean_dec(x_14); +x_39 = lean_alloc_ctor(0, 9, 0); +lean_ctor_set(x_39, 0, x_16); +lean_ctor_set(x_39, 1, x_38); +lean_ctor_set(x_39, 2, x_17); +lean_ctor_set(x_39, 3, x_18); +lean_ctor_set(x_39, 4, x_19); +lean_ctor_set(x_39, 5, x_20); +lean_ctor_set(x_39, 6, x_21); +lean_ctor_set(x_39, 7, x_22); +lean_ctor_set(x_39, 8, x_23); +x_40 = lean_st_ref_set(x_4, x_39, x_15); +lean_dec(x_4); +x_41 = lean_ctor_get(x_40, 1); +lean_inc(x_41); +if (lean_is_exclusive(x_40)) { + lean_ctor_release(x_40, 0); + lean_ctor_release(x_40, 1); + x_42 = x_40; +} else { + lean_dec_ref(x_40); + x_42 = lean_box(0); +} +if (lean_is_scalar(x_42)) { + x_43 = lean_alloc_ctor(0, 2, 0); +} else { + x_43 = x_42; +} +lean_ctor_set(x_43, 0, x_11); +lean_ctor_set(x_43, 1, x_41); +return x_43; +} +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_10, 0); +lean_inc(x_44); +x_45 = lean_ctor_get(x_10, 1); +lean_inc(x_45); +lean_dec(x_10); +lean_inc(x_4); +x_46 = l_Lean_Elab_logException___at_Lean_Elab_Command_runLinters___spec__1(x_44, x_3, x_4, x_45); +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_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; +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 = lean_st_ref_take(x_4, 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 = lean_ctor_get(x_7, 0); +lean_inc(x_52); +x_53 = lean_ctor_get(x_7, 2); +lean_inc(x_53); +x_54 = lean_ctor_get(x_7, 3); +lean_inc(x_54); +x_55 = lean_ctor_get(x_7, 4); +lean_inc(x_55); +x_56 = lean_ctor_get(x_7, 5); +lean_inc(x_56); +x_57 = lean_ctor_get(x_7, 6); +lean_inc(x_57); +x_58 = lean_ctor_get(x_7, 7); +lean_inc(x_58); +x_59 = lean_ctor_get(x_7, 8); +lean_inc(x_59); +lean_dec(x_7); +x_60 = !lean_is_exclusive(x_50); +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; uint8_t x_70; +x_61 = lean_ctor_get(x_50, 8); +lean_dec(x_61); +x_62 = lean_ctor_get(x_50, 7); +lean_dec(x_62); +x_63 = lean_ctor_get(x_50, 6); +lean_dec(x_63); +x_64 = lean_ctor_get(x_50, 5); +lean_dec(x_64); +x_65 = lean_ctor_get(x_50, 4); +lean_dec(x_65); +x_66 = lean_ctor_get(x_50, 3); +lean_dec(x_66); +x_67 = lean_ctor_get(x_50, 2); +lean_dec(x_67); +x_68 = lean_ctor_get(x_50, 0); +lean_dec(x_68); +lean_ctor_set(x_50, 8, x_59); +lean_ctor_set(x_50, 7, x_58); +lean_ctor_set(x_50, 6, x_57); +lean_ctor_set(x_50, 5, x_56); +lean_ctor_set(x_50, 4, x_55); +lean_ctor_set(x_50, 3, x_54); +lean_ctor_set(x_50, 2, x_53); +lean_ctor_set(x_50, 0, x_52); +x_69 = lean_st_ref_set(x_4, x_50, x_51); +lean_dec(x_4); +x_70 = !lean_is_exclusive(x_69); +if (x_70 == 0) +{ +lean_object* x_71; +x_71 = lean_ctor_get(x_69, 0); +lean_dec(x_71); +lean_ctor_set(x_69, 0, x_47); +return x_69; +} +else +{ +lean_object* x_72; lean_object* x_73; +x_72 = lean_ctor_get(x_69, 1); +lean_inc(x_72); +lean_dec(x_69); +x_73 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_73, 0, x_47); +lean_ctor_set(x_73, 1, x_72); +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; +x_74 = lean_ctor_get(x_50, 1); +lean_inc(x_74); +lean_dec(x_50); +x_75 = lean_alloc_ctor(0, 9, 0); +lean_ctor_set(x_75, 0, x_52); +lean_ctor_set(x_75, 1, x_74); +lean_ctor_set(x_75, 2, x_53); +lean_ctor_set(x_75, 3, x_54); +lean_ctor_set(x_75, 4, x_55); +lean_ctor_set(x_75, 5, x_56); +lean_ctor_set(x_75, 6, x_57); +lean_ctor_set(x_75, 7, x_58); +lean_ctor_set(x_75, 8, x_59); +x_76 = lean_st_ref_set(x_4, x_75, x_51); +lean_dec(x_4); +x_77 = lean_ctor_get(x_76, 1); +lean_inc(x_77); +if (lean_is_exclusive(x_76)) { + lean_ctor_release(x_76, 0); + lean_ctor_release(x_76, 1); + x_78 = x_76; +} else { + lean_dec_ref(x_76); + x_78 = lean_box(0); +} +if (lean_is_scalar(x_78)) { + x_79 = lean_alloc_ctor(0, 2, 0); +} else { + x_79 = x_78; +} +lean_ctor_set(x_79, 0, x_47); +lean_ctor_set(x_79, 1, x_77); +return x_79; +} +} +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; uint8_t x_93; +x_80 = lean_ctor_get(x_46, 0); +lean_inc(x_80); +x_81 = lean_ctor_get(x_46, 1); +lean_inc(x_81); +lean_dec(x_46); +x_82 = lean_st_ref_take(x_4, x_81); +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_ctor_get(x_7, 0); +lean_inc(x_85); +x_86 = lean_ctor_get(x_7, 2); +lean_inc(x_86); +x_87 = lean_ctor_get(x_7, 3); +lean_inc(x_87); +x_88 = lean_ctor_get(x_7, 4); +lean_inc(x_88); +x_89 = lean_ctor_get(x_7, 5); +lean_inc(x_89); +x_90 = lean_ctor_get(x_7, 6); +lean_inc(x_90); +x_91 = lean_ctor_get(x_7, 7); +lean_inc(x_91); +x_92 = lean_ctor_get(x_7, 8); +lean_inc(x_92); +lean_dec(x_7); +x_93 = !lean_is_exclusive(x_83); +if (x_93 == 0) +{ +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; uint8_t x_103; +x_94 = lean_ctor_get(x_83, 8); +lean_dec(x_94); +x_95 = lean_ctor_get(x_83, 7); +lean_dec(x_95); +x_96 = lean_ctor_get(x_83, 6); +lean_dec(x_96); +x_97 = lean_ctor_get(x_83, 5); +lean_dec(x_97); +x_98 = lean_ctor_get(x_83, 4); +lean_dec(x_98); +x_99 = lean_ctor_get(x_83, 3); +lean_dec(x_99); +x_100 = lean_ctor_get(x_83, 2); +lean_dec(x_100); +x_101 = lean_ctor_get(x_83, 0); +lean_dec(x_101); +lean_ctor_set(x_83, 8, x_92); +lean_ctor_set(x_83, 7, x_91); +lean_ctor_set(x_83, 6, x_90); +lean_ctor_set(x_83, 5, x_89); +lean_ctor_set(x_83, 4, x_88); +lean_ctor_set(x_83, 3, x_87); +lean_ctor_set(x_83, 2, x_86); +lean_ctor_set(x_83, 0, x_85); +x_102 = lean_st_ref_set(x_4, x_83, x_84); +lean_dec(x_4); +x_103 = !lean_is_exclusive(x_102); +if (x_103 == 0) +{ +lean_object* x_104; +x_104 = lean_ctor_get(x_102, 0); +lean_dec(x_104); +lean_ctor_set_tag(x_102, 1); +lean_ctor_set(x_102, 0, x_80); +return x_102; +} +else +{ +lean_object* x_105; lean_object* x_106; +x_105 = lean_ctor_get(x_102, 1); +lean_inc(x_105); +lean_dec(x_102); +x_106 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_106, 0, x_80); +lean_ctor_set(x_106, 1, x_105); +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; +x_107 = lean_ctor_get(x_83, 1); +lean_inc(x_107); +lean_dec(x_83); +x_108 = lean_alloc_ctor(0, 9, 0); +lean_ctor_set(x_108, 0, x_85); +lean_ctor_set(x_108, 1, x_107); +lean_ctor_set(x_108, 2, x_86); +lean_ctor_set(x_108, 3, x_87); +lean_ctor_set(x_108, 4, x_88); +lean_ctor_set(x_108, 5, x_89); +lean_ctor_set(x_108, 6, x_90); +lean_ctor_set(x_108, 7, x_91); +lean_ctor_set(x_108, 8, x_92); +x_109 = lean_st_ref_set(x_4, x_108, x_84); +lean_dec(x_4); +x_110 = lean_ctor_get(x_109, 1); +lean_inc(x_110); +if (lean_is_exclusive(x_109)) { + lean_ctor_release(x_109, 0); + lean_ctor_release(x_109, 1); + x_111 = x_109; +} else { + lean_dec_ref(x_109); + x_111 = lean_box(0); +} +if (lean_is_scalar(x_111)) { + x_112 = lean_alloc_ctor(1, 2, 0); +} else { + x_112 = x_111; + lean_ctor_set_tag(x_112, 1); +} +lean_ctor_set(x_112, 0, x_80); +lean_ctor_set(x_112, 1, x_110); +return x_112; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___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) { +_start: +{ +uint8_t x_10; +x_10 = lean_usize_dec_lt(x_5, x_4); +if (x_10 == 0) +{ +lean_object* x_11; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_2); +lean_dec(x_1); +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_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; +lean_dec(x_6); +x_12 = lean_array_uget(x_3, x_5); +lean_inc(x_12); +x_13 = lean_alloc_closure((void*)(l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__11___lambda__1___boxed), 5, 1); +lean_closure_set(x_13, 0, x_12); +lean_inc(x_1); +x_14 = lean_alloc_closure((void*)(l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__11___lambda__2), 5, 2); +lean_closure_set(x_14, 0, x_12); +lean_closure_set(x_14, 1, x_1); +x_15 = 1; +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_2); +x_16 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4(x_2, x_13, x_14, x_15, x_7, x_8, x_9); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; size_t x_18; size_t x_19; lean_object* x_20; +x_17 = lean_ctor_get(x_16, 1); +lean_inc(x_17); +lean_dec(x_16); +x_18 = 1; +x_19 = lean_usize_add(x_5, x_18); +x_20 = lean_box(0); +x_5 = x_19; +x_6 = x_20; +x_9 = x_17; +goto _start; +} +else +{ +uint8_t x_22; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_2); +lean_dec(x_1); +x_22 = !lean_is_exclusive(x_16); +if (x_22 == 0) +{ +return x_16; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_16, 0); +x_24 = lean_ctor_get(x_16, 1); +lean_inc(x_24); +lean_inc(x_23); +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; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__13(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; +x_5 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__5___closed__1; +x_6 = lean_st_ref_get(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_9 = lean_st_ref_get(x_3, x_8); +x_10 = !lean_is_exclusive(x_9); +if (x_10 == 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; lean_object* x_17; +x_11 = lean_ctor_get(x_9, 0); +x_12 = lean_ctor_get(x_11, 2); +lean_inc(x_12); +lean_dec(x_11); +x_13 = l_Lean_Elab_Command_instInhabitedScope; +x_14 = l_List_head_x21___rarg(x_13, x_12); +lean_dec(x_12); +x_15 = lean_ctor_get(x_14, 1); +lean_inc(x_15); +lean_dec(x_14); +x_16 = l___private_Lean_Util_Trace_0__Lean_checkTraceOption(x_7, x_15, x_1); +lean_dec(x_15); +lean_dec(x_7); +x_17 = lean_box(x_16); +lean_ctor_set(x_9, 0, x_17); +return x_9; +} +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_25; lean_object* x_26; +x_18 = lean_ctor_get(x_9, 0); +x_19 = lean_ctor_get(x_9, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_9); +x_20 = lean_ctor_get(x_18, 2); +lean_inc(x_20); +lean_dec(x_18); +x_21 = l_Lean_Elab_Command_instInhabitedScope; +x_22 = l_List_head_x21___rarg(x_21, x_20); +lean_dec(x_20); +x_23 = lean_ctor_get(x_22, 1); +lean_inc(x_23); +lean_dec(x_22); +x_24 = l___private_Lean_Util_Trace_0__Lean_checkTraceOption(x_7, x_23, x_1); +lean_dec(x_23); +lean_dec(x_7); +x_25 = lean_box(x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_19); +return x_26; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Command_runLinters___spec__14___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; lean_object* x_9; uint8_t x_10; +x_3 = lean_st_ref_get(x_1, x_2); +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 = lean_ctor_get(x_4, 8); +lean_inc(x_6); +lean_dec(x_4); +x_7 = lean_st_ref_take(x_1, x_5); +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 = !lean_is_exclusive(x_8); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_11 = lean_ctor_get(x_8, 8); +lean_dec(x_11); +x_12 = l_Lean_Elab_Command_State_messages___default___closed__3; +lean_ctor_set(x_8, 8, x_12); +x_13 = lean_st_ref_set(x_1, x_8, x_9); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +lean_object* x_15; +x_15 = lean_ctor_get(x_13, 0); +lean_dec(x_15); +lean_ctor_set(x_13, 0, x_6); +return x_13; +} +else +{ +lean_object* x_16; lean_object* x_17; +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); +lean_dec(x_13); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_6); +lean_ctor_set(x_17, 1, x_16); +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; +x_18 = lean_ctor_get(x_8, 0); +x_19 = lean_ctor_get(x_8, 1); +x_20 = lean_ctor_get(x_8, 2); +x_21 = lean_ctor_get(x_8, 3); +x_22 = lean_ctor_get(x_8, 4); +x_23 = lean_ctor_get(x_8, 5); +x_24 = lean_ctor_get(x_8, 6); +x_25 = lean_ctor_get(x_8, 7); +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_dec(x_8); +x_26 = l_Lean_Elab_Command_State_messages___default___closed__3; +x_27 = lean_alloc_ctor(0, 9, 0); +lean_ctor_set(x_27, 0, x_18); +lean_ctor_set(x_27, 1, x_19); +lean_ctor_set(x_27, 2, x_20); +lean_ctor_set(x_27, 3, x_21); +lean_ctor_set(x_27, 4, x_22); +lean_ctor_set(x_27, 5, x_23); +lean_ctor_set(x_27, 6, x_24); +lean_ctor_set(x_27, 7, x_25); +lean_ctor_set(x_27, 8, x_26); +x_28 = lean_st_ref_set(x_1, x_27, x_9); +x_29 = lean_ctor_get(x_28, 1); +lean_inc(x_29); +if (lean_is_exclusive(x_28)) { + lean_ctor_release(x_28, 0); + lean_ctor_release(x_28, 1); + x_30 = x_28; +} else { + lean_dec_ref(x_28); + x_30 = lean_box(0); +} +if (lean_is_scalar(x_30)) { + x_31 = lean_alloc_ctor(0, 2, 0); +} else { + x_31 = x_30; +} +lean_ctor_set(x_31, 0, x_6); +lean_ctor_set(x_31, 1, x_29); +return x_31; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Command_runLinters___spec__14(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Command_runLinters___spec__14___rarg___boxed), 2, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_withSeconds___at_Lean_Elab_Command_runLinters___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; +x_5 = lean_io_mono_nanos_now(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_apply_3(x_1, x_2, x_3, x_7); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +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_dec(x_8); +x_11 = lean_io_mono_nanos_now(x_10); +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; double x_17; double x_18; double x_19; lean_object* x_20; lean_object* x_21; +x_13 = lean_ctor_get(x_11, 0); +x_14 = lean_nat_sub(x_13, x_6); +lean_dec(x_6); +lean_dec(x_13); +x_15 = 0; +x_16 = lean_unsigned_to_nat(0u); +x_17 = l_Float_ofScientific(x_14, x_15, x_16); +lean_dec(x_14); +x_18 = l_Lean_withSeconds___at_Lean_Elab_Command_runLinters___spec__7___closed__1; +x_19 = lean_float_div(x_17, x_18); +x_20 = lean_box_float(x_19); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_9); +lean_ctor_set(x_21, 1, x_20); +lean_ctor_set(x_11, 0, x_21); +return x_11; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; lean_object* x_26; double x_27; double x_28; double x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_22 = lean_ctor_get(x_11, 0); +x_23 = lean_ctor_get(x_11, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_11); +x_24 = lean_nat_sub(x_22, x_6); +lean_dec(x_6); +lean_dec(x_22); +x_25 = 0; +x_26 = lean_unsigned_to_nat(0u); +x_27 = l_Float_ofScientific(x_24, x_25, x_26); +lean_dec(x_24); +x_28 = l_Lean_withSeconds___at_Lean_Elab_Command_runLinters___spec__7___closed__1; +x_29 = lean_float_div(x_27, x_28); +x_30 = lean_box_float(x_29); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_9); +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_23); +return x_32; +} +} +else +{ +uint8_t x_33; +lean_dec(x_6); +x_33 = !lean_is_exclusive(x_8); +if (x_33 == 0) +{ +return x_8; +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_8, 0); +x_35 = lean_ctor_get(x_8, 1); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_8); +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_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Elab_Command_runLinters___spec__17(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) { +_start: +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_9 = lean_st_ref_take(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_is_exclusive(x_10); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; size_t x_16; size_t 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; +x_13 = lean_ctor_get(x_10, 8); +x_14 = l_Lean_PersistentArray_toArray___rarg(x_13); +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___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___spec__1(x_16, x_17, x_14); +x_19 = lean_alloc_ctor(9, 3, 1); +lean_ctor_set(x_19, 0, x_2); +lean_ctor_set(x_19, 1, x_4); +lean_ctor_set(x_19, 2, x_18); +lean_ctor_set_uint8(x_19, sizeof(void*)*3, x_5); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_3); +lean_ctor_set(x_20, 1, x_19); +x_21 = l_Lean_PersistentArray_push___rarg(x_1, x_20); +lean_ctor_set(x_10, 8, x_21); +x_22 = lean_st_ref_set(x_7, x_10, x_11); +x_23 = !lean_is_exclusive(x_22); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; +x_24 = lean_ctor_get(x_22, 0); +lean_dec(x_24); +x_25 = lean_box(0); +lean_ctor_set(x_22, 0, x_25); +return x_22; +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_22, 1); +lean_inc(x_26); +lean_dec(x_22); +x_27 = lean_box(0); +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; +} +} +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; lean_object* x_37; lean_object* x_38; lean_object* x_39; size_t x_40; size_t x_41; 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_29 = lean_ctor_get(x_10, 0); +x_30 = lean_ctor_get(x_10, 1); +x_31 = lean_ctor_get(x_10, 2); +x_32 = lean_ctor_get(x_10, 3); +x_33 = lean_ctor_get(x_10, 4); +x_34 = lean_ctor_get(x_10, 5); +x_35 = lean_ctor_get(x_10, 6); +x_36 = lean_ctor_get(x_10, 7); +x_37 = lean_ctor_get(x_10, 8); +lean_inc(x_37); +lean_inc(x_36); +lean_inc(x_35); +lean_inc(x_34); +lean_inc(x_33); +lean_inc(x_32); +lean_inc(x_31); +lean_inc(x_30); +lean_inc(x_29); +lean_dec(x_10); +x_38 = l_Lean_PersistentArray_toArray___rarg(x_37); +x_39 = lean_array_get_size(x_38); +x_40 = lean_usize_of_nat(x_39); +lean_dec(x_39); +x_41 = 0; +x_42 = l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___spec__1(x_40, x_41, x_38); +x_43 = lean_alloc_ctor(9, 3, 1); +lean_ctor_set(x_43, 0, x_2); +lean_ctor_set(x_43, 1, x_4); +lean_ctor_set(x_43, 2, x_42); +lean_ctor_set_uint8(x_43, sizeof(void*)*3, x_5); +x_44 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_44, 0, x_3); +lean_ctor_set(x_44, 1, x_43); +x_45 = l_Lean_PersistentArray_push___rarg(x_1, x_44); +x_46 = lean_alloc_ctor(0, 9, 0); +lean_ctor_set(x_46, 0, x_29); +lean_ctor_set(x_46, 1, x_30); +lean_ctor_set(x_46, 2, x_31); +lean_ctor_set(x_46, 3, x_32); +lean_ctor_set(x_46, 4, x_33); +lean_ctor_set(x_46, 5, x_34); +lean_ctor_set(x_46, 6, x_35); +lean_ctor_set(x_46, 7, x_36); +lean_ctor_set(x_46, 8, x_45); +x_47 = lean_st_ref_set(x_7, x_46, x_11); +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___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Command_runLinters___spec__16(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) { +_start: +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_9 = l_Lean_Elab_Command_getRef(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_replaceRef(x_3, x_10); +lean_dec(x_10); +x_13 = !lean_is_exclusive(x_6); +if (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_6, 6); +lean_dec(x_14); +lean_ctor_set(x_6, 6, x_12); +x_15 = l_Lean_addMessageContextPartial___at_Lean_Elab_Command_instAddMessageContextCommandElabM___spec__1(x_4, x_6, x_7, x_11); +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___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Elab_Command_runLinters___spec__17(x_1, x_2, x_3, x_16, x_5, x_6, x_7, x_17); +lean_dec(x_6); +return x_18; +} +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_ctor_get(x_6, 0); +x_20 = lean_ctor_get(x_6, 1); +x_21 = lean_ctor_get(x_6, 2); +x_22 = lean_ctor_get(x_6, 3); +x_23 = lean_ctor_get(x_6, 4); +x_24 = lean_ctor_get(x_6, 5); +x_25 = lean_ctor_get(x_6, 7); +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_dec(x_6); +x_26 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_26, 0, x_19); +lean_ctor_set(x_26, 1, x_20); +lean_ctor_set(x_26, 2, x_21); +lean_ctor_set(x_26, 3, x_22); +lean_ctor_set(x_26, 4, x_23); +lean_ctor_set(x_26, 5, x_24); +lean_ctor_set(x_26, 6, x_12); +lean_ctor_set(x_26, 7, x_25); +x_27 = l_Lean_addMessageContextPartial___at_Lean_Elab_Command_instAddMessageContextCommandElabM___spec__1(x_4, x_26, x_7, x_11); +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_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Elab_Command_runLinters___spec__17(x_1, x_2, x_3, x_28, x_5, x_26, x_7, x_29); +lean_dec(x_26); +return x_30; +} +} +} +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Command_runLinters___spec__18(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_5; lean_object* x_6; +x_5 = lean_ctor_get(x_1, 0); +lean_inc(x_5); +x_6 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_6, 0, x_5); +lean_ctor_set(x_6, 1, x_4); +return x_6; +} +else +{ +lean_object* x_7; lean_object* x_8; +x_7 = lean_ctor_get(x_1, 0); +lean_inc(x_7); +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; +} +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__12___lambda__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) { +_start: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; +lean_inc(x_8); +x_11 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Command_runLinters___spec__16(x_1, x_2, x_3, x_6, x_4, x_8, x_9, x_10); +x_12 = lean_ctor_get(x_11, 1); +lean_inc(x_12); +lean_dec(x_11); +x_13 = l_MonadExcept_ofExcept___at_Lean_Elab_Command_runLinters___spec__18(x_5, x_8, x_9, x_12); +lean_dec(x_8); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__12___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, uint8_t x_7, double 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_dec(x_9); +x_13 = l_Lean_Elab_Command_getRef(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); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_2); +x_16 = lean_apply_4(x_1, x_2, x_10, x_11, x_15); +if (lean_obj_tag(x_16) == 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); +x_19 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3___closed__1; +x_20 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_6, x_19); +lean_dec(x_6); +if (x_20 == 0) +{ +if (x_7 == 0) +{ +lean_object* x_21; lean_object* x_22; +x_21 = lean_box(0); +x_22 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__12___lambda__1(x_3, x_4, x_14, x_5, x_2, x_17, x_21, x_10, x_11, x_18); +lean_dec(x_11); +lean_dec(x_2); +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; +x_23 = lean_float_to_string(x_8); +x_24 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_24, 0, x_23); +x_25 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_25, 0, x_24); +x_26 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3___closed__3; +x_27 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_25); +x_28 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3___closed__5; +x_29 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +x_30 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_17); +x_31 = l_Lean_Elab_logException___at_Lean_Elab_Command_runLinters___spec__1___closed__3; +x_32 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +x_33 = lean_box(0); +x_34 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__12___lambda__1(x_3, x_4, x_14, x_5, x_2, x_32, x_33, x_10, x_11, x_18); +lean_dec(x_11); +lean_dec(x_2); +return x_34; +} +} +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; +x_35 = lean_float_to_string(x_8); +x_36 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_36, 0, x_35); +x_37 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_37, 0, x_36); +x_38 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3___closed__3; +x_39 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_37); +x_40 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3___closed__5; +x_41 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_40); +x_42 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_17); +x_43 = l_Lean_Elab_logException___at_Lean_Elab_Command_runLinters___spec__1___closed__3; +x_44 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_44, 0, x_42); +lean_ctor_set(x_44, 1, x_43); +x_45 = lean_box(0); +x_46 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__12___lambda__1(x_3, x_4, x_14, x_5, x_2, x_44, x_45, x_10, x_11, x_18); +lean_dec(x_11); +lean_dec(x_2); +return x_46; +} +} +else +{ +uint8_t x_47; +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_47 = !lean_is_exclusive(x_16); +if (x_47 == 0) +{ +return x_16; +} +else +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_16, 0); +x_49 = lean_ctor_get(x_16, 1); +lean_inc(x_49); +lean_inc(x_48); +lean_dec(x_16); +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; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__12___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t 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) { +_start: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +lean_dec(x_7); +x_11 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Command_runLinters___spec__14___rarg(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_alloc_closure((void*)(l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__1), 4, 1); +lean_closure_set(x_14, 0, x_1); +lean_inc(x_9); +lean_inc(x_8); +x_15 = l_Lean_withSeconds___at_Lean_Elab_Command_runLinters___spec__15(x_14, x_8, x_9, x_13); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; lean_object* x_67; uint8_t x_68; +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); +x_19 = lean_ctor_get(x_16, 1); +lean_inc(x_19); +lean_dec(x_16); +x_67 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__4___closed__1; +x_68 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_5, x_67); +if (x_68 == 0) +{ +uint8_t x_69; +x_69 = 0; +x_20 = x_69; +goto block_66; +} +else +{ +double x_70; double x_71; uint8_t x_72; +x_70 = l_Lean_trace_profiler_threshold_getSecs(x_5); +x_71 = lean_unbox_float(x_19); +x_72 = lean_float_decLt(x_70, x_71); +x_20 = x_72; +goto block_66; +} +block_66: +{ +if (x_6 == 0) +{ +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +lean_dec(x_19); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +x_21 = lean_st_ref_take(x_9, x_17); +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_is_exclusive(x_22); +if (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_22, 8); +x_26 = l_Lean_PersistentArray_append___rarg(x_12, x_25); +lean_ctor_set(x_22, 8, x_26); +x_27 = lean_st_ref_set(x_9, x_22, x_23); +x_28 = lean_ctor_get(x_27, 1); +lean_inc(x_28); +lean_dec(x_27); +x_29 = l_MonadExcept_ofExcept___at_Lean_Elab_Command_runLinters___spec__18(x_18, x_8, x_9, x_28); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_18); +if (lean_obj_tag(x_29) == 0) +{ +uint8_t x_30; +x_30 = !lean_is_exclusive(x_29); +if (x_30 == 0) +{ +return x_29; +} +else +{ +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 +{ +uint8_t x_34; +x_34 = !lean_is_exclusive(x_29); +if (x_34 == 0) +{ +return x_29; +} +else +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_29, 0); +x_36 = lean_ctor_get(x_29, 1); +lean_inc(x_36); +lean_inc(x_35); +lean_dec(x_29); +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_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; 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_38 = lean_ctor_get(x_22, 0); +x_39 = lean_ctor_get(x_22, 1); +x_40 = lean_ctor_get(x_22, 2); +x_41 = lean_ctor_get(x_22, 3); +x_42 = lean_ctor_get(x_22, 4); +x_43 = lean_ctor_get(x_22, 5); +x_44 = lean_ctor_get(x_22, 6); +x_45 = lean_ctor_get(x_22, 7); +x_46 = lean_ctor_get(x_22, 8); +lean_inc(x_46); +lean_inc(x_45); +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_dec(x_22); +x_47 = l_Lean_PersistentArray_append___rarg(x_12, x_46); +x_48 = lean_alloc_ctor(0, 9, 0); +lean_ctor_set(x_48, 0, x_38); +lean_ctor_set(x_48, 1, x_39); +lean_ctor_set(x_48, 2, x_40); +lean_ctor_set(x_48, 3, x_41); +lean_ctor_set(x_48, 4, x_42); +lean_ctor_set(x_48, 5, x_43); +lean_ctor_set(x_48, 6, x_44); +lean_ctor_set(x_48, 7, x_45); +lean_ctor_set(x_48, 8, x_47); +x_49 = lean_st_ref_set(x_9, x_48, x_23); +x_50 = lean_ctor_get(x_49, 1); +lean_inc(x_50); +lean_dec(x_49); +x_51 = l_MonadExcept_ofExcept___at_Lean_Elab_Command_runLinters___spec__18(x_18, x_8, x_9, x_50); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_18); +if (lean_obj_tag(x_51) == 0) +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_52 = lean_ctor_get(x_51, 0); +lean_inc(x_52); +x_53 = lean_ctor_get(x_51, 1); +lean_inc(x_53); +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_51); + x_54 = lean_box(0); +} +if (lean_is_scalar(x_54)) { + x_55 = lean_alloc_ctor(0, 2, 0); +} else { + x_55 = x_54; +} +lean_ctor_set(x_55, 0, x_52); +lean_ctor_set(x_55, 1, x_53); +return x_55; +} +else +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_56 = lean_ctor_get(x_51, 0); +lean_inc(x_56); +x_57 = lean_ctor_get(x_51, 1); +lean_inc(x_57); +if (lean_is_exclusive(x_51)) { + lean_ctor_release(x_51, 0); + lean_ctor_release(x_51, 1); + x_58 = x_51; +} else { + lean_dec_ref(x_51); + x_58 = lean_box(0); } -x_99 = l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__4___closed__1; -if (lean_is_scalar(x_98)) { - x_100 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_58)) { + x_59 = lean_alloc_ctor(1, 2, 0); } else { - x_100 = x_98; + x_59 = x_58; +} +lean_ctor_set(x_59, 0, x_56); +lean_ctor_set(x_59, 1, x_57); +return x_59; } -lean_ctor_set(x_100, 0, x_99); -lean_ctor_set(x_100, 1, x_97); -x_11 = x_100; -goto block_22; } } 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_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; uint8_t x_114; -x_101 = lean_ctor_get(x_65, 0); -lean_inc(x_101); -x_102 = lean_ctor_get(x_65, 1); -lean_inc(x_102); -lean_dec(x_65); -x_103 = lean_st_ref_take(x_7, x_102); -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_ctor_get(x_24, 0); -lean_inc(x_106); -x_107 = lean_ctor_get(x_24, 2); -lean_inc(x_107); -x_108 = lean_ctor_get(x_24, 3); -lean_inc(x_108); -x_109 = lean_ctor_get(x_24, 4); -lean_inc(x_109); -x_110 = lean_ctor_get(x_24, 5); -lean_inc(x_110); -x_111 = lean_ctor_get(x_24, 6); -lean_inc(x_111); -x_112 = lean_ctor_get(x_24, 7); -lean_inc(x_112); -x_113 = lean_ctor_get(x_24, 8); -lean_inc(x_113); -lean_dec(x_24); -x_114 = !lean_is_exclusive(x_104); -if (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; lean_object* x_122; lean_object* x_123; uint8_t x_124; -x_115 = lean_ctor_get(x_104, 8); -lean_dec(x_115); -x_116 = lean_ctor_get(x_104, 7); -lean_dec(x_116); -x_117 = lean_ctor_get(x_104, 6); -lean_dec(x_117); -x_118 = lean_ctor_get(x_104, 5); -lean_dec(x_118); -x_119 = lean_ctor_get(x_104, 4); -lean_dec(x_119); -x_120 = lean_ctor_get(x_104, 3); -lean_dec(x_120); -x_121 = lean_ctor_get(x_104, 2); -lean_dec(x_121); -x_122 = lean_ctor_get(x_104, 0); -lean_dec(x_122); -lean_ctor_set(x_104, 8, x_113); -lean_ctor_set(x_104, 7, x_112); -lean_ctor_set(x_104, 6, x_111); -lean_ctor_set(x_104, 5, x_110); -lean_ctor_set(x_104, 4, x_109); -lean_ctor_set(x_104, 3, x_108); -lean_ctor_set(x_104, 2, x_107); -lean_ctor_set(x_104, 0, x_106); -x_123 = lean_st_ref_set(x_7, x_104, x_105); -x_124 = !lean_is_exclusive(x_123); -if (x_124 == 0) -{ -lean_object* x_125; -x_125 = lean_ctor_get(x_123, 0); -lean_dec(x_125); -lean_ctor_set_tag(x_123, 1); -lean_ctor_set(x_123, 0, x_101); -x_11 = x_123; -goto block_22; -} -else -{ -lean_object* x_126; lean_object* x_127; -x_126 = lean_ctor_get(x_123, 1); -lean_inc(x_126); -lean_dec(x_123); -x_127 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_127, 0, x_101); -lean_ctor_set(x_127, 1, x_126); -x_11 = x_127; -goto block_22; +lean_object* x_60; double x_61; lean_object* x_62; +x_60 = lean_box(0); +x_61 = lean_unbox_float(x_19); +lean_dec(x_19); +x_62 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__12___lambda__2(x_2, x_18, x_12, x_3, x_4, x_5, x_20, x_61, x_60, x_8, x_9, x_17); +return x_62; } } else { -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_128 = lean_ctor_get(x_104, 1); -lean_inc(x_128); -lean_dec(x_104); -x_129 = lean_alloc_ctor(0, 9, 0); -lean_ctor_set(x_129, 0, x_106); -lean_ctor_set(x_129, 1, x_128); -lean_ctor_set(x_129, 2, x_107); -lean_ctor_set(x_129, 3, x_108); -lean_ctor_set(x_129, 4, x_109); -lean_ctor_set(x_129, 5, x_110); -lean_ctor_set(x_129, 6, x_111); -lean_ctor_set(x_129, 7, x_112); -lean_ctor_set(x_129, 8, x_113); -x_130 = lean_st_ref_set(x_7, x_129, x_105); -x_131 = lean_ctor_get(x_130, 1); -lean_inc(x_131); -if (lean_is_exclusive(x_130)) { - lean_ctor_release(x_130, 0); - lean_ctor_release(x_130, 1); - x_132 = x_130; -} else { - lean_dec_ref(x_130); - x_132 = lean_box(0); -} -if (lean_is_scalar(x_132)) { - x_133 = lean_alloc_ctor(1, 2, 0); -} else { - x_133 = x_132; - lean_ctor_set_tag(x_133, 1); +lean_object* x_63; double x_64; lean_object* x_65; +x_63 = lean_box(0); +x_64 = lean_unbox_float(x_19); +lean_dec(x_19); +x_65 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__12___lambda__2(x_2, x_18, x_12, x_3, x_4, x_5, x_20, x_64, x_63, x_8, x_9, x_17); +return x_65; } -lean_ctor_set(x_133, 0, x_101); -lean_ctor_set(x_133, 1, x_131); -x_11 = x_133; -goto block_22; } } +else +{ +uint8_t x_73; +lean_dec(x_12); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +x_73 = !lean_is_exclusive(x_15); +if (x_73 == 0) +{ +return x_15; } -block_22: +else { -if (lean_obj_tag(x_11) == 0) +lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_74 = lean_ctor_get(x_15, 0); +x_75 = lean_ctor_get(x_15, 1); +lean_inc(x_75); +lean_inc(x_74); +lean_dec(x_15); +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_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__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, lean_object* x_7) { +_start: { -lean_object* x_12; lean_object* x_13; lean_object* x_14; size_t x_15; size_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_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_st_ref_get(x_6, 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_dec(x_8); +x_11 = lean_ctor_get(x_9, 2); +lean_inc(x_11); +lean_dec(x_9); +x_12 = l_Lean_Elab_Command_instInhabitedScope; +x_13 = l_List_head_x21___rarg(x_12, x_11); lean_dec(x_11); -x_14 = lean_ctor_get(x_12, 0); +x_14 = lean_ctor_get(x_13, 1); lean_inc(x_14); -lean_dec(x_12); -x_15 = 1; -x_16 = lean_usize_add(x_4, x_15); -x_4 = x_16; -x_5 = x_14; -x_8 = x_13; -goto _start; -} -else +lean_dec(x_13); +lean_inc(x_1); +x_15 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__13(x_1, x_5, x_6, x_10); +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_unbox(x_16); +if (x_17 == 0) { -uint8_t x_18; -lean_dec(x_7); -lean_dec(x_6); +lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_18 = lean_ctor_get(x_15, 1); +lean_inc(x_18); +lean_dec(x_15); +x_19 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__4___closed__1; +x_20 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_14, x_19); +if (x_20 == 0) +{ +lean_object* x_21; +lean_dec(x_16); +lean_dec(x_14); +lean_dec(x_2); lean_dec(x_1); -x_18 = !lean_is_exclusive(x_11); -if (x_18 == 0) +x_21 = lean_apply_3(x_3, x_5, x_6, x_18); +if (lean_obj_tag(x_21) == 0) { -return x_11; +uint8_t x_22; +x_22 = !lean_is_exclusive(x_21); +if (x_22 == 0) +{ +return x_21; } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_11, 0); -x_20 = lean_ctor_get(x_11, 1); -lean_inc(x_20); -lean_inc(x_19); -lean_dec(x_11); -x_21 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_21, 0, x_19); -lean_ctor_set(x_21, 1, x_20); +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; +} +} +else +{ +uint8_t x_26; +x_26 = !lean_is_exclusive(x_21); +if (x_26 == 0) +{ return x_21; } +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_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 +{ +lean_object* x_30; uint8_t x_31; lean_object* x_32; +x_30 = lean_box(0); +x_31 = lean_unbox(x_16); +lean_dec(x_16); +x_32 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__12___lambda__3(x_3, x_2, x_1, x_4, x_14, x_31, x_30, x_5, x_6, x_18); +return x_32; } } +else +{ +lean_object* x_33; lean_object* x_34; uint8_t x_35; lean_object* x_36; +x_33 = lean_ctor_get(x_15, 1); +lean_inc(x_33); +lean_dec(x_15); +x_34 = lean_box(0); +x_35 = lean_unbox(x_16); +lean_dec(x_16); +x_36 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__12___lambda__3(x_3, x_2, x_1, x_4, x_14, x_35, x_34, x_5, x_6, x_33); +return x_36; } } } -LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Elab_Command_runLinters___spec__5___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_EXPORT lean_object* l_Lean_profileitM___at_Lean_Elab_Command_runLinters___spec__19___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) { _start: { lean_object* x_8; lean_object* x_9; @@ -7919,175 +10985,205 @@ x_9 = l_Lean_profileitIOUnsafe___rarg(x_1, x_2, x_8, x_4, x_7); return x_9; } } -LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Elab_Command_runLinters___spec__5(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Elab_Command_runLinters___spec__19(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_profileitM___at_Lean_Elab_Command_runLinters___spec__5___rarg___boxed), 7, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_profileitM___at_Lean_Elab_Command_runLinters___spec__19___rarg___boxed), 7, 0); +return x_2; +} +} +static lean_object* _init_l_Lean_Elab_Command_runLinters___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("running linters", 15); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Command_runLinters___lambda__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_Command_runLinters___lambda__1___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } LEAN_EXPORT lean_object* l_Lean_Elab_Command_runLinters___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; uint8_t x_7; -x_5 = l_Lean_Elab_Command_addLinter___closed__1; -x_6 = lean_st_ref_get(x_5, x_4); -x_7 = !lean_is_exclusive(x_6); -if (x_7 == 0) +lean_object* x_5; lean_object* x_6; +x_5 = l_Lean_Elab_Command_runLinters___lambda__1___closed__2; +x_6 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_6, 0, x_5); +lean_ctor_set(x_6, 1, x_4); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Command_runLinters___lambda__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_8; lean_object* x_9; uint8_t x_10; -x_8 = lean_ctor_get(x_6, 0); -x_9 = lean_ctor_get(x_6, 1); -x_10 = l_Array_isEmpty___rarg(x_8); -if (x_10 == 0) +lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_6 = l_Lean_Elab_Command_addLinter___closed__1; +x_7 = lean_st_ref_get(x_6, x_5); +x_8 = !lean_is_exclusive(x_7); +if (x_8 == 0) { -lean_object* x_11; size_t x_12; size_t x_13; lean_object* x_14; lean_object* x_15; -lean_free_object(x_6); -x_11 = lean_array_get_size(x_8); -x_12 = lean_usize_of_nat(x_11); -lean_dec(x_11); -x_13 = 0; -x_14 = lean_box(0); -x_15 = l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__4(x_1, x_8, x_12, x_13, x_14, x_2, x_3, x_9); -lean_dec(x_8); -if (lean_obj_tag(x_15) == 0) +lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_9 = lean_ctor_get(x_7, 0); +x_10 = lean_ctor_get(x_7, 1); +x_11 = l_Array_isEmpty___rarg(x_9); +if (x_11 == 0) { -uint8_t x_16; -x_16 = !lean_is_exclusive(x_15); -if (x_16 == 0) +lean_object* x_12; size_t x_13; size_t x_14; lean_object* x_15; lean_object* x_16; +lean_free_object(x_7); +x_12 = lean_array_get_size(x_9); +x_13 = lean_usize_of_nat(x_12); +lean_dec(x_12); +x_14 = 0; +x_15 = lean_box(0); +x_16 = l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__11(x_1, x_2, x_9, x_13, x_14, x_15, x_3, x_4, x_10); +lean_dec(x_9); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_17; -x_17 = lean_ctor_get(x_15, 0); -lean_dec(x_17); -lean_ctor_set(x_15, 0, x_14); -return x_15; +uint8_t x_17; +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) +{ +lean_object* x_18; +x_18 = lean_ctor_get(x_16, 0); +lean_dec(x_18); +lean_ctor_set(x_16, 0, x_15); +return x_16; } else { -lean_object* x_18; lean_object* x_19; -x_18 = lean_ctor_get(x_15, 1); -lean_inc(x_18); -lean_dec(x_15); -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_14); -lean_ctor_set(x_19, 1, x_18); -return x_19; +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_16, 1); +lean_inc(x_19); +lean_dec(x_16); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_15); +lean_ctor_set(x_20, 1, x_19); +return x_20; } } else { -uint8_t x_20; -x_20 = !lean_is_exclusive(x_15); -if (x_20 == 0) +uint8_t x_21; +x_21 = !lean_is_exclusive(x_16); +if (x_21 == 0) { -return x_15; +return x_16; } 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_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_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_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; } } } else { -lean_object* x_24; -lean_dec(x_8); +lean_object* x_25; +lean_dec(x_9); +lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_24 = lean_box(0); -lean_ctor_set(x_6, 0, x_24); -return x_6; +lean_dec(x_2); +lean_dec(x_1); +x_25 = lean_box(0); +lean_ctor_set(x_7, 0, x_25); +return x_7; } } else { -lean_object* x_25; lean_object* x_26; uint8_t x_27; -x_25 = lean_ctor_get(x_6, 0); -x_26 = lean_ctor_get(x_6, 1); +lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_26 = lean_ctor_get(x_7, 0); +x_27 = lean_ctor_get(x_7, 1); +lean_inc(x_27); lean_inc(x_26); -lean_inc(x_25); -lean_dec(x_6); -x_27 = l_Array_isEmpty___rarg(x_25); -if (x_27 == 0) +lean_dec(x_7); +x_28 = l_Array_isEmpty___rarg(x_26); +if (x_28 == 0) { -lean_object* x_28; size_t x_29; size_t x_30; lean_object* x_31; lean_object* x_32; -x_28 = lean_array_get_size(x_25); -x_29 = lean_usize_of_nat(x_28); -lean_dec(x_28); -x_30 = 0; -x_31 = lean_box(0); -x_32 = l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__4(x_1, x_25, x_29, x_30, x_31, x_2, x_3, x_26); -lean_dec(x_25); -if (lean_obj_tag(x_32) == 0) +lean_object* x_29; size_t x_30; size_t x_31; lean_object* x_32; lean_object* x_33; +x_29 = lean_array_get_size(x_26); +x_30 = lean_usize_of_nat(x_29); +lean_dec(x_29); +x_31 = 0; +x_32 = lean_box(0); +x_33 = l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__11(x_1, x_2, x_26, x_30, x_31, x_32, x_3, x_4, x_27); +lean_dec(x_26); +if (lean_obj_tag(x_33) == 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); -if (lean_is_exclusive(x_32)) { - lean_ctor_release(x_32, 0); - lean_ctor_release(x_32, 1); - x_34 = x_32; +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_33, 1); +lean_inc(x_34); +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_32); - x_34 = lean_box(0); + lean_dec_ref(x_33); + x_35 = lean_box(0); } -if (lean_is_scalar(x_34)) { - x_35 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_35)) { + x_36 = lean_alloc_ctor(0, 2, 0); } else { - x_35 = x_34; + x_36 = x_35; } -lean_ctor_set(x_35, 0, x_31); -lean_ctor_set(x_35, 1, x_33); -return x_35; +lean_ctor_set(x_36, 0, x_32); +lean_ctor_set(x_36, 1, x_34); +return x_36; } 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); -lean_inc(x_36); -x_37 = lean_ctor_get(x_32, 1); +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_37 = lean_ctor_get(x_33, 0); lean_inc(x_37); -if (lean_is_exclusive(x_32)) { - lean_ctor_release(x_32, 0); - lean_ctor_release(x_32, 1); - x_38 = x_32; +x_38 = lean_ctor_get(x_33, 1); +lean_inc(x_38); +if (lean_is_exclusive(x_33)) { + lean_ctor_release(x_33, 0); + lean_ctor_release(x_33, 1); + x_39 = x_33; } else { - lean_dec_ref(x_32); - x_38 = lean_box(0); + lean_dec_ref(x_33); + x_39 = lean_box(0); } -if (lean_is_scalar(x_38)) { - x_39 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_39)) { + x_40 = lean_alloc_ctor(1, 2, 0); } else { - x_39 = x_38; + x_40 = x_39; } -lean_ctor_set(x_39, 0, x_36); -lean_ctor_set(x_39, 1, x_37); -return x_39; +lean_ctor_set(x_40, 0, x_37); +lean_ctor_set(x_40, 1, x_38); +return x_40; } } else { -lean_object* x_40; lean_object* x_41; -lean_dec(x_25); +lean_object* x_41; lean_object* x_42; +lean_dec(x_26); +lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -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_26); -return x_41; +x_41 = lean_box(0); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_27); +return x_42; } } } @@ -8096,6 +11192,14 @@ static lean_object* _init_l_Lean_Elab_Command_runLinters___closed__1() { _start: { lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Command_runLinters___lambda__1___boxed), 4, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Command_runLinters___closed__2() { +_start: +{ +lean_object* x_1; x_1 = lean_mk_string_from_bytes("linting", 7); return x_1; } @@ -8103,7 +11207,7 @@ return x_1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_runLinters(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_5; lean_object* x_6; lean_object* x_7; 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; x_5 = lean_st_ref_get(x_3, x_4); x_6 = lean_ctor_get(x_5, 0); lean_inc(x_6); @@ -8119,13 +11223,23 @@ lean_dec(x_8); x_11 = lean_ctor_get(x_10, 1); lean_inc(x_11); lean_dec(x_10); -x_12 = lean_alloc_closure((void*)(l_Lean_Elab_Command_runLinters___lambda__1), 4, 1); -lean_closure_set(x_12, 0, x_1); -x_13 = l_Lean_Elab_Command_runLinters___closed__1; -x_14 = lean_box(0); -x_15 = l_Lean_profileitM___at_Lean_Elab_Command_runLinters___spec__5___rarg(x_13, x_11, x_12, x_14, x_2, x_3, x_7); +x_12 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__3; +x_13 = lean_alloc_closure((void*)(l_Lean_Elab_Command_runLinters___lambda__2), 5, 2); +lean_closure_set(x_13, 0, x_1); +lean_closure_set(x_13, 1, x_12); +x_14 = l_Lean_Elab_Command_runLinters___closed__1; +x_15 = 1; +x_16 = lean_box(x_15); +x_17 = lean_alloc_closure((void*)(l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__12___boxed), 7, 4); +lean_closure_set(x_17, 0, x_12); +lean_closure_set(x_17, 1, x_14); +lean_closure_set(x_17, 2, x_13); +lean_closure_set(x_17, 3, x_16); +x_18 = l_Lean_Elab_Command_runLinters___closed__2; +x_19 = lean_box(0); +x_20 = l_Lean_profileitM___at_Lean_Elab_Command_runLinters___spec__19___rarg(x_18, x_11, x_17, x_19, x_2, x_3, x_7); lean_dec(x_11); -return x_15; +return x_20; } } LEAN_EXPORT lean_object* l_Lean_logAt___at_Lean_Elab_Command_runLinters___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) { @@ -8151,27 +11265,271 @@ x_7 = l_Lean_log___at_Lean_Elab_Command_runLinters___spec__3(x_1, x_6, x_3, x_4, return x_7; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___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_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__5___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_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__5(x_1, x_2, x_3, x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Command_runLinters___spec__6___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Command_runLinters___spec__6___rarg(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Command_runLinters___spec__6___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Command_runLinters___spec__6(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Elab_Command_runLinters___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) { +_start: +{ +uint8_t x_9; lean_object* x_10; +x_9 = lean_unbox(x_5); +lean_dec(x_5); +x_10 = l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Elab_Command_runLinters___spec__9(x_1, x_2, x_3, x_4, x_9, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_6); +return x_10; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Command_runLinters___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) { +_start: +{ +uint8_t x_9; lean_object* x_10; +x_9 = lean_unbox(x_5); +lean_dec(x_5); +x_10 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Command_runLinters___spec__8(x_1, x_2, x_3, x_4, x_9, x_6, x_7, x_8); +lean_dec(x_7); +return x_10; +} +} +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Command_runLinters___spec__10___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_MonadExcept_ofExcept___at_Lean_Elab_Command_runLinters___spec__10(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_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___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: +{ +uint8_t x_11; lean_object* x_12; +x_11 = lean_unbox(x_4); +lean_dec(x_4); +x_12 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__2(x_1, x_2, x_3, x_11, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_7); +lean_dec(x_5); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___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) { +_start: +{ +uint8_t x_13; uint8_t x_14; double x_15; lean_object* x_16; +x_13 = lean_unbox(x_5); +lean_dec(x_5); +x_14 = lean_unbox(x_7); +lean_dec(x_7); +x_15 = lean_unbox_float(x_8); +lean_dec(x_8); +x_16 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3(x_1, x_2, x_3, x_4, x_13, x_6, x_14, x_15, x_9, x_10, x_11, x_12); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___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: +{ +uint8_t x_11; uint8_t x_12; lean_object* x_13; +x_11 = lean_unbox(x_4); +lean_dec(x_4); +x_12 = lean_unbox(x_6); +lean_dec(x_6); +x_13 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__4(x_1, x_2, x_3, x_11, x_5, x_12, x_7, x_8, x_9, x_10); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___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) { +_start: +{ +uint8_t x_8; lean_object* x_9; +x_8 = lean_unbox(x_4); +lean_dec(x_4); +x_9 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4(x_1, x_2, x_3, x_8, x_5, x_6, x_7); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__11___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_Elab_Command_runLinters___spec__11___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_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___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_10; size_t x_11; lean_object* x_12; +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_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__11(x_1, x_2, x_3, x_10, x_11, x_6, x_7, x_8, x_9); +lean_dec(x_3); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___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_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__13(x_1, x_2, x_3, x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Command_runLinters___spec__14___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Command_runLinters___spec__14___rarg(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Command_runLinters___spec__14___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Command_runLinters___spec__14(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Elab_Command_runLinters___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) { +_start: +{ +uint8_t x_9; lean_object* x_10; +x_9 = lean_unbox(x_5); +lean_dec(x_5); +x_10 = l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Elab_Command_runLinters___spec__17(x_1, x_2, x_3, x_4, x_9, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_6); +return x_10; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Command_runLinters___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) { +_start: +{ +uint8_t x_9; lean_object* x_10; +x_9 = lean_unbox(x_5); +lean_dec(x_5); +x_10 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Command_runLinters___spec__16(x_1, x_2, x_3, x_4, x_9, x_6, x_7, x_8); +lean_dec(x_7); +return x_10; +} +} +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Command_runLinters___spec__18___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_MonadExcept_ofExcept___at_Lean_Elab_Command_runLinters___spec__18(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_withTraceNode___at_Lean_Elab_Command_runLinters___spec__12___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: +{ +uint8_t x_11; lean_object* x_12; +x_11 = lean_unbox(x_4); +lean_dec(x_4); +x_12 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__12___lambda__1(x_1, x_2, x_3, x_11, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_7); +lean_dec(x_5); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__12___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: +{ +uint8_t x_13; uint8_t x_14; double x_15; lean_object* x_16; +x_13 = lean_unbox(x_5); +lean_dec(x_5); +x_14 = lean_unbox(x_7); +lean_dec(x_7); +x_15 = lean_unbox_float(x_8); +lean_dec(x_8); +x_16 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__12___lambda__2(x_1, x_2, x_3, x_4, x_13, x_6, x_14, x_15, x_9, x_10, x_11, x_12); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__12___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; uint8_t x_12; lean_object* x_13; +x_11 = lean_unbox(x_4); +lean_dec(x_4); +x_12 = lean_unbox(x_6); +lean_dec(x_6); +x_13 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__12___lambda__3(x_1, x_2, x_3, x_11, x_5, x_12, x_7, x_8, x_9, x_10); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___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) { +_start: +{ +uint8_t x_8; lean_object* x_9; +x_8 = lean_unbox(x_4); +lean_dec(x_4); +x_9 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__12(x_1, x_2, x_3, x_8, x_5, x_6, x_7); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Elab_Command_runLinters___spec__19___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { -size_t x_9; size_t x_10; lean_object* x_11; -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_Command_runLinters___spec__4(x_1, x_2, x_9, x_10, x_5, x_6, x_7, x_8); +lean_object* x_8; +x_8 = l_Lean_profileitM___at_Lean_Elab_Command_runLinters___spec__19___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7); lean_dec(x_2); -return x_11; +lean_dec(x_1); +return x_8; } } -LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Elab_Command_runLinters___spec__5___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_Elab_Command_runLinters___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_8; -x_8 = l_Lean_profileitM___at_Lean_Elab_Command_runLinters___spec__5___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_object* x_5; +x_5 = l_Lean_Elab_Command_runLinters___lambda__1(x_1, x_2, x_3, x_4); +lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -return x_8; +return x_5; } } LEAN_EXPORT lean_object* l_Lean_Elab_Command_getCurrMacroScope(lean_object* x_1, lean_object* x_2, lean_object* x_3) { @@ -8508,47 +11866,15 @@ return x_3; static lean_object* _init_l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__5() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Lean", 4); -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__6() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Parser", 6); -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__7() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Command", 7); -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__8() { -_start: -{ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__5; -x_2 = l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__6; -x_3 = l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__7; +x_1 = l___auto____x40_Lean_Elab_Command___hyg_448____closed__1; +x_2 = l___auto____x40_Lean_Elab_Command___hyg_448____closed__2; +x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__6; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__9() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Elab", 4); -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__10() { +static lean_object* _init_l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__6() { _start: { lean_object* x_1; @@ -8556,19 +11882,19 @@ x_1 = lean_mk_string_from_bytes("CommandElab", 11); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__11() { +static lean_object* _init_l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__7() { _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_Elab_Command_mkCommandElabAttributeUnsafe___closed__5; -x_2 = l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__9; -x_3 = l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__7; -x_4 = l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__10; +x_1 = l___auto____x40_Lean_Elab_Command___hyg_448____closed__1; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__1; +x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__6; +x_4 = l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__6; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__12() { +static lean_object* _init_l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__8() { _start: { lean_object* x_1; @@ -8582,14 +11908,14 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkCommandElabAttributeUnsafe(lean_o 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_3 = l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__2; x_4 = l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__4; -x_5 = l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__8; -x_6 = l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__11; -x_7 = l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__12; +x_5 = l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__5; +x_6 = l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__7; +x_7 = l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__8; x_8 = l_Lean_Elab_mkElabAttribute___rarg(x_3, x_4, x_5, x_6, x_7, x_1, x_2); return x_8; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3670____closed__1() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3931____closed__1() { _start: { lean_object* x_1; @@ -8597,23 +11923,23 @@ x_1 = lean_mk_string_from_bytes("commandElabAttribute", 20); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3670____closed__2() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3931____closed__2() { _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_Elab_Command_mkCommandElabAttributeUnsafe___closed__5; -x_2 = l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__9; -x_3 = l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__7; -x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3670____closed__1; +x_1 = l___auto____x40_Lean_Elab_Command___hyg_448____closed__1; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__1; +x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__6; +x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3931____closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3670_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3931_(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_3670____closed__2; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3931____closed__2; x_3 = l_Lean_Elab_Command_mkCommandElabAttributeUnsafe(x_2, x_1); return x_3; } @@ -10590,7 +13916,7 @@ lean_dec(x_1); return x_4; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4380____closed__1() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4641____closed__1() { _start: { lean_object* x_1; @@ -10598,17 +13924,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_4380____closed__2() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4641____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_4380____closed__1; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4641____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4380____closed__3() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4641____closed__3() { _start: { lean_object* x_1; @@ -10616,13 +13942,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_4380____closed__4() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4641____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_4380____closed__3; +x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4641____closed__3; x_4 = lean_box(x_1); x_5 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_5, 0, x_4); @@ -10631,170 +13957,56 @@ lean_ctor_set(x_5, 2, x_3); return x_5; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4380____closed__5() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4641____closed__5() { _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_Elab_Command_mkCommandElabAttributeUnsafe___closed__5; -x_2 = l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__9; -x_3 = l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__7; -x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4380____closed__1; +x_1 = l___auto____x40_Lean_Elab_Command___hyg_448____closed__1; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__1; +x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__6; +x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4641____closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4380_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4641_(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_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4380____closed__2; -x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4380____closed__4; -x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4380____closed__5; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4641____closed__2; +x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4641____closed__4; +x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4641____closed__5; x_5 = l_Lean_Option_register___at_Lean_Elab_initFn____x40_Lean_Elab_AutoBound___hyg_7____spec__1(x_2, x_3, x_4, x_1); return x_5; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__1() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4678____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__9; -x_2 = l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__12; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__1; +x_2 = l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__8; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____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_mkCommandElabAttributeUnsafe___closed__5; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__3() { -_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_4417____closed__2; -x_2 = l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__9; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____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_Command___hyg_4417____closed__3; -x_2 = l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__7; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__5() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("initFn", 6); -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__6() { -_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_4417____closed__4; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__5; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__7() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("_@", 2); -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__8() { -_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_4417____closed__6; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__7; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__9() { -_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_4417____closed__8; -x_2 = l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__5; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__10() { -_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_4417____closed__9; -x_2 = l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__9; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__11() { -_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_4417____closed__10; -x_2 = l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__7; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__12() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("_hyg", 4); -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__13() { -_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_4417____closed__11; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__12; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__14() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4678____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_4417____closed__13; -x_2 = lean_unsigned_to_nat(4417u); +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__16; +x_2 = lean_unsigned_to_nat(4678u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4678_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__1; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4678____closed__1; x_3 = 0; -x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__14; +x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4678____closed__2; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } @@ -10902,7 +14114,7 @@ else { lean_object* x_349; uint8_t x_350; x_349 = l_Lean_logAt___at_Lean_Elab_Command_runLinters___spec__2___closed__1; -x_350 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_14, x_349); +x_350 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_14, x_349); lean_dec(x_14); if (x_350 == 0) { @@ -12156,76 +15368,7 @@ return x_8; } } } -static lean_object* _init_l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_elabCommand___spec__7___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_inheritedTraceOptions; -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_elabCommand___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; lean_object* x_8; lean_object* x_9; uint8_t x_10; -x_5 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_elabCommand___spec__7___closed__1; -x_6 = lean_st_ref_get(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_9 = lean_st_ref_get(x_3, x_8); -x_10 = !lean_is_exclusive(x_9); -if (x_10 == 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; lean_object* x_17; -x_11 = lean_ctor_get(x_9, 0); -x_12 = lean_ctor_get(x_11, 2); -lean_inc(x_12); -lean_dec(x_11); -x_13 = l_Lean_Elab_Command_instInhabitedScope; -x_14 = l_List_head_x21___rarg(x_13, x_12); -lean_dec(x_12); -x_15 = lean_ctor_get(x_14, 1); -lean_inc(x_15); -lean_dec(x_14); -x_16 = l___private_Lean_Util_Trace_0__Lean_checkTraceOption(x_7, x_15, x_1); -lean_dec(x_15); -lean_dec(x_7); -x_17 = lean_box(x_16); -lean_ctor_set(x_9, 0, x_17); -return x_9; -} -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_25; lean_object* x_26; -x_18 = lean_ctor_get(x_9, 0); -x_19 = lean_ctor_get(x_9, 1); -lean_inc(x_19); -lean_inc(x_18); -lean_dec(x_9); -x_20 = lean_ctor_get(x_18, 2); -lean_inc(x_20); -lean_dec(x_18); -x_21 = l_Lean_Elab_Command_instInhabitedScope; -x_22 = l_List_head_x21___rarg(x_21, x_20); -lean_dec(x_20); -x_23 = lean_ctor_get(x_22, 1); -lean_inc(x_23); -lean_dec(x_22); -x_24 = l___private_Lean_Util_Trace_0__Lean_checkTraceOption(x_7, x_23, x_1); -lean_dec(x_23); -lean_dec(x_7); -x_25 = lean_box(x_24); -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_19); -return x_26; -} -} -} -LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Elab_Command_elabCommand___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_addTrace___at_Lean_Elab_Command_elabCommand___spec__7(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; uint8_t x_15; @@ -12354,7 +15497,7 @@ return x_48; } } } -LEAN_EXPORT lean_object* l_List_forM___at_Lean_Elab_Command_elabCommand___spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_List_forM___at_Lean_Elab_Command_elabCommand___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) == 0) @@ -12380,7 +15523,7 @@ x_10 = lean_ctor_get(x_7, 1); lean_inc(x_10); lean_dec(x_7); lean_inc(x_9); -x_11 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_elabCommand___spec__7(x_9, x_2, x_3, x_4); +x_11 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__13(x_9, x_2, x_3, x_4); x_12 = lean_ctor_get(x_11, 0); lean_inc(x_12); x_13 = lean_unbox(x_12); @@ -12407,7 +15550,7 @@ x_17 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_17, 0, x_10); x_18 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_18, 0, x_17); -x_19 = l_Lean_addTrace___at_Lean_Elab_Command_elabCommand___spec__8(x_9, x_18, x_2, x_3, x_16); +x_19 = l_Lean_addTrace___at_Lean_Elab_Command_elabCommand___spec__7(x_9, x_18, x_2, x_3, x_16); x_20 = lean_ctor_get(x_19, 1); lean_inc(x_20); lean_dec(x_19); @@ -12418,7 +15561,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabCommand___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_Command_elabCommand___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; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; @@ -12471,7 +15614,7 @@ return x_20; } } } -LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabCommand___spec__10(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_elabCommand___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; @@ -12491,7 +15634,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_elabCommand___spec__11(x_2, x_3, x_4, x_8); +x_12 = l_Lean_throwError___at_Lean_Elab_Command_elabCommand___spec__10(x_2, x_3, x_4, x_8); lean_dec(x_4); return x_12; } @@ -12522,13 +15665,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_elabCommand___spec__11(x_2, x_20, x_4, x_8); +x_21 = l_Lean_throwError___at_Lean_Elab_Command_elabCommand___spec__10(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_elabCommand___spec__12(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_elabCommand___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; @@ -12542,7 +15685,7 @@ lean_ctor_set(x_7, 1, x_4); return x_7; } } -static lean_object* _init_l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__13___rarg___closed__1() { +static lean_object* _init_l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__12___rarg___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -12554,22 +15697,22 @@ lean_ctor_set(x_3, 1, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__13___rarg(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___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_Command_elabCommand___spec__13___rarg___closed__1; +x_2 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__12___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_elabCommand___spec__13(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___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_elabCommand___spec__13___rarg), 1, 0); +x_3 = lean_alloc_closure((void*)(l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__12___rarg), 1, 0); return x_3; } } @@ -12910,7 +16053,7 @@ x_53 = lean_ctor_get(x_44, 1); lean_inc(x_53); lean_dec(x_44); x_54 = l_List_reverse___rarg(x_53); -x_55 = l_List_forM___at_Lean_Elab_Command_elabCommand___spec__9(x_54, x_2, x_3, x_52); +x_55 = l_List_forM___at_Lean_Elab_Command_elabCommand___spec__8(x_54, x_2, x_3, x_52); lean_dec(x_3); lean_dec(x_2); x_56 = !lean_is_exclusive(x_55); @@ -12972,7 +16115,7 @@ x_71 = lean_ctor_get(x_44, 1); lean_inc(x_71); lean_dec(x_44); x_72 = l_List_reverse___rarg(x_71); -x_73 = l_List_forM___at_Lean_Elab_Command_elabCommand___spec__9(x_72, x_2, x_3, x_70); +x_73 = l_List_forM___at_Lean_Elab_Command_elabCommand___spec__8(x_72, x_2, x_3, x_70); lean_dec(x_3); lean_dec(x_2); x_74 = lean_ctor_get(x_73, 1); @@ -13018,14 +16161,14 @@ x_82 = lean_alloc_ctor(3, 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_elabCommand___spec__10(x_78, x_83, x_2, x_3, x_36); +x_84 = l_Lean_throwErrorAt___at_Lean_Elab_Command_elabCommand___spec__9(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_elabCommand___spec__12(x_78, x_2, x_3, x_36); +x_85 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabCommand___spec__11(x_78, x_2, x_3, x_36); lean_dec(x_3); lean_dec(x_2); return x_85; @@ -13036,13 +16179,13 @@ else lean_object* x_86; lean_dec(x_3); lean_dec(x_2); -x_86 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__13___rarg(x_36); +x_86 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__12___rarg(x_36); return x_86; } } } } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabCommand___spec__16(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_elabCommand___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; @@ -13095,7 +16238,7 @@ return x_20; } } } -LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabCommand___spec__15(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_elabCommand___spec__14(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; @@ -13115,7 +16258,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_elabCommand___spec__16(x_2, x_3, x_4, x_8); +x_12 = l_Lean_throwError___at_Lean_Elab_Command_elabCommand___spec__15(x_2, x_3, x_4, x_8); lean_dec(x_4); return x_12; } @@ -13146,13 +16289,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_elabCommand___spec__16(x_2, x_20, x_4, x_8); +x_21 = l_Lean_throwError___at_Lean_Elab_Command_elabCommand___spec__15(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_elabCommand___spec__17(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_elabCommand___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; @@ -13166,26 +16309,26 @@ lean_ctor_set(x_7, 1, x_4); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__18___rarg(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__17___rarg(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__13___rarg___closed__1; +x_2 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__12___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_elabCommand___spec__18(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__17(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_elabCommand___spec__18___rarg), 1, 0); +x_3 = lean_alloc_closure((void*)(l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__17___rarg), 1, 0); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabCommand___spec__14(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_elabCommand___spec__13(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; @@ -13319,7 +16462,7 @@ x_53 = lean_ctor_get(x_44, 1); lean_inc(x_53); lean_dec(x_44); x_54 = l_List_reverse___rarg(x_53); -x_55 = l_List_forM___at_Lean_Elab_Command_elabCommand___spec__9(x_54, x_2, x_3, x_52); +x_55 = l_List_forM___at_Lean_Elab_Command_elabCommand___spec__8(x_54, x_2, x_3, x_52); lean_dec(x_3); lean_dec(x_2); x_56 = !lean_is_exclusive(x_55); @@ -13381,7 +16524,7 @@ x_71 = lean_ctor_get(x_44, 1); lean_inc(x_71); lean_dec(x_44); x_72 = l_List_reverse___rarg(x_71); -x_73 = l_List_forM___at_Lean_Elab_Command_elabCommand___spec__9(x_72, x_2, x_3, x_70); +x_73 = l_List_forM___at_Lean_Elab_Command_elabCommand___spec__8(x_72, x_2, x_3, x_70); lean_dec(x_3); lean_dec(x_2); x_74 = lean_ctor_get(x_73, 1); @@ -13427,14 +16570,14 @@ x_82 = lean_alloc_ctor(3, 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_elabCommand___spec__15(x_78, x_83, x_2, x_3, x_36); +x_84 = l_Lean_throwErrorAt___at_Lean_Elab_Command_elabCommand___spec__14(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_elabCommand___spec__17(x_78, x_2, x_3, x_36); +x_85 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabCommand___spec__16(x_78, x_2, x_3, x_36); lean_dec(x_3); lean_dec(x_2); return x_85; @@ -13445,13 +16588,13 @@ else lean_object* x_86; lean_dec(x_3); lean_dec(x_2); -x_86 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__18___rarg(x_36); +x_86 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__17___rarg(x_36); return x_86; } } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__19(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_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__18(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; @@ -13699,10 +16842,21 @@ return x_52; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCommand___lambda__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; +x_6 = l_Lean_MessageData_ofSyntax(x_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); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCommand___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_inc(x_4); lean_inc(x_3); -x_6 = l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabCommand___spec__14(x_1, x_3, x_4, x_5); +x_6 = l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabCommand___spec__13(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; @@ -13744,7 +16898,7 @@ return x_14; } } } -static lean_object* _init_l_Lean_Elab_Command_elabCommand___lambda__3___closed__1() { +static lean_object* _init_l_Lean_Elab_Command_elabCommand___lambda__4___closed__1() { _start: { lean_object* x_1; @@ -13752,7 +16906,7 @@ x_1 = l_Lean_Elab_Command_commandElabAttribute; return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_elabCommand___lambda__3___closed__2() { +static lean_object* _init_l_Lean_Elab_Command_elabCommand___lambda__4___closed__2() { _start: { lean_object* x_1; @@ -13760,16 +16914,16 @@ x_1 = lean_mk_string_from_bytes("elaboration function for '", 26); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_elabCommand___lambda__3___closed__3() { +static lean_object* _init_l_Lean_Elab_Command_elabCommand___lambda__4___closed__3() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Command_elabCommand___lambda__3___closed__2; +x_1 = l_Lean_Elab_Command_elabCommand___lambda__4___closed__2; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Command_elabCommand___lambda__3___closed__4() { +static lean_object* _init_l_Lean_Elab_Command_elabCommand___lambda__4___closed__4() { _start: { lean_object* x_1; @@ -13777,175 +16931,140 @@ x_1 = lean_mk_string_from_bytes("' has not been implemented", 26); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_elabCommand___lambda__3___closed__5() { +static lean_object* _init_l_Lean_Elab_Command_elabCommand___lambda__4___closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Command_elabCommand___lambda__3___closed__4; +x_1 = l_Lean_Elab_Command_elabCommand___lambda__4___closed__4; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCommand___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_EXPORT lean_object* l_Lean_Elab_Command_elabCommand___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_7; lean_object* x_41; lean_object* x_42; uint8_t x_43; -lean_inc(x_3); -x_41 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_elabCommand___spec__7(x_3, x_4, x_5, x_6); -x_42 = lean_ctor_get(x_41, 0); -lean_inc(x_42); -x_43 = lean_unbox(x_42); -lean_dec(x_42); -if (x_43 == 0) -{ -lean_object* x_44; -lean_dec(x_3); -x_44 = lean_ctor_get(x_41, 1); -lean_inc(x_44); -lean_dec(x_41); -x_7 = x_44; -goto block_40; -} -else -{ -lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_45 = lean_ctor_get(x_41, 1); -lean_inc(x_45); -lean_dec(x_41); -lean_inc(x_1); -x_46 = l_Lean_MessageData_ofSyntax(x_1); -x_47 = l_Lean_addTrace___at_Lean_Elab_Command_elabCommand___spec__8(x_3, x_46, x_4, x_5, x_45); -x_48 = lean_ctor_get(x_47, 1); -lean_inc(x_48); -lean_dec(x_47); -x_7 = x_48; -goto block_40; -} -block_40: -{ -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_st_ref_get(x_5, x_7); -x_9 = lean_ctor_get(x_8, 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; +x_6 = lean_st_ref_get(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 = lean_ctor_get(x_7, 0); lean_inc(x_9); -x_10 = lean_ctor_get(x_8, 1); -lean_inc(x_10); -lean_dec(x_8); -x_11 = lean_ctor_get(x_9, 0); -lean_inc(x_11); lean_inc(x_1); -lean_inc(x_11); -x_12 = lean_alloc_closure((void*)(l_Lean_Elab_expandMacroImpl_x3f___boxed), 4, 2); -lean_closure_set(x_12, 0, x_11); -lean_closure_set(x_12, 1, x_1); -lean_inc(x_5); +lean_inc(x_9); +x_10 = lean_alloc_closure((void*)(l_Lean_Elab_expandMacroImpl_x3f___boxed), 4, 2); +lean_closure_set(x_10, 0, x_9); +lean_closure_set(x_10, 1, x_1); lean_inc(x_4); -x_13 = l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabCommand___spec__6(x_12, x_4, x_5, x_10); -if (lean_obj_tag(x_13) == 0) +lean_inc(x_3); +x_11 = l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabCommand___spec__6(x_10, x_3, x_4, x_8); +if (lean_obj_tag(x_11) == 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_12; +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +if (lean_obj_tag(x_12) == 0) { -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_13, 1); -lean_inc(x_15); -lean_dec(x_13); -x_16 = l_Lean_Elab_Command_elabCommand___lambda__3___closed__1; +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +lean_dec(x_11); +x_14 = l_Lean_Elab_Command_elabCommand___lambda__4___closed__1; lean_inc(x_2); -x_17 = l_Lean_KeyedDeclsAttribute_getEntries___rarg(x_16, x_11, x_2); -if (lean_obj_tag(x_17) == 0) +x_15 = l_Lean_KeyedDeclsAttribute_getEntries___rarg(x_14, x_9, x_2); +if (lean_obj_tag(x_15) == 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_dec(x_9); -x_18 = l_Lean_MessageData_ofName(x_2); -x_19 = l_Lean_Elab_Command_elabCommand___lambda__3___closed__3; +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_dec(x_7); +x_16 = l_Lean_MessageData_ofName(x_2); +x_17 = l_Lean_Elab_Command_elabCommand___lambda__4___closed__3; +x_18 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_16); +x_19 = l_Lean_Elab_Command_elabCommand___lambda__4___closed__5; x_20 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_18); -x_21 = l_Lean_Elab_Command_elabCommand___lambda__3___closed__5; -x_22 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_22, 0, x_20); -lean_ctor_set(x_22, 1, x_21); -x_23 = lean_alloc_closure((void*)(l_Lean_throwError___at___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___spec__1___boxed), 4, 1); +lean_ctor_set(x_20, 0, x_18); +lean_ctor_set(x_20, 1, x_19); +x_21 = lean_alloc_closure((void*)(l_Lean_throwError___at___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___spec__1___boxed), 4, 1); +lean_closure_set(x_21, 0, x_20); +x_22 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___closed__4; +x_23 = lean_alloc_closure((void*)(l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkInfoTree___boxed), 6, 2); lean_closure_set(x_23, 0, x_22); -x_24 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___closed__4; -x_25 = lean_alloc_closure((void*)(l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkInfoTree___boxed), 6, 2); -lean_closure_set(x_25, 0, x_24); -lean_closure_set(x_25, 1, x_1); -x_26 = l_Lean_Elab_withInfoTreeContext___at___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___spec__2(x_23, x_25, x_4, x_5, x_15); -return x_26; +lean_closure_set(x_23, 1, x_1); +x_24 = l_Lean_Elab_withInfoTreeContext___at___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___spec__2(x_21, x_23, x_3, x_4, x_13); +return x_24; } else { -lean_object* x_27; +lean_object* x_25; lean_dec(x_2); -x_27 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing(x_9, x_1, x_17, x_4, x_5, x_15); -return x_27; +x_25 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing(x_7, x_1, x_15, x_3, x_4, x_13); +return x_25; } } 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_dec(x_11); +lean_object* x_26; lean_object* x_27; lean_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_9); +lean_dec(x_7); lean_dec(x_2); -x_28 = lean_ctor_get(x_14, 0); +x_26 = lean_ctor_get(x_12, 0); +lean_inc(x_26); +lean_dec(x_12); +x_27 = lean_ctor_get(x_11, 1); +lean_inc(x_27); +lean_dec(x_11); +x_28 = lean_ctor_get(x_26, 0); lean_inc(x_28); -lean_dec(x_14); -x_29 = lean_ctor_get(x_13, 1); +x_29 = lean_ctor_get(x_26, 1); lean_inc(x_29); -lean_dec(x_13); -x_30 = lean_ctor_get(x_28, 0); -lean_inc(x_30); -x_31 = lean_ctor_get(x_28, 1); -lean_inc(x_31); -lean_dec(x_28); -x_32 = lean_alloc_closure((void*)(l_liftExcept___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__8___boxed), 3, 1); -lean_closure_set(x_32, 0, x_31); +lean_dec(x_26); +x_30 = lean_alloc_closure((void*)(l_liftExcept___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__8___boxed), 3, 1); +lean_closure_set(x_30, 0, x_29); lean_inc(x_1); -x_33 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand___lambda__2), 5, 2); -lean_closure_set(x_33, 0, x_32); -lean_closure_set(x_33, 1, x_1); -x_34 = lean_alloc_closure((void*)(l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkInfoTree___boxed), 6, 2); -lean_closure_set(x_34, 0, x_30); -lean_closure_set(x_34, 1, x_1); -x_35 = l_Lean_Elab_withInfoTreeContext___at___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___spec__2(x_33, x_34, x_4, x_5, x_29); -return x_35; +x_31 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand___lambda__3), 5, 2); +lean_closure_set(x_31, 0, x_30); +lean_closure_set(x_31, 1, x_1); +x_32 = lean_alloc_closure((void*)(l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkInfoTree___boxed), 6, 2); +lean_closure_set(x_32, 0, x_28); +lean_closure_set(x_32, 1, x_1); +x_33 = l_Lean_Elab_withInfoTreeContext___at___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___spec__2(x_31, x_32, x_3, x_4, x_27); +return x_33; } } else { -uint8_t x_36; -lean_dec(x_11); +uint8_t x_34; lean_dec(x_9); -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_36 = !lean_is_exclusive(x_13); -if (x_36 == 0) +x_34 = !lean_is_exclusive(x_11); +if (x_34 == 0) { -return x_13; +return x_11; } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_13, 0); -x_38 = lean_ctor_get(x_13, 1); -lean_inc(x_38); -lean_inc(x_37); -lean_dec(x_13); -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_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_11, 0); +x_36 = lean_ctor_get(x_11, 1); +lean_inc(x_36); +lean_inc(x_35); +lean_dec(x_11); +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; } } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCommand___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCommand___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; @@ -13987,25 +17106,7 @@ 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 = l_Lean_Name_str___override(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); +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand___lambda__5___boxed), 3, 0); return x_1; } } @@ -14028,53 +17129,46 @@ 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_Elab_Command_elabCommand___closed__5; +x_7 = l___auto____x40_Lean_Elab_Command___hyg_448____closed__9; 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_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_dec(x_6); -x_9 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__1; lean_inc(x_1); -x_10 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand___lambda__3), 6, 3); +x_9 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand___lambda__2___boxed), 5, 1); +lean_closure_set(x_9, 0, x_1); +lean_inc(x_1); +x_10 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand___lambda__4), 5, 2); lean_closure_set(x_10, 0, x_1); lean_closure_set(x_10, 1, x_5); -lean_closure_set(x_10, 2, x_9); -x_11 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand___lambda__1), 5, 2); -lean_closure_set(x_11, 0, x_1); -lean_closure_set(x_11, 1, x_10); -x_12 = l_Lean_Elab_withLogging___at_Lean_Elab_Command_elabCommand___spec__2(x_11, x_2, x_3, x_4); -return x_12; +x_11 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4678____closed__1; +x_12 = 1; +x_13 = lean_box(x_12); +x_14 = lean_alloc_closure((void*)(l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__12___boxed), 7, 4); +lean_closure_set(x_14, 0, x_11); +lean_closure_set(x_14, 1, x_9); +lean_closure_set(x_14, 2, x_10); +lean_closure_set(x_14, 3, x_13); +x_15 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand___lambda__1), 5, 2); +lean_closure_set(x_15, 0, x_1); +lean_closure_set(x_15, 1, x_14); +x_16 = l_Lean_Elab_withLogging___at_Lean_Elab_Command_elabCommand___spec__2(x_15, x_2, x_3, x_4); +return x_16; } else { -lean_object* x_13; lean_object* x_14; uint8_t x_15; +lean_object* x_17; lean_object* x_18; uint8_t x_19; lean_dec(x_5); -x_13 = lean_array_get_size(x_6); -x_14 = lean_unsigned_to_nat(0u); -x_15 = lean_nat_dec_lt(x_14, x_13); -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__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); -x_18 = l_Lean_Elab_withLogging___at_Lean_Elab_Command_elabCommand___spec__2(x_17, x_2, x_3, x_4); -return x_18; -} -else -{ -uint8_t x_19; -x_19 = lean_nat_dec_le(x_13, x_13); +x_17 = lean_array_get_size(x_6); +x_18 = lean_unsigned_to_nat(0u); +x_19 = lean_nat_dec_lt(x_18, x_17); if (x_19 == 0) { lean_object* x_20; lean_object* x_21; lean_object* x_22; -lean_dec(x_13); +lean_dec(x_17); lean_dec(x_6); -x_20 = l_Lean_Elab_Command_elabCommand___closed__6; +x_20 = l_Lean_Elab_Command_elabCommand___closed__4; 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); @@ -14083,43 +17177,59 @@ return x_22; } else { -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_23 = lean_usize_of_nat(x_13); -lean_dec(x_13); -x_24 = lean_box(0); -x_25 = l_Lean_Elab_Command_elabCommand___boxed__const__1; -x_26 = lean_box_usize(x_23); -x_27 = lean_alloc_closure((void*)(l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__19___boxed), 7, 4); -lean_closure_set(x_27, 0, x_6); -lean_closure_set(x_27, 1, x_25); -lean_closure_set(x_27, 2, x_26); -lean_closure_set(x_27, 3, x_24); -x_28 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand___lambda__1), 5, 2); -lean_closure_set(x_28, 0, x_1); -lean_closure_set(x_28, 1, x_27); -x_29 = l_Lean_Elab_withLogging___at_Lean_Elab_Command_elabCommand___spec__2(x_28, x_2, x_3, x_4); -return x_29; +uint8_t x_23; +x_23 = lean_nat_dec_le(x_17, x_17); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +lean_dec(x_17); +lean_dec(x_6); +x_24 = l_Lean_Elab_Command_elabCommand___closed__4; +x_25 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand___lambda__1), 5, 2); +lean_closure_set(x_25, 0, x_1); +lean_closure_set(x_25, 1, x_24); +x_26 = l_Lean_Elab_withLogging___at_Lean_Elab_Command_elabCommand___spec__2(x_25, x_2, x_3, x_4); +return x_26; +} +else +{ +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; +x_27 = lean_usize_of_nat(x_17); +lean_dec(x_17); +x_28 = lean_box(0); +x_29 = l_Lean_Elab_Command_elabCommand___boxed__const__1; +x_30 = lean_box_usize(x_27); +x_31 = lean_alloc_closure((void*)(l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__18___boxed), 7, 4); +lean_closure_set(x_31, 0, x_6); +lean_closure_set(x_31, 1, x_29); +lean_closure_set(x_31, 2, x_30); +lean_closure_set(x_31, 3, x_28); +x_32 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand___lambda__1), 5, 2); +lean_closure_set(x_32, 0, x_1); +lean_closure_set(x_32, 1, x_31); +x_33 = l_Lean_Elab_withLogging___at_Lean_Elab_Command_elabCommand___spec__2(x_32, x_2, x_3, x_4); +return x_33; } } } } else { -lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_30 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___closed__4; +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_34 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___closed__4; lean_inc(x_1); -x_31 = lean_alloc_closure((void*)(l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkInfoTree___boxed), 6, 2); -lean_closure_set(x_31, 0, x_30); -lean_closure_set(x_31, 1, x_1); -x_32 = l_Lean_Elab_Command_elabCommand___closed__3; -x_33 = lean_alloc_closure((void*)(l_Lean_Elab_withInfoTreeContext___at___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___spec__2), 5, 2); -lean_closure_set(x_33, 0, x_32); -lean_closure_set(x_33, 1, x_31); -x_34 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand___lambda__1), 5, 2); -lean_closure_set(x_34, 0, x_1); -lean_closure_set(x_34, 1, x_33); -x_35 = l_Lean_Elab_withLogging___at_Lean_Elab_Command_elabCommand___spec__2(x_34, x_2, x_3, x_4); -return x_35; +x_35 = lean_alloc_closure((void*)(l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkInfoTree___boxed), 6, 2); +lean_closure_set(x_35, 0, x_34); +lean_closure_set(x_35, 1, x_1); +x_36 = l_Lean_Elab_Command_elabCommand___closed__3; +x_37 = lean_alloc_closure((void*)(l_Lean_Elab_withInfoTreeContext___at___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___spec__2), 5, 2); +lean_closure_set(x_37, 0, x_36); +lean_closure_set(x_37, 1, x_35); +x_38 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabCommand___lambda__1), 5, 2); +lean_closure_set(x_38, 0, x_1); +lean_closure_set(x_38, 1, x_37); +x_39 = l_Lean_Elab_withLogging___at_Lean_Elab_Command_elabCommand___spec__2(x_38, x_2, x_3, x_4); +return x_39; } } } @@ -14156,60 +17266,50 @@ x_7 = l_Lean_log___at_Lean_Elab_Command_elabCommand___spec__5(x_1, x_6, x_3, x_4 return x_7; } } -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_elabCommand___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_isTracingEnabledFor___at_Lean_Elab_Command_elabCommand___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_addTrace___at_Lean_Elab_Command_elabCommand___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_addTrace___at_Lean_Elab_Command_elabCommand___spec__7___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_addTrace___at_Lean_Elab_Command_elabCommand___spec__8(x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean_addTrace___at_Lean_Elab_Command_elabCommand___spec__7(x_1, x_2, x_3, x_4, x_5); lean_dec(x_4); lean_dec(x_3); return x_6; } } -LEAN_EXPORT lean_object* l_List_forM___at_Lean_Elab_Command_elabCommand___spec__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_List_forM___at_Lean_Elab_Command_elabCommand___spec__8___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_List_forM___at_Lean_Elab_Command_elabCommand___spec__9(x_1, x_2, x_3, x_4); +x_5 = l_List_forM___at_Lean_Elab_Command_elabCommand___spec__8(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_throwError___at_Lean_Elab_Command_elabCommand___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_Command_elabCommand___spec__10___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_elabCommand___spec__11(x_1, x_2, x_3, x_4); +x_5 = l_Lean_throwError___at_Lean_Elab_Command_elabCommand___spec__10(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_elabCommand___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_throwMaxRecDepthAt___at_Lean_Elab_Command_elabCommand___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_elabCommand___spec__12(x_1, x_2, x_3, x_4); +x_5 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabCommand___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_elabCommand___spec__13___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___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_elabCommand___spec__13(x_1, x_2); +x_3 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__12(x_1, x_2); lean_dec(x_2); lean_dec(x_1); return x_3; @@ -14251,36 +17351,36 @@ lean_dec(x_5); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabCommand___spec__16___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_Command_elabCommand___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_Command_elabCommand___spec__16(x_1, x_2, x_3, x_4); +x_5 = l_Lean_throwError___at_Lean_Elab_Command_elabCommand___spec__15(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_elabCommand___spec__17___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_elabCommand___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_throwMaxRecDepthAt___at_Lean_Elab_Command_elabCommand___spec__17(x_1, x_2, x_3, x_4); +x_5 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabCommand___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_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__18___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__17___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__18(x_1, x_2); +x_3 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__17(x_1, x_2); lean_dec(x_2); lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___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) { _start: { size_t x_8; size_t x_9; lean_object* x_10; @@ -14288,22 +17388,33 @@ 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_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__19(x_1, x_8, x_9, x_4, x_5, x_6, x_7); +x_10 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__18(x_1, x_8, x_9, x_4, x_5, x_6, x_7); lean_dec(x_1); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCommand___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCommand___lambda__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_Lean_Elab_Command_elabCommand___lambda__2(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_Lean_Elab_Command_elabCommand___lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_Lean_Elab_Command_elabCommand___lambda__4(x_1, x_2, x_3); +x_4 = l_Lean_Elab_Command_elabCommand___lambda__5(x_1, x_2, x_3); lean_dec(x_2); lean_dec(x_1); return x_4; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4918____closed__1() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_5179____closed__1() { _start: { lean_object* x_1; @@ -14311,33 +17422,33 @@ 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_4918____closed__2() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_5179____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__9; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4918____closed__1; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__1; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_5179____closed__1; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4918____closed__3() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_5179____closed__3() { _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_4417____closed__13; -x_2 = lean_unsigned_to_nat(4918u); +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__16; +x_2 = lean_unsigned_to_nat(5179u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4918_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_5179_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4918____closed__2; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_5179____closed__2; x_3 = 0; -x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4918____closed__3; +x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_5179____closed__3; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } @@ -14754,12 +17865,22 @@ static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__9; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__1; x_2 = l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCommandTopLevel___spec__6___closed__1; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCommandTopLevel___spec__6___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCommandTopLevel___spec__6(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: { @@ -14787,7 +17908,7 @@ if (lean_is_exclusive(x_5)) { x_12 = lean_box(0); } x_21 = l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCommandTopLevel___spec__6___closed__2; -x_22 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_elabCommand___spec__7(x_21, x_6, x_7, x_8); +x_22 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__13(x_21, x_6, x_7, x_8); x_23 = lean_ctor_get(x_22, 0); lean_inc(x_23); x_24 = lean_unbox(x_23); @@ -14799,7 +17920,7 @@ lean_dec(x_11); x_25 = lean_ctor_get(x_22, 1); lean_inc(x_25); lean_dec(x_22); -x_26 = l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__4___closed__1; +x_26 = l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCommandTopLevel___spec__6___closed__3; x_13 = x_26; x_14 = x_25; goto block_20; @@ -14826,7 +17947,7 @@ x_33 = l_Lean_addTrace___at_Lean_Elab_Command_elabCommandTopLevel___spec__2(x_21 x_34 = lean_ctor_get(x_33, 1); lean_inc(x_34); lean_dec(x_33); -x_35 = l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__4___closed__1; +x_35 = l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCommandTopLevel___spec__6___closed__3; x_13 = x_35; x_14 = x_34; goto block_20; @@ -15144,7 +18265,7 @@ if (lean_is_exclusive(x_5)) { x_12 = lean_box(0); } x_21 = l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCommandTopLevel___spec__6___closed__2; -x_22 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_elabCommand___spec__7(x_21, x_6, x_7, x_8); +x_22 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__13(x_21, x_6, x_7, x_8); x_23 = lean_ctor_get(x_22, 0); lean_inc(x_23); x_24 = lean_unbox(x_23); @@ -15156,7 +18277,7 @@ lean_dec(x_11); x_25 = lean_ctor_get(x_22, 1); lean_inc(x_25); lean_dec(x_22); -x_26 = l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__4___closed__1; +x_26 = l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCommandTopLevel___spec__6___closed__3; x_13 = x_26; x_14 = x_25; goto block_20; @@ -15183,7 +18304,7 @@ x_33 = l_Lean_addTrace___at_Lean_Elab_Command_elabCommandTopLevel___spec__2(x_21 x_34 = lean_ctor_get(x_33, 1); lean_inc(x_34); lean_dec(x_33); -x_35 = l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__4___closed__1; +x_35 = l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCommandTopLevel___spec__6___closed__3; x_13 = x_35; x_14 = x_34; goto block_20; @@ -15496,7 +18617,7 @@ static lean_object* _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabC _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__9; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__1; x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommandTopLevel___spec__12___lambda__1___closed__1; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; @@ -15506,29 +18627,21 @@ static lean_object* _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabC _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Tactic", 6); -return x_1; -} -} -static lean_object* _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommandTopLevel___spec__12___lambda__1___closed__4() { -_start: -{ -lean_object* x_1; x_1 = lean_mk_string_from_bytes("unsolvedGoals", 13); return x_1; } } -static lean_object* _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommandTopLevel___spec__12___lambda__1___closed__5() { +static lean_object* _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommandTopLevel___spec__12___lambda__1___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommandTopLevel___spec__12___lambda__1___closed__3; -x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommandTopLevel___spec__12___lambda__1___closed__4; +x_1 = l___auto____x40_Lean_Elab_Command___hyg_448____closed__3; +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommandTopLevel___spec__12___lambda__1___closed__3; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommandTopLevel___spec__12___lambda__1___closed__6() { +static lean_object* _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommandTopLevel___spec__12___lambda__1___closed__5() { _start: { lean_object* x_1; @@ -15536,12 +18649,12 @@ x_1 = lean_mk_string_from_bytes("_traceMsg", 9); return x_1; } } -static lean_object* _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommandTopLevel___spec__12___lambda__1___closed__7() { +static lean_object* _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommandTopLevel___spec__12___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_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommandTopLevel___spec__12___lambda__1___closed__6; +x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommandTopLevel___spec__12___lambda__1___closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -15555,12 +18668,12 @@ x_3 = lean_name_eq(x_1, x_2); if (x_3 == 0) { lean_object* x_4; uint8_t x_5; -x_4 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommandTopLevel___spec__12___lambda__1___closed__5; +x_4 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommandTopLevel___spec__12___lambda__1___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_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommandTopLevel___spec__12___lambda__1___closed__7; +x_6 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommandTopLevel___spec__12___lambda__1___closed__6; x_7 = l_Lean_Name_isSuffixOf(x_6, x_1); return x_7; } @@ -16106,7 +19219,7 @@ return x_57; } } } -static lean_object* _init_l_Lean_Elab_Command_elabCommandTopLevel___lambda__2___closed__1() { +static lean_object* _init_l_Lean_Elab_Command_elabCommandTopLevel___closed__1() { _start: { lean_object* x_1; @@ -16114,496 +19227,642 @@ x_1 = l_Lean_Elab_Command_showPartialSyntaxErrors; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCommandTopLevel___lambda__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_Command_elabCommandTopLevel(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; -lean_dec(x_2); -x_6 = lean_st_ref_take(x_4, x_5); -x_7 = lean_ctor_get(x_6, 0); +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; +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); -x_8 = lean_ctor_get(x_6, 1); -lean_inc(x_8); +lean_dec(x_5); +x_8 = l_Lean_replaceRef(x_1, x_6); lean_dec(x_6); -x_9 = !lean_is_exclusive(x_7); +x_9 = !lean_is_exclusive(x_2); 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; -x_10 = lean_ctor_get(x_7, 1); -x_11 = l_Lean_Elab_Command_State_messages___default___closed__3; -lean_ctor_set(x_7, 1, x_11); -x_12 = lean_st_ref_set(x_4, x_7, x_8); -x_13 = lean_ctor_get(x_12, 1); +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_2, 6); +lean_dec(x_10); +lean_ctor_set(x_2, 6, x_8); +x_11 = lean_st_ref_take(x_3, x_7); +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_12); -x_14 = l_Lean_Elab_getResetInfoTrees___at___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___spec__3___rarg(x_4, x_13); -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_4); -lean_inc(x_3); -lean_inc(x_1); -x_17 = l_Lean_Elab_Command_elabCommand(x_1, x_3, x_4, x_16); -if (lean_obj_tag(x_17) == 0) +lean_dec(x_11); +x_14 = !lean_is_exclusive(x_12); +if (x_14 == 0) { -lean_object* x_18; lean_object* x_19; lean_object* x_20; +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_15 = lean_ctor_get(x_12, 1); +x_16 = l_Lean_Elab_Command_State_messages___default___closed__3; +lean_ctor_set(x_12, 1, x_16); +x_17 = lean_st_ref_set(x_3, x_12, x_13); x_18 = lean_ctor_get(x_17, 1); lean_inc(x_18); lean_dec(x_17); -lean_inc(x_1); -x_19 = lean_alloc_closure((void*)(l_Lean_Elab_Command_runLinters), 4, 1); -lean_closure_set(x_19, 0, x_1); -lean_inc(x_4); +x_19 = l_Lean_Elab_getResetInfoTrees___at___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___spec__3___rarg(x_3, 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); lean_inc(x_3); -x_20 = l_Lean_Elab_withLogging___at_Lean_Elab_Command_elabCommand___spec__2(x_19, x_3, x_4, x_18); -if (lean_obj_tag(x_20) == 0) +lean_inc(x_2); +lean_inc(x_1); +x_22 = l_Lean_Elab_Command_elabCommand(x_1, x_2, x_3, x_21); +if (lean_obj_tag(x_22) == 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; uint8_t x_34; -x_21 = lean_ctor_get(x_20, 1); -lean_inc(x_21); -lean_dec(x_20); -x_22 = lean_st_ref_get(x_4, x_21); -x_23 = lean_ctor_get(x_22, 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_22, 1); -lean_inc(x_24); lean_dec(x_22); -x_25 = lean_ctor_get(x_23, 1); -lean_inc(x_25); -lean_dec(x_23); -x_26 = lean_st_ref_get(x_4, 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_1); +x_24 = lean_alloc_closure((void*)(l_Lean_Elab_Command_runLinters), 4, 1); +lean_closure_set(x_24, 0, x_1); +lean_inc(x_3); +lean_inc(x_2); +x_25 = l_Lean_Elab_withLogging___at_Lean_Elab_Command_elabCommand___spec__2(x_24, x_2, x_3, 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; uint8_t x_39; +x_26 = lean_ctor_get(x_25, 1); +lean_inc(x_26); +lean_dec(x_25); +x_27 = lean_st_ref_get(x_3, x_26); +x_28 = lean_ctor_get(x_27, 0); lean_inc(x_28); -lean_dec(x_26); -x_29 = lean_ctor_get(x_27, 2); +x_29 = lean_ctor_get(x_27, 1); lean_inc(x_29); lean_dec(x_27); -x_30 = l_Lean_Elab_Command_instInhabitedScope; -x_31 = l_List_head_x21___rarg(x_30, x_29); -lean_dec(x_29); -x_32 = lean_ctor_get(x_31, 1); +x_30 = lean_ctor_get(x_28, 1); +lean_inc(x_30); +lean_dec(x_28); +x_31 = lean_st_ref_get(x_3, x_29); +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_33 = l_Lean_Elab_Command_elabCommandTopLevel___lambda__2___closed__1; -x_34 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_32, x_33); +x_34 = lean_ctor_get(x_32, 2); +lean_inc(x_34); lean_dec(x_32); -if (x_34 == 0) +x_35 = l_Lean_Elab_Command_instInhabitedScope; +x_36 = l_List_head_x21___rarg(x_35, x_34); +lean_dec(x_34); +x_37 = lean_ctor_get(x_36, 1); +lean_inc(x_37); +lean_dec(x_36); +x_38 = l_Lean_Elab_Command_elabCommandTopLevel___closed__1; +x_39 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_37, x_38); +lean_dec(x_37); +if (x_39 == 0) { -uint8_t x_35; -lean_inc(x_10); -x_35 = l_Lean_PersistentArray_anyM___at_Lean_MessageLog_hasErrors___spec__1(x_10); -if (x_35 == 0) +uint8_t x_40; +lean_inc(x_15); +x_40 = l_Lean_PersistentArray_anyM___at_Lean_MessageLog_hasErrors___spec__1(x_15); +if (x_40 == 0) { -lean_object* x_36; lean_object* x_37; +lean_object* x_41; lean_object* x_42; lean_dec(x_1); -x_36 = lean_box(0); -x_37 = l_Lean_Elab_Command_elabCommandTopLevel___lambda__1(x_10, x_15, x_25, x_36, x_3, x_4, x_28); -return x_37; +x_41 = lean_box(0); +x_42 = l_Lean_Elab_Command_elabCommandTopLevel___lambda__1(x_15, x_20, x_30, x_41, x_2, x_3, x_33); +return x_42; } else { -lean_object* x_38; uint8_t x_39; -x_38 = l_Lean_Syntax_hasMissing(x_1); -x_39 = lean_unbox(x_38); -lean_dec(x_38); -if (x_39 == 0) +lean_object* x_43; uint8_t x_44; +x_43 = l_Lean_Syntax_hasMissing(x_1); +x_44 = lean_unbox(x_43); +lean_dec(x_43); +if (x_44 == 0) { -lean_object* x_40; lean_object* x_41; -x_40 = lean_box(0); -x_41 = l_Lean_Elab_Command_elabCommandTopLevel___lambda__1(x_10, x_15, x_25, x_40, x_3, x_4, x_28); -return x_41; +lean_object* x_45; lean_object* x_46; +x_45 = lean_box(0); +x_46 = l_Lean_Elab_Command_elabCommandTopLevel___lambda__1(x_15, x_20, x_30, x_45, x_2, x_3, x_33); +return x_46; } else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_42 = lean_unsigned_to_nat(0u); -x_43 = l_Lean_PersistentArray_foldlM___at_Lean_Elab_Command_elabCommandTopLevel___spec__8(x_25, x_11, x_42); -x_44 = lean_box(0); -x_45 = l_Lean_Elab_Command_elabCommandTopLevel___lambda__1(x_10, x_15, x_43, x_44, x_3, x_4, x_28); -return x_45; +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_47 = lean_unsigned_to_nat(0u); +x_48 = l_Lean_PersistentArray_foldlM___at_Lean_Elab_Command_elabCommandTopLevel___spec__8(x_30, x_16, x_47); +x_49 = lean_box(0); +x_50 = l_Lean_Elab_Command_elabCommandTopLevel___lambda__1(x_15, x_20, x_48, x_49, x_2, x_3, x_33); +return x_50; } } } else { -lean_object* x_46; lean_object* x_47; +lean_object* x_51; lean_object* x_52; lean_dec(x_1); -x_46 = lean_box(0); -x_47 = l_Lean_Elab_Command_elabCommandTopLevel___lambda__1(x_10, x_15, x_25, x_46, x_3, x_4, x_28); -return x_47; +x_51 = lean_box(0); +x_52 = l_Lean_Elab_Command_elabCommandTopLevel___lambda__1(x_15, x_20, x_30, x_51, x_2, x_3, x_33); +return x_52; } } else { -uint8_t x_48; +uint8_t x_53; +lean_dec(x_20); lean_dec(x_15); -lean_dec(x_10); -lean_dec(x_4); +lean_dec(x_2); lean_dec(x_3); lean_dec(x_1); -x_48 = !lean_is_exclusive(x_20); -if (x_48 == 0) +x_53 = !lean_is_exclusive(x_25); +if (x_53 == 0) { -return x_20; +return x_25; } else { -lean_object* x_49; lean_object* x_50; lean_object* x_51; -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, 2, 0); -lean_ctor_set(x_51, 0, x_49); -lean_ctor_set(x_51, 1, x_50); -return x_51; +lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_54 = lean_ctor_get(x_25, 0); +x_55 = lean_ctor_get(x_25, 1); +lean_inc(x_55); +lean_inc(x_54); +lean_dec(x_25); +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 { -uint8_t x_52; +uint8_t x_57; +lean_dec(x_20); lean_dec(x_15); -lean_dec(x_10); -lean_dec(x_4); +lean_dec(x_2); lean_dec(x_3); lean_dec(x_1); -x_52 = !lean_is_exclusive(x_17); -if (x_52 == 0) +x_57 = !lean_is_exclusive(x_22); +if (x_57 == 0) { -return x_17; +return x_22; } else { -lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_53 = lean_ctor_get(x_17, 0); -x_54 = lean_ctor_get(x_17, 1); -lean_inc(x_54); -lean_inc(x_53); -lean_dec(x_17); -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_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_22, 0); +x_59 = lean_ctor_get(x_22, 1); +lean_inc(x_59); +lean_inc(x_58); +lean_dec(x_22); +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 { -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; -x_56 = lean_ctor_get(x_7, 0); -x_57 = lean_ctor_get(x_7, 1); -x_58 = lean_ctor_get(x_7, 2); -x_59 = lean_ctor_get(x_7, 3); -x_60 = lean_ctor_get(x_7, 4); -x_61 = lean_ctor_get(x_7, 5); -x_62 = lean_ctor_get(x_7, 6); -x_63 = lean_ctor_get(x_7, 7); -x_64 = lean_ctor_get(x_7, 8); +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_61 = lean_ctor_get(x_12, 0); +x_62 = lean_ctor_get(x_12, 1); +x_63 = lean_ctor_get(x_12, 2); +x_64 = lean_ctor_get(x_12, 3); +x_65 = lean_ctor_get(x_12, 4); +x_66 = lean_ctor_get(x_12, 5); +x_67 = lean_ctor_get(x_12, 6); +x_68 = lean_ctor_get(x_12, 7); +x_69 = lean_ctor_get(x_12, 8); +lean_inc(x_69); +lean_inc(x_68); +lean_inc(x_67); +lean_inc(x_66); +lean_inc(x_65); lean_inc(x_64); lean_inc(x_63); lean_inc(x_62); lean_inc(x_61); -lean_inc(x_60); -lean_inc(x_59); -lean_inc(x_58); -lean_inc(x_57); -lean_inc(x_56); -lean_dec(x_7); -x_65 = l_Lean_Elab_Command_State_messages___default___closed__3; -x_66 = lean_alloc_ctor(0, 9, 0); -lean_ctor_set(x_66, 0, x_56); -lean_ctor_set(x_66, 1, x_65); -lean_ctor_set(x_66, 2, x_58); -lean_ctor_set(x_66, 3, x_59); -lean_ctor_set(x_66, 4, x_60); -lean_ctor_set(x_66, 5, x_61); -lean_ctor_set(x_66, 6, x_62); -lean_ctor_set(x_66, 7, x_63); -lean_ctor_set(x_66, 8, x_64); -x_67 = lean_st_ref_set(x_4, x_66, x_8); -x_68 = lean_ctor_get(x_67, 1); -lean_inc(x_68); -lean_dec(x_67); -x_69 = l_Lean_Elab_getResetInfoTrees___at___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___spec__3___rarg(x_4, x_68); -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); -lean_inc(x_4); -lean_inc(x_3); -lean_inc(x_1); -x_72 = l_Lean_Elab_Command_elabCommand(x_1, x_3, x_4, x_71); -if (lean_obj_tag(x_72) == 0) -{ -lean_object* x_73; lean_object* x_74; lean_object* x_75; +lean_dec(x_12); +x_70 = l_Lean_Elab_Command_State_messages___default___closed__3; +x_71 = lean_alloc_ctor(0, 9, 0); +lean_ctor_set(x_71, 0, x_61); +lean_ctor_set(x_71, 1, x_70); +lean_ctor_set(x_71, 2, x_63); +lean_ctor_set(x_71, 3, x_64); +lean_ctor_set(x_71, 4, x_65); +lean_ctor_set(x_71, 5, x_66); +lean_ctor_set(x_71, 6, x_67); +lean_ctor_set(x_71, 7, x_68); +lean_ctor_set(x_71, 8, x_69); +x_72 = lean_st_ref_set(x_3, x_71, x_13); x_73 = lean_ctor_get(x_72, 1); lean_inc(x_73); lean_dec(x_72); -lean_inc(x_1); -x_74 = lean_alloc_closure((void*)(l_Lean_Elab_Command_runLinters), 4, 1); -lean_closure_set(x_74, 0, x_1); -lean_inc(x_4); +x_74 = l_Lean_Elab_getResetInfoTrees___at___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___spec__3___rarg(x_3, 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); +lean_dec(x_74); lean_inc(x_3); -x_75 = l_Lean_Elab_withLogging___at_Lean_Elab_Command_elabCommand___spec__2(x_74, x_3, x_4, x_73); -if (lean_obj_tag(x_75) == 0) +lean_inc(x_2); +lean_inc(x_1); +x_77 = l_Lean_Elab_Command_elabCommand(x_1, x_2, x_3, x_76); +if (lean_obj_tag(x_77) == 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; uint8_t x_89; -x_76 = lean_ctor_get(x_75, 1); -lean_inc(x_76); -lean_dec(x_75); -x_77 = lean_st_ref_get(x_4, x_76); -x_78 = lean_ctor_get(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); -x_79 = lean_ctor_get(x_77, 1); -lean_inc(x_79); lean_dec(x_77); -x_80 = lean_ctor_get(x_78, 1); -lean_inc(x_80); -lean_dec(x_78); -x_81 = lean_st_ref_get(x_4, x_79); -x_82 = lean_ctor_get(x_81, 0); -lean_inc(x_82); -x_83 = lean_ctor_get(x_81, 1); +lean_inc(x_1); +x_79 = lean_alloc_closure((void*)(l_Lean_Elab_Command_runLinters), 4, 1); +lean_closure_set(x_79, 0, x_1); +lean_inc(x_3); +lean_inc(x_2); +x_80 = l_Lean_Elab_withLogging___at_Lean_Elab_Command_elabCommand___spec__2(x_79, x_2, x_3, 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; 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; uint8_t x_94; +x_81 = lean_ctor_get(x_80, 1); +lean_inc(x_81); +lean_dec(x_80); +x_82 = lean_st_ref_get(x_3, x_81); +x_83 = lean_ctor_get(x_82, 0); lean_inc(x_83); -lean_dec(x_81); -x_84 = lean_ctor_get(x_82, 2); +x_84 = lean_ctor_get(x_82, 1); lean_inc(x_84); lean_dec(x_82); -x_85 = l_Lean_Elab_Command_instInhabitedScope; -x_86 = l_List_head_x21___rarg(x_85, x_84); -lean_dec(x_84); -x_87 = lean_ctor_get(x_86, 1); +x_85 = lean_ctor_get(x_83, 1); +lean_inc(x_85); +lean_dec(x_83); +x_86 = lean_st_ref_get(x_3, x_84); +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 = l_Lean_Elab_Command_elabCommandTopLevel___lambda__2___closed__1; -x_89 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_87, x_88); +x_89 = lean_ctor_get(x_87, 2); +lean_inc(x_89); lean_dec(x_87); -if (x_89 == 0) +x_90 = l_Lean_Elab_Command_instInhabitedScope; +x_91 = l_List_head_x21___rarg(x_90, x_89); +lean_dec(x_89); +x_92 = lean_ctor_get(x_91, 1); +lean_inc(x_92); +lean_dec(x_91); +x_93 = l_Lean_Elab_Command_elabCommandTopLevel___closed__1; +x_94 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_92, x_93); +lean_dec(x_92); +if (x_94 == 0) { -uint8_t x_90; -lean_inc(x_57); -x_90 = l_Lean_PersistentArray_anyM___at_Lean_MessageLog_hasErrors___spec__1(x_57); -if (x_90 == 0) +uint8_t x_95; +lean_inc(x_62); +x_95 = l_Lean_PersistentArray_anyM___at_Lean_MessageLog_hasErrors___spec__1(x_62); +if (x_95 == 0) { -lean_object* x_91; lean_object* x_92; +lean_object* x_96; lean_object* x_97; lean_dec(x_1); -x_91 = lean_box(0); -x_92 = l_Lean_Elab_Command_elabCommandTopLevel___lambda__1(x_57, x_70, x_80, x_91, x_3, x_4, x_83); -return x_92; +x_96 = lean_box(0); +x_97 = l_Lean_Elab_Command_elabCommandTopLevel___lambda__1(x_62, x_75, x_85, x_96, x_2, x_3, x_88); +return x_97; } else { -lean_object* x_93; uint8_t x_94; -x_93 = l_Lean_Syntax_hasMissing(x_1); -x_94 = lean_unbox(x_93); -lean_dec(x_93); -if (x_94 == 0) +lean_object* x_98; uint8_t x_99; +x_98 = l_Lean_Syntax_hasMissing(x_1); +x_99 = lean_unbox(x_98); +lean_dec(x_98); +if (x_99 == 0) { -lean_object* x_95; lean_object* x_96; -x_95 = lean_box(0); -x_96 = l_Lean_Elab_Command_elabCommandTopLevel___lambda__1(x_57, x_70, x_80, x_95, x_3, x_4, x_83); -return x_96; +lean_object* x_100; lean_object* x_101; +x_100 = lean_box(0); +x_101 = l_Lean_Elab_Command_elabCommandTopLevel___lambda__1(x_62, x_75, x_85, x_100, x_2, x_3, x_88); +return x_101; } else { -lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; -x_97 = lean_unsigned_to_nat(0u); -x_98 = l_Lean_PersistentArray_foldlM___at_Lean_Elab_Command_elabCommandTopLevel___spec__8(x_80, x_65, x_97); -x_99 = lean_box(0); -x_100 = l_Lean_Elab_Command_elabCommandTopLevel___lambda__1(x_57, x_70, x_98, x_99, x_3, x_4, x_83); -return x_100; +lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; +x_102 = lean_unsigned_to_nat(0u); +x_103 = l_Lean_PersistentArray_foldlM___at_Lean_Elab_Command_elabCommandTopLevel___spec__8(x_85, x_70, x_102); +x_104 = lean_box(0); +x_105 = l_Lean_Elab_Command_elabCommandTopLevel___lambda__1(x_62, x_75, x_103, x_104, x_2, x_3, x_88); +return x_105; } } } else { -lean_object* x_101; lean_object* x_102; +lean_object* x_106; lean_object* x_107; lean_dec(x_1); -x_101 = lean_box(0); -x_102 = l_Lean_Elab_Command_elabCommandTopLevel___lambda__1(x_57, x_70, x_80, x_101, x_3, x_4, x_83); -return x_102; +x_106 = lean_box(0); +x_107 = l_Lean_Elab_Command_elabCommandTopLevel___lambda__1(x_62, x_75, x_85, x_106, x_2, x_3, x_88); +return x_107; } } else { -lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; -lean_dec(x_70); -lean_dec(x_57); -lean_dec(x_4); +lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; +lean_dec(x_75); +lean_dec(x_62); +lean_dec(x_2); lean_dec(x_3); lean_dec(x_1); -x_103 = lean_ctor_get(x_75, 0); -lean_inc(x_103); -x_104 = lean_ctor_get(x_75, 1); -lean_inc(x_104); -if (lean_is_exclusive(x_75)) { - lean_ctor_release(x_75, 0); - lean_ctor_release(x_75, 1); - x_105 = x_75; +x_108 = lean_ctor_get(x_80, 0); +lean_inc(x_108); +x_109 = lean_ctor_get(x_80, 1); +lean_inc(x_109); +if (lean_is_exclusive(x_80)) { + lean_ctor_release(x_80, 0); + lean_ctor_release(x_80, 1); + x_110 = x_80; } else { - lean_dec_ref(x_75); - x_105 = lean_box(0); + lean_dec_ref(x_80); + x_110 = lean_box(0); } -if (lean_is_scalar(x_105)) { - x_106 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_110)) { + x_111 = lean_alloc_ctor(1, 2, 0); } else { - x_106 = x_105; + x_111 = x_110; } -lean_ctor_set(x_106, 0, x_103); -lean_ctor_set(x_106, 1, x_104); -return x_106; +lean_ctor_set(x_111, 0, x_108); +lean_ctor_set(x_111, 1, x_109); +return x_111; } } else { -lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; -lean_dec(x_70); -lean_dec(x_57); -lean_dec(x_4); +lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; +lean_dec(x_75); +lean_dec(x_62); +lean_dec(x_2); lean_dec(x_3); lean_dec(x_1); -x_107 = lean_ctor_get(x_72, 0); -lean_inc(x_107); -x_108 = lean_ctor_get(x_72, 1); -lean_inc(x_108); -if (lean_is_exclusive(x_72)) { - lean_ctor_release(x_72, 0); - lean_ctor_release(x_72, 1); - x_109 = x_72; +x_112 = lean_ctor_get(x_77, 0); +lean_inc(x_112); +x_113 = lean_ctor_get(x_77, 1); +lean_inc(x_113); +if (lean_is_exclusive(x_77)) { + lean_ctor_release(x_77, 0); + lean_ctor_release(x_77, 1); + x_114 = x_77; } else { - lean_dec_ref(x_72); - x_109 = lean_box(0); + lean_dec_ref(x_77); + x_114 = lean_box(0); } -if (lean_is_scalar(x_109)) { - x_110 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_114)) { + x_115 = lean_alloc_ctor(1, 2, 0); } else { - x_110 = x_109; + x_115 = x_114; } -lean_ctor_set(x_110, 0, x_107); -lean_ctor_set(x_110, 1, x_108); -return x_110; +lean_ctor_set(x_115, 0, x_112); +lean_ctor_set(x_115, 1, x_113); +return x_115; } } } +else +{ +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; +x_116 = lean_ctor_get(x_2, 0); +x_117 = lean_ctor_get(x_2, 1); +x_118 = lean_ctor_get(x_2, 2); +x_119 = lean_ctor_get(x_2, 3); +x_120 = lean_ctor_get(x_2, 4); +x_121 = lean_ctor_get(x_2, 5); +x_122 = lean_ctor_get(x_2, 7); +lean_inc(x_122); +lean_inc(x_121); +lean_inc(x_120); +lean_inc(x_119); +lean_inc(x_118); +lean_inc(x_117); +lean_inc(x_116); +lean_dec(x_2); +x_123 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_123, 0, x_116); +lean_ctor_set(x_123, 1, x_117); +lean_ctor_set(x_123, 2, x_118); +lean_ctor_set(x_123, 3, x_119); +lean_ctor_set(x_123, 4, x_120); +lean_ctor_set(x_123, 5, x_121); +lean_ctor_set(x_123, 6, x_8); +lean_ctor_set(x_123, 7, x_122); +x_124 = lean_st_ref_take(x_3, x_7); +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_125, 0); +lean_inc(x_127); +x_128 = lean_ctor_get(x_125, 1); +lean_inc(x_128); +x_129 = lean_ctor_get(x_125, 2); +lean_inc(x_129); +x_130 = lean_ctor_get(x_125, 3); +lean_inc(x_130); +x_131 = lean_ctor_get(x_125, 4); +lean_inc(x_131); +x_132 = lean_ctor_get(x_125, 5); +lean_inc(x_132); +x_133 = lean_ctor_get(x_125, 6); +lean_inc(x_133); +x_134 = lean_ctor_get(x_125, 7); +lean_inc(x_134); +x_135 = lean_ctor_get(x_125, 8); +lean_inc(x_135); +if (lean_is_exclusive(x_125)) { + lean_ctor_release(x_125, 0); + lean_ctor_release(x_125, 1); + lean_ctor_release(x_125, 2); + lean_ctor_release(x_125, 3); + lean_ctor_release(x_125, 4); + lean_ctor_release(x_125, 5); + lean_ctor_release(x_125, 6); + lean_ctor_release(x_125, 7); + lean_ctor_release(x_125, 8); + x_136 = x_125; +} else { + lean_dec_ref(x_125); + x_136 = lean_box(0); } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCommandTopLevel(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: +x_137 = l_Lean_Elab_Command_State_messages___default___closed__3; +if (lean_is_scalar(x_136)) { + x_138 = lean_alloc_ctor(0, 9, 0); +} else { + x_138 = x_136; +} +lean_ctor_set(x_138, 0, x_127); +lean_ctor_set(x_138, 1, x_137); +lean_ctor_set(x_138, 2, x_129); +lean_ctor_set(x_138, 3, x_130); +lean_ctor_set(x_138, 4, x_131); +lean_ctor_set(x_138, 5, x_132); +lean_ctor_set(x_138, 6, x_133); +lean_ctor_set(x_138, 7, x_134); +lean_ctor_set(x_138, 8, x_135); +x_139 = lean_st_ref_set(x_3, x_138, x_126); +x_140 = lean_ctor_get(x_139, 1); +lean_inc(x_140); +lean_dec(x_139); +x_141 = l_Lean_Elab_getResetInfoTrees___at___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___spec__3___rarg(x_3, x_140); +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_3); +lean_inc(x_123); +lean_inc(x_1); +x_144 = l_Lean_Elab_Command_elabCommand(x_1, x_123, x_3, x_143); +if (lean_obj_tag(x_144) == 0) { -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_4918____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); -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); -x_10 = !lean_is_exclusive(x_2); -if (x_10 == 0) +lean_object* x_145; lean_object* x_146; lean_object* x_147; +x_145 = lean_ctor_get(x_144, 1); +lean_inc(x_145); +lean_dec(x_144); +lean_inc(x_1); +x_146 = lean_alloc_closure((void*)(l_Lean_Elab_Command_runLinters), 4, 1); +lean_closure_set(x_146, 0, x_1); +lean_inc(x_3); +lean_inc(x_123); +x_147 = l_Lean_Elab_withLogging___at_Lean_Elab_Command_elabCommand___spec__2(x_146, x_123, x_3, x_145); +if (lean_obj_tag(x_147) == 0) +{ +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; uint8_t x_161; +x_148 = lean_ctor_get(x_147, 1); +lean_inc(x_148); +lean_dec(x_147); +x_149 = lean_st_ref_get(x_3, 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_150, 1); +lean_inc(x_152); +lean_dec(x_150); +x_153 = lean_st_ref_get(x_3, x_151); +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_ctor_get(x_154, 2); +lean_inc(x_156); +lean_dec(x_154); +x_157 = l_Lean_Elab_Command_instInhabitedScope; +x_158 = l_List_head_x21___rarg(x_157, x_156); +lean_dec(x_156); +x_159 = lean_ctor_get(x_158, 1); +lean_inc(x_159); +lean_dec(x_158); +x_160 = l_Lean_Elab_Command_elabCommandTopLevel___closed__1; +x_161 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_159, x_160); +lean_dec(x_159); +if (x_161 == 0) { -lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_11 = lean_ctor_get(x_2, 6); -lean_dec(x_11); -lean_ctor_set(x_2, 6, x_9); -x_12 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_elabCommand___spec__7(x_5, x_2, x_3, x_8); -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_unbox(x_13); -lean_dec(x_13); -if (x_14 == 0) +uint8_t x_162; +lean_inc(x_128); +x_162 = l_Lean_PersistentArray_anyM___at_Lean_MessageLog_hasErrors___spec__1(x_128); +if (x_162 == 0) { -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_12, 1); -lean_inc(x_15); -lean_dec(x_12); -x_16 = lean_box(0); -x_17 = l_Lean_Elab_Command_elabCommandTopLevel___lambda__2(x_1, x_16, x_2, x_3, x_15); -return x_17; +lean_object* x_163; lean_object* x_164; +lean_dec(x_1); +x_163 = lean_box(0); +x_164 = l_Lean_Elab_Command_elabCommandTopLevel___lambda__1(x_128, x_142, x_152, x_163, x_123, x_3, x_155); +return x_164; } 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; -x_18 = lean_ctor_get(x_12, 1); -lean_inc(x_18); -lean_dec(x_12); -lean_inc(x_1); -x_19 = l_Lean_MessageData_ofSyntax(x_1); -x_20 = l_Lean_addTrace___at_Lean_Elab_Command_elabCommand___spec__8(x_5, x_19, x_2, x_3, x_18); -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 = l_Lean_Elab_Command_elabCommandTopLevel___lambda__2(x_1, x_21, x_2, x_3, x_22); -return x_23; +lean_object* x_165; uint8_t x_166; +x_165 = l_Lean_Syntax_hasMissing(x_1); +x_166 = lean_unbox(x_165); +lean_dec(x_165); +if (x_166 == 0) +{ +lean_object* x_167; lean_object* x_168; +x_167 = lean_box(0); +x_168 = l_Lean_Elab_Command_elabCommandTopLevel___lambda__1(x_128, x_142, x_152, x_167, x_123, x_3, x_155); +return x_168; +} +else +{ +lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; +x_169 = lean_unsigned_to_nat(0u); +x_170 = l_Lean_PersistentArray_foldlM___at_Lean_Elab_Command_elabCommandTopLevel___spec__8(x_152, x_137, x_169); +x_171 = lean_box(0); +x_172 = l_Lean_Elab_Command_elabCommandTopLevel___lambda__1(x_128, x_142, x_170, x_171, x_123, x_3, x_155); +return x_172; +} } } 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; uint8_t x_34; -x_24 = lean_ctor_get(x_2, 0); -x_25 = lean_ctor_get(x_2, 1); -x_26 = lean_ctor_get(x_2, 2); -x_27 = lean_ctor_get(x_2, 3); -x_28 = lean_ctor_get(x_2, 4); -x_29 = lean_ctor_get(x_2, 5); -x_30 = lean_ctor_get(x_2, 7); -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_dec(x_2); -x_31 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_31, 0, x_24); -lean_ctor_set(x_31, 1, x_25); -lean_ctor_set(x_31, 2, x_26); -lean_ctor_set(x_31, 3, x_27); -lean_ctor_set(x_31, 4, x_28); -lean_ctor_set(x_31, 5, x_29); -lean_ctor_set(x_31, 6, x_9); -lean_ctor_set(x_31, 7, x_30); -x_32 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_elabCommand___spec__7(x_5, x_31, x_3, x_8); -x_33 = lean_ctor_get(x_32, 0); -lean_inc(x_33); -x_34 = lean_unbox(x_33); -lean_dec(x_33); -if (x_34 == 0) +lean_object* x_173; lean_object* x_174; +lean_dec(x_1); +x_173 = lean_box(0); +x_174 = l_Lean_Elab_Command_elabCommandTopLevel___lambda__1(x_128, x_142, x_152, x_173, x_123, x_3, x_155); +return x_174; +} +} +else { -lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_35 = lean_ctor_get(x_32, 1); -lean_inc(x_35); -lean_dec(x_32); -x_36 = lean_box(0); -x_37 = l_Lean_Elab_Command_elabCommandTopLevel___lambda__2(x_1, x_36, x_31, x_3, x_35); -return x_37; +lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; +lean_dec(x_142); +lean_dec(x_128); +lean_dec(x_123); +lean_dec(x_3); +lean_dec(x_1); +x_175 = lean_ctor_get(x_147, 0); +lean_inc(x_175); +x_176 = lean_ctor_get(x_147, 1); +lean_inc(x_176); +if (lean_is_exclusive(x_147)) { + lean_ctor_release(x_147, 0); + lean_ctor_release(x_147, 1); + x_177 = x_147; +} else { + lean_dec_ref(x_147); + x_177 = lean_box(0); +} +if (lean_is_scalar(x_177)) { + x_178 = lean_alloc_ctor(1, 2, 0); +} else { + x_178 = x_177; +} +lean_ctor_set(x_178, 0, x_175); +lean_ctor_set(x_178, 1, x_176); +return x_178; +} } 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_32, 1); -lean_inc(x_38); -lean_dec(x_32); -lean_inc(x_1); -x_39 = l_Lean_MessageData_ofSyntax(x_1); -x_40 = l_Lean_addTrace___at_Lean_Elab_Command_elabCommand___spec__8(x_5, x_39, x_31, x_3, 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 = l_Lean_Elab_Command_elabCommandTopLevel___lambda__2(x_1, x_41, x_31, x_3, x_42); -return x_43; +lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; +lean_dec(x_142); +lean_dec(x_128); +lean_dec(x_123); +lean_dec(x_3); +lean_dec(x_1); +x_179 = lean_ctor_get(x_144, 0); +lean_inc(x_179); +x_180 = lean_ctor_get(x_144, 1); +lean_inc(x_180); +if (lean_is_exclusive(x_144)) { + lean_ctor_release(x_144, 0); + lean_ctor_release(x_144, 1); + x_181 = x_144; +} else { + lean_dec_ref(x_144); + x_181 = lean_box(0); +} +if (lean_is_scalar(x_181)) { + x_182 = lean_alloc_ctor(1, 2, 0); +} else { + x_182 = x_181; +} +lean_ctor_set(x_182, 0, x_179); +lean_ctor_set(x_182, 1, x_180); +return x_182; } } } @@ -16964,31 +20223,23 @@ static lean_object* _init_l_Lean_Elab_Command_getBracketedBinderIds___closed__1( _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Term", 4); -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_Command_getBracketedBinderIds___closed__2() { -_start: -{ -lean_object* x_1; x_1 = lean_mk_string_from_bytes("explicitBinder", 14); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_getBracketedBinderIds___closed__3() { +static lean_object* _init_l_Lean_Elab_Command_getBracketedBinderIds___closed__2() { _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_Elab_Command_mkCommandElabAttributeUnsafe___closed__5; -x_2 = l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__6; -x_3 = l_Lean_Elab_Command_getBracketedBinderIds___closed__1; -x_4 = l_Lean_Elab_Command_getBracketedBinderIds___closed__2; +x_1 = l___auto____x40_Lean_Elab_Command___hyg_448____closed__1; +x_2 = l___auto____x40_Lean_Elab_Command___hyg_448____closed__2; +x_3 = l___auto____x40_Lean_Elab_Command___hyg_448____closed__14; +x_4 = l_Lean_Elab_Command_getBracketedBinderIds___closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Elab_Command_getBracketedBinderIds___closed__4() { +static lean_object* _init_l_Lean_Elab_Command_getBracketedBinderIds___closed__3() { _start: { lean_object* x_1; @@ -16996,19 +20247,19 @@ x_1 = lean_mk_string_from_bytes("implicitBinder", 14); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_getBracketedBinderIds___closed__5() { +static lean_object* _init_l_Lean_Elab_Command_getBracketedBinderIds___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_Elab_Command_mkCommandElabAttributeUnsafe___closed__5; -x_2 = l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__6; -x_3 = l_Lean_Elab_Command_getBracketedBinderIds___closed__1; -x_4 = l_Lean_Elab_Command_getBracketedBinderIds___closed__4; +x_1 = l___auto____x40_Lean_Elab_Command___hyg_448____closed__1; +x_2 = l___auto____x40_Lean_Elab_Command___hyg_448____closed__2; +x_3 = l___auto____x40_Lean_Elab_Command___hyg_448____closed__14; +x_4 = l_Lean_Elab_Command_getBracketedBinderIds___closed__3; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Elab_Command_getBracketedBinderIds___closed__6() { +static lean_object* _init_l_Lean_Elab_Command_getBracketedBinderIds___closed__5() { _start: { lean_object* x_1; @@ -17016,19 +20267,19 @@ x_1 = lean_mk_string_from_bytes("instBinder", 10); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_getBracketedBinderIds___closed__7() { +static lean_object* _init_l_Lean_Elab_Command_getBracketedBinderIds___closed__6() { _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_Elab_Command_mkCommandElabAttributeUnsafe___closed__5; -x_2 = l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__6; -x_3 = l_Lean_Elab_Command_getBracketedBinderIds___closed__1; -x_4 = l_Lean_Elab_Command_getBracketedBinderIds___closed__6; +x_1 = l___auto____x40_Lean_Elab_Command___hyg_448____closed__1; +x_2 = l___auto____x40_Lean_Elab_Command___hyg_448____closed__2; +x_3 = l___auto____x40_Lean_Elab_Command___hyg_448____closed__14; +x_4 = l_Lean_Elab_Command_getBracketedBinderIds___closed__5; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Elab_Command_getBracketedBinderIds___closed__8() { +static lean_object* _init_l_Lean_Elab_Command_getBracketedBinderIds___closed__7() { _start: { lean_object* x_1; lean_object* x_2; @@ -17037,11 +20288,11 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Command_getBracketedBinderIds___closed__9() { +static lean_object* _init_l_Lean_Elab_Command_getBracketedBinderIds___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_getBracketedBinderIds___closed__8; +x_1 = l_Lean_Elab_Command_getBracketedBinderIds___closed__7; x_2 = lean_box(0); x_3 = lean_array_push(x_1, x_2); return x_3; @@ -17051,19 +20302,19 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_getBracketedBinderIds(lean_object* _start: { lean_object* x_2; uint8_t x_3; -x_2 = l_Lean_Elab_Command_getBracketedBinderIds___closed__3; +x_2 = l_Lean_Elab_Command_getBracketedBinderIds___closed__2; lean_inc(x_1); x_3 = l_Lean_Syntax_isOfKind(x_1, x_2); if (x_3 == 0) { lean_object* x_4; uint8_t x_5; -x_4 = l_Lean_Elab_Command_getBracketedBinderIds___closed__5; +x_4 = l_Lean_Elab_Command_getBracketedBinderIds___closed__4; lean_inc(x_1); x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); if (x_5 == 0) { lean_object* x_6; uint8_t x_7; -x_6 = l_Lean_Elab_Command_getBracketedBinderIds___closed__7; +x_6 = l_Lean_Elab_Command_getBracketedBinderIds___closed__6; lean_inc(x_1); x_7 = l_Lean_Syntax_isOfKind(x_1, x_6); if (x_7 == 0) @@ -17096,7 +20347,7 @@ return x_15; else { lean_object* x_16; -x_16 = l_Lean_Elab_Command_getBracketedBinderIds___closed__9; +x_16 = l_Lean_Elab_Command_getBracketedBinderIds___closed__8; return x_16; } } @@ -17108,7 +20359,7 @@ x_18 = l_Lean_Syntax_getArg(x_10, x_17); lean_dec(x_10); x_19 = l_Lean_Syntax_getId(x_18); lean_dec(x_18); -x_20 = l_Lean_Elab_Command_getBracketedBinderIds___closed__8; +x_20 = l_Lean_Elab_Command_getBracketedBinderIds___closed__7; x_21 = lean_array_push(x_20, x_19); return x_21; } @@ -18678,94 +21929,347 @@ if (lean_is_exclusive(x_96)) { lean_dec_ref(x_96); x_98 = lean_box(0); } -if (lean_is_scalar(x_98)) { - x_99 = lean_alloc_ctor(1, 2, 0); -} else { - x_99 = x_98; - lean_ctor_set_tag(x_99, 1); +if (lean_is_scalar(x_98)) { + x_99 = lean_alloc_ctor(1, 2, 0); +} else { + x_99 = x_98; + lean_ctor_set_tag(x_99, 1); +} +lean_ctor_set(x_99, 0, x_68); +lean_ctor_set(x_99, 1, x_97); +return x_99; +} +} +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; 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; +x_100 = lean_ctor_get(x_79, 0); +x_101 = lean_ctor_get(x_79, 1); +x_102 = lean_ctor_get(x_79, 2); +x_103 = lean_ctor_get(x_79, 3); +x_104 = lean_ctor_get(x_79, 4); +x_105 = lean_ctor_get(x_79, 5); +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_dec(x_79); +x_106 = lean_ctor_get_uint8(x_80, sizeof(void*)*2); +x_107 = lean_ctor_get(x_80, 0); +lean_inc(x_107); +if (lean_is_exclusive(x_80)) { + lean_ctor_release(x_80, 0); + lean_ctor_release(x_80, 1); + x_108 = x_80; +} else { + lean_dec_ref(x_80); + x_108 = lean_box(0); +} +x_109 = l_Lean_PersistentArray_append___rarg(x_17, x_76); +if (lean_is_scalar(x_108)) { + x_110 = lean_alloc_ctor(0, 2, 1); +} else { + x_110 = x_108; +} +lean_ctor_set(x_110, 0, x_107); +lean_ctor_set(x_110, 1, x_109); +lean_ctor_set_uint8(x_110, sizeof(void*)*2, x_106); +x_111 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_111, 0, x_100); +lean_ctor_set(x_111, 1, x_101); +lean_ctor_set(x_111, 2, x_102); +lean_ctor_set(x_111, 3, x_103); +lean_ctor_set(x_111, 4, x_104); +lean_ctor_set(x_111, 5, x_105); +lean_ctor_set(x_111, 6, x_110); +x_112 = lean_st_ref_set(x_7, x_111, x_81); +lean_dec(x_7); +x_113 = lean_ctor_get(x_112, 1); +lean_inc(x_113); +if (lean_is_exclusive(x_112)) { + lean_ctor_release(x_112, 0); + lean_ctor_release(x_112, 1); + x_114 = x_112; +} else { + lean_dec_ref(x_112); + x_114 = lean_box(0); +} +if (lean_is_scalar(x_114)) { + x_115 = lean_alloc_ctor(1, 2, 0); +} else { + x_115 = x_114; + lean_ctor_set_tag(x_115, 1); +} +lean_ctor_set(x_115, 0, x_68); +lean_ctor_set(x_115, 1, x_113); +return x_115; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_withSaveInfoContext___at_Lean_Elab_Command_liftTermElabM___spec__1(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Elab_withSaveInfoContext___at_Lean_Elab_Command_liftTermElabM___spec__1___rarg), 8, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftTermElabM___spec__14(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 = l_Lean_PersistentArray_mapMAux___at_Lean_Elab_Command_liftTermElabM___spec__13(x_1, 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_Lean_Elab_Command_liftTermElabM___spec__15(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; uint8_t x_9; +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_is_exclusive(x_6); +if (x_9 == 0) +{ +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_10 = lean_ctor_get(x_6, 0); +x_11 = lean_ctor_get(x_1, 6); +x_12 = l_Lean_replaceRef(x_10, x_11); +lean_dec(x_10); +lean_ctor_set(x_6, 0, x_12); +x_13 = 1; +x_14 = lean_usize_add(x_3, x_13); +x_15 = lean_array_uset(x_8, x_3, x_6); +x_3 = x_14; +x_4 = x_15; +goto _start; +} +else +{ +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; +x_17 = lean_ctor_get(x_6, 0); +x_18 = lean_ctor_get(x_6, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_6); +x_19 = lean_ctor_get(x_1, 6); +x_20 = l_Lean_replaceRef(x_17, x_19); +lean_dec(x_17); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_18); +x_22 = 1; +x_23 = lean_usize_add(x_3, x_22); +x_24 = lean_array_uset(x_8, x_3, x_21); +x_3 = x_23; +x_4 = x_24; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapMAux___at_Lean_Elab_Command_liftTermElabM___spec__13(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_3; +x_3 = !lean_is_exclusive(x_2); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; size_t x_6; size_t x_7; lean_object* x_8; +x_4 = lean_ctor_get(x_2, 0); +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_Lean_Elab_Command_liftTermElabM___spec__14(x_1, x_6, x_7, x_4); +lean_ctor_set(x_2, 0, x_8); +return x_2; +} +else +{ +lean_object* x_9; lean_object* x_10; size_t x_11; size_t x_12; lean_object* x_13; lean_object* x_14; +x_9 = lean_ctor_get(x_2, 0); +lean_inc(x_9); +lean_dec(x_2); +x_10 = lean_array_get_size(x_9); +x_11 = lean_usize_of_nat(x_10); +lean_dec(x_10); +x_12 = 0; +x_13 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftTermElabM___spec__14(x_1, x_11, x_12, x_9); +x_14 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_14, 0, x_13); +return x_14; } -lean_ctor_set(x_99, 0, x_68); -lean_ctor_set(x_99, 1, x_97); -return x_99; } +else +{ +uint8_t x_15; +x_15 = !lean_is_exclusive(x_2); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; size_t x_18; size_t x_19; lean_object* x_20; +x_16 = lean_ctor_get(x_2, 0); +x_17 = lean_array_get_size(x_16); +x_18 = lean_usize_of_nat(x_17); +lean_dec(x_17); +x_19 = 0; +x_20 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftTermElabM___spec__15(x_1, x_18, x_19, x_16); +lean_ctor_set(x_2, 0, x_20); +return x_2; } 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; 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; -x_100 = lean_ctor_get(x_79, 0); -x_101 = lean_ctor_get(x_79, 1); -x_102 = lean_ctor_get(x_79, 2); -x_103 = lean_ctor_get(x_79, 3); -x_104 = lean_ctor_get(x_79, 4); -x_105 = lean_ctor_get(x_79, 5); -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_dec(x_79); -x_106 = lean_ctor_get_uint8(x_80, sizeof(void*)*2); -x_107 = lean_ctor_get(x_80, 0); -lean_inc(x_107); -if (lean_is_exclusive(x_80)) { - lean_ctor_release(x_80, 0); - lean_ctor_release(x_80, 1); - x_108 = x_80; -} else { - lean_dec_ref(x_80); - x_108 = lean_box(0); +lean_object* x_21; lean_object* x_22; size_t x_23; size_t x_24; lean_object* x_25; lean_object* x_26; +x_21 = lean_ctor_get(x_2, 0); +lean_inc(x_21); +lean_dec(x_2); +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_mapMUnsafe_map___at_Lean_Elab_Command_liftTermElabM___spec__15(x_1, x_23, x_24, x_21); +x_26 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_26, 0, x_25); +return x_26; } -x_109 = l_Lean_PersistentArray_append___rarg(x_17, x_76); -if (lean_is_scalar(x_108)) { - x_110 = lean_alloc_ctor(0, 2, 1); -} else { - x_110 = x_108; } -lean_ctor_set(x_110, 0, x_107); -lean_ctor_set(x_110, 1, x_109); -lean_ctor_set_uint8(x_110, sizeof(void*)*2, x_106); -x_111 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_111, 0, x_100); -lean_ctor_set(x_111, 1, x_101); -lean_ctor_set(x_111, 2, x_102); -lean_ctor_set(x_111, 3, x_103); -lean_ctor_set(x_111, 4, x_104); -lean_ctor_set(x_111, 5, x_105); -lean_ctor_set(x_111, 6, x_110); -x_112 = lean_st_ref_set(x_7, x_111, x_81); -lean_dec(x_7); -x_113 = lean_ctor_get(x_112, 1); -lean_inc(x_113); -if (lean_is_exclusive(x_112)) { - lean_ctor_release(x_112, 0); - lean_ctor_release(x_112, 1); - x_114 = x_112; -} else { - lean_dec_ref(x_112); - x_114 = lean_box(0); } -if (lean_is_scalar(x_114)) { - x_115 = lean_alloc_ctor(1, 2, 0); -} else { - x_115 = x_114; - lean_ctor_set_tag(x_115, 1); } -lean_ctor_set(x_115, 0, x_68); -lean_ctor_set(x_115, 1, x_113); -return x_115; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftTermElabM___spec__16(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; uint8_t x_9; +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_is_exclusive(x_6); +if (x_9 == 0) +{ +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_10 = lean_ctor_get(x_6, 0); +x_11 = lean_ctor_get(x_1, 6); +x_12 = l_Lean_replaceRef(x_10, x_11); +lean_dec(x_10); +lean_ctor_set(x_6, 0, x_12); +x_13 = 1; +x_14 = lean_usize_add(x_3, x_13); +x_15 = lean_array_uset(x_8, x_3, x_6); +x_3 = x_14; +x_4 = x_15; +goto _start; } +else +{ +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; +x_17 = lean_ctor_get(x_6, 0); +x_18 = lean_ctor_get(x_6, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_6); +x_19 = lean_ctor_get(x_1, 6); +x_20 = l_Lean_replaceRef(x_17, x_19); +lean_dec(x_17); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_18); +x_22 = 1; +x_23 = lean_usize_add(x_3, x_22); +x_24 = lean_array_uset(x_8, x_3, x_21); +x_3 = x_23; +x_4 = x_24; +goto _start; } } } } -LEAN_EXPORT lean_object* l_Lean_Elab_withSaveInfoContext___at_Lean_Elab_Command_liftTermElabM___spec__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapM___at_Lean_Elab_Command_liftTermElabM___spec__12(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Elab_withSaveInfoContext___at_Lean_Elab_Command_liftTermElabM___spec__1___rarg), 8, 0); +uint8_t x_3; +x_3 = !lean_is_exclusive(x_2); +if (x_3 == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; size_t x_8; size_t x_9; lean_object* x_10; +x_4 = lean_ctor_get(x_2, 0); +x_5 = lean_ctor_get(x_2, 1); +x_6 = l_Lean_PersistentArray_mapMAux___at_Lean_Elab_Command_liftTermElabM___spec__13(x_1, x_4); +x_7 = lean_array_get_size(x_5); +x_8 = lean_usize_of_nat(x_7); +lean_dec(x_7); +x_9 = 0; +x_10 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftTermElabM___spec__16(x_1, x_8, x_9, x_5); +lean_ctor_set(x_2, 1, x_10); +lean_ctor_set(x_2, 0, x_6); return x_2; } +else +{ +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; size_t x_18; size_t x_19; lean_object* x_20; lean_object* x_21; +x_11 = lean_ctor_get(x_2, 0); +x_12 = lean_ctor_get(x_2, 1); +x_13 = lean_ctor_get(x_2, 2); +x_14 = lean_ctor_get_usize(x_2, 4); +x_15 = lean_ctor_get(x_2, 3); +lean_inc(x_15); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_dec(x_2); +x_16 = l_Lean_PersistentArray_mapMAux___at_Lean_Elab_Command_liftTermElabM___spec__13(x_1, x_11); +x_17 = lean_array_get_size(x_12); +x_18 = lean_usize_of_nat(x_17); +lean_dec(x_17); +x_19 = 0; +x_20 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftTermElabM___spec__16(x_1, x_18, x_19, x_12); +x_21 = lean_alloc_ctor(0, 4, sizeof(size_t)*1); +lean_ctor_set(x_21, 0, x_16); +lean_ctor_set(x_21, 1, x_20); +lean_ctor_set(x_21, 2, x_13); +lean_ctor_set(x_21, 3, x_15); +lean_ctor_set_usize(x_21, 4, x_14); +return x_21; +} +} } LEAN_EXPORT lean_object* l_Lean_Elab_Command_liftTermElabM___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) { _start: @@ -19008,7 +22512,7 @@ return x_5; LEAN_EXPORT lean_object* l_Lean_Elab_Command_liftTermElabM___rarg(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; uint8_t x_25; +lean_object* x_5; lean_object* x_6; lean_object* x_7; 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; uint8_t x_26; x_5 = lean_st_ref_get(x_3, x_4); x_6 = lean_ctor_get(x_5, 0); lean_inc(x_6); @@ -19031,529 +22535,530 @@ x_14 = lean_ctor_get(x_6, 6); lean_inc(x_14); x_15 = lean_ctor_get(x_6, 7); lean_inc(x_15); -x_16 = l_Lean_Elab_Command_instInhabitedScope; -x_17 = l_List_head_x21___rarg(x_16, x_12); +x_16 = lean_ctor_get(x_6, 8); +lean_inc(x_16); +x_17 = l_Lean_Elab_Command_instInhabitedScope; +x_18 = l_List_head_x21___rarg(x_17, x_12); lean_dec(x_12); -x_18 = lean_alloc_closure((void*)(l_Lean_Elab_Command_liftTermElabM___rarg___lambda__1), 8, 1); -lean_closure_set(x_18, 0, x_1); -x_19 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkTermContext(x_2, x_6); -x_20 = lean_ctor_get(x_17, 4); -lean_inc(x_20); -lean_dec(x_17); -x_21 = lean_box(0); +x_19 = lean_alloc_closure((void*)(l_Lean_Elab_Command_liftTermElabM___rarg___lambda__1), 8, 1); +lean_closure_set(x_19, 0, x_1); +x_20 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkTermContext(x_2, x_6); +x_21 = lean_ctor_get(x_18, 4); +lean_inc(x_21); +lean_dec(x_18); x_22 = lean_box(0); -x_23 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_23, 0, x_20); -lean_ctor_set(x_23, 1, x_21); -lean_ctor_set(x_23, 2, x_22); -lean_ctor_set(x_23, 3, x_21); -lean_ctor_set(x_23, 4, x_22); -x_24 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkCoreContext(x_2, x_6, x_9); +x_23 = lean_box(0); +x_24 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_24, 0, x_21); +lean_ctor_set(x_24, 1, x_22); +lean_ctor_set(x_24, 2, x_23); +lean_ctor_set(x_24, 3, x_22); +lean_ctor_set(x_24, 4, x_23); +x_25 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkCoreContext(x_2, x_6, x_9); lean_dec(x_6); -x_25 = !lean_is_exclusive(x_15); -if (x_25 == 0) +x_26 = !lean_is_exclusive(x_15); +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; 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_15, 1); -lean_dec(x_26); -x_27 = lean_ctor_get(x_15, 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; +x_27 = lean_ctor_get(x_15, 1); lean_dec(x_27); -x_28 = l_Lean_Elab_Command_State_infoState___default___closed__3; -x_29 = l_Lean_Elab_Command_State_messages___default___closed__3; -lean_ctor_set(x_15, 1, x_29); -lean_ctor_set(x_15, 0, x_28); -x_30 = l_Lean_Elab_Command_liftCoreM___rarg___closed__1; -x_31 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_31, 0, x_11); -lean_ctor_set(x_31, 1, x_13); -lean_ctor_set(x_31, 2, x_14); -lean_ctor_set(x_31, 3, x_29); -lean_ctor_set(x_31, 4, x_30); -lean_ctor_set(x_31, 5, x_29); -lean_ctor_set(x_31, 6, x_15); -x_32 = lean_st_mk_ref(x_31, x_10); -x_33 = lean_ctor_get(x_32, 0); -lean_inc(x_33); -x_34 = lean_ctor_get(x_32, 1); +x_28 = lean_ctor_get(x_15, 0); +lean_dec(x_28); +x_29 = l_Lean_Elab_Command_State_infoState___default___closed__3; +x_30 = l_Lean_Elab_Command_State_messages___default___closed__3; +lean_ctor_set(x_15, 1, x_30); +lean_ctor_set(x_15, 0, x_29); +x_31 = l_Lean_Elab_Command_liftCoreM___rarg___closed__1; +x_32 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_32, 0, x_11); +lean_ctor_set(x_32, 1, x_13); +lean_ctor_set(x_32, 2, x_14); +lean_ctor_set(x_32, 3, x_16); +lean_ctor_set(x_32, 4, x_31); +lean_ctor_set(x_32, 5, x_30); +lean_ctor_set(x_32, 6, x_15); +x_33 = lean_st_mk_ref(x_32, x_10); +x_34 = lean_ctor_get(x_33, 0); lean_inc(x_34); -lean_dec(x_32); -x_35 = l_Lean_Elab_Command_liftTermElabM___rarg___closed__15; -x_36 = lean_st_mk_ref(x_35, x_34); -x_37 = lean_ctor_get(x_36, 0); -lean_inc(x_37); -x_38 = lean_ctor_get(x_36, 1); +x_35 = lean_ctor_get(x_33, 1); +lean_inc(x_35); +lean_dec(x_33); +x_36 = l_Lean_Elab_Command_liftTermElabM___rarg___closed__15; +x_37 = lean_st_mk_ref(x_36, x_35); +x_38 = lean_ctor_get(x_37, 0); lean_inc(x_38); -lean_dec(x_36); -x_39 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkMetaContext; -lean_inc(x_33); -lean_inc(x_37); -x_40 = l_Lean_Elab_Term_TermElabM_run___rarg(x_18, x_19, x_23, x_39, x_37, x_24, x_33, x_38); -if (lean_obj_tag(x_40) == 0) +x_39 = lean_ctor_get(x_37, 1); +lean_inc(x_39); +lean_dec(x_37); +x_40 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkMetaContext; +lean_inc(x_34); +lean_inc(x_38); +x_41 = l_Lean_Elab_Term_TermElabM_run___rarg(x_19, x_20, x_24, x_40, x_38, x_25, x_34, x_39); +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; 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; uint8_t x_58; -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; 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; uint8_t x_59; +x_42 = lean_ctor_get(x_41, 0); lean_inc(x_42); -lean_dec(x_40); -x_43 = lean_st_ref_get(x_37, x_42); -lean_dec(x_37); -x_44 = lean_ctor_get(x_43, 1); -lean_inc(x_44); -lean_dec(x_43); -x_45 = lean_st_ref_get(x_33, x_44); -lean_dec(x_33); -x_46 = lean_ctor_get(x_45, 0); -lean_inc(x_46); -x_47 = lean_ctor_get(x_45, 1); +x_43 = lean_ctor_get(x_41, 1); +lean_inc(x_43); +lean_dec(x_41); +x_44 = lean_st_ref_get(x_38, x_43); +lean_dec(x_38); +x_45 = lean_ctor_get(x_44, 1); +lean_inc(x_45); +lean_dec(x_44); +x_46 = lean_st_ref_get(x_34, x_45); +lean_dec(x_34); +x_47 = lean_ctor_get(x_46, 0); lean_inc(x_47); -lean_dec(x_45); -x_48 = lean_ctor_get(x_41, 0); +x_48 = lean_ctor_get(x_46, 1); lean_inc(x_48); -lean_dec(x_41); -x_49 = lean_st_ref_take(x_3, x_47); -x_50 = lean_ctor_get(x_49, 0); -lean_inc(x_50); -x_51 = lean_ctor_get(x_49, 1); +lean_dec(x_46); +x_49 = lean_ctor_get(x_42, 0); +lean_inc(x_49); +lean_dec(x_42); +x_50 = lean_st_ref_take(x_3, x_48); +x_51 = lean_ctor_get(x_50, 0); lean_inc(x_51); -lean_dec(x_49); -x_52 = lean_ctor_get(x_46, 0); +x_52 = lean_ctor_get(x_50, 1); lean_inc(x_52); -x_53 = lean_ctor_get(x_46, 1); +lean_dec(x_50); +x_53 = lean_ctor_get(x_47, 0); lean_inc(x_53); -x_54 = lean_ctor_get(x_46, 2); +x_54 = lean_ctor_get(x_47, 1); lean_inc(x_54); -x_55 = lean_ctor_get(x_46, 3); +x_55 = lean_ctor_get(x_47, 2); lean_inc(x_55); -x_56 = lean_ctor_get(x_46, 5); +x_56 = lean_ctor_get(x_47, 3); lean_inc(x_56); -x_57 = lean_ctor_get(x_46, 6); +x_57 = lean_ctor_get(x_47, 5); lean_inc(x_57); -lean_dec(x_46); -x_58 = !lean_is_exclusive(x_50); -if (x_58 == 0) +x_58 = lean_ctor_get(x_47, 6); +lean_inc(x_58); +lean_dec(x_47); +x_59 = !lean_is_exclusive(x_51); +if (x_59 == 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; uint8_t x_66; -x_59 = lean_ctor_get(x_50, 1); -x_60 = lean_ctor_get(x_50, 7); -x_61 = lean_ctor_get(x_50, 6); -lean_dec(x_61); -x_62 = lean_ctor_get(x_50, 3); +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; +x_60 = lean_ctor_get(x_51, 1); +x_61 = lean_ctor_get(x_51, 7); +x_62 = lean_ctor_get(x_51, 8); lean_dec(x_62); -x_63 = lean_ctor_get(x_50, 0); +x_63 = lean_ctor_get(x_51, 6); lean_dec(x_63); -x_64 = l_Lean_PersistentArray_append___rarg(x_59, x_56); -x_65 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore(x_2, x_64, x_55); -x_66 = !lean_is_exclusive(x_60); -if (x_66 == 0) -{ -lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; -x_67 = lean_ctor_get(x_60, 1); -x_68 = lean_ctor_get(x_57, 1); -lean_inc(x_68); -lean_dec(x_57); -x_69 = l_Lean_PersistentArray_append___rarg(x_67, x_68); -lean_ctor_set(x_60, 1, x_69); -lean_ctor_set(x_50, 6, x_54); -lean_ctor_set(x_50, 3, x_53); -lean_ctor_set(x_50, 1, x_65); -lean_ctor_set(x_50, 0, x_52); -x_70 = lean_st_ref_set(x_3, x_50, x_51); -if (lean_obj_tag(x_48) == 0) -{ -uint8_t x_71; -x_71 = !lean_is_exclusive(x_70); -if (x_71 == 0) -{ -lean_object* x_72; lean_object* x_73; -x_72 = lean_ctor_get(x_70, 0); -lean_dec(x_72); -x_73 = lean_ctor_get(x_48, 0); -lean_inc(x_73); -lean_dec(x_48); -lean_ctor_set_tag(x_70, 1); -lean_ctor_set(x_70, 0, x_73); -return x_70; -} -else +x_64 = lean_ctor_get(x_51, 3); +lean_dec(x_64); +x_65 = lean_ctor_get(x_51, 0); +lean_dec(x_65); +x_66 = l_Lean_PersistentArray_append___rarg(x_60, x_57); +x_67 = !lean_is_exclusive(x_61); +if (x_67 == 0) { -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_ctor_get(x_48, 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_61, 1); +x_69 = lean_ctor_get(x_58, 1); +lean_inc(x_69); +lean_dec(x_58); +x_70 = l_Lean_PersistentArray_append___rarg(x_68, x_69); +lean_ctor_set(x_61, 1, x_70); +x_71 = l_Lean_PersistentArray_mapM___at_Lean_Elab_Command_liftTermElabM___spec__12(x_2, x_56); +lean_ctor_set(x_51, 8, x_71); +lean_ctor_set(x_51, 6, x_55); +lean_ctor_set(x_51, 3, x_54); +lean_ctor_set(x_51, 1, x_66); +lean_ctor_set(x_51, 0, x_53); +x_72 = lean_st_ref_set(x_3, x_51, x_52); +if (lean_obj_tag(x_49) == 0) +{ +uint8_t x_73; +x_73 = !lean_is_exclusive(x_72); +if (x_73 == 0) +{ +lean_object* x_74; lean_object* x_75; +x_74 = lean_ctor_get(x_72, 0); +lean_dec(x_74); +x_75 = lean_ctor_get(x_49, 0); lean_inc(x_75); -lean_dec(x_48); -x_76 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_76, 0, x_75); -lean_ctor_set(x_76, 1, x_74); -return x_76; +lean_dec(x_49); +lean_ctor_set_tag(x_72, 1); +lean_ctor_set(x_72, 0, x_75); +return x_72; +} +else +{ +lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_76 = lean_ctor_get(x_72, 1); +lean_inc(x_76); +lean_dec(x_72); +x_77 = lean_ctor_get(x_49, 0); +lean_inc(x_77); +lean_dec(x_49); +x_78 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_78, 0, x_77); +lean_ctor_set(x_78, 1, x_76); +return x_78; } } else { -uint8_t x_77; -x_77 = !lean_is_exclusive(x_70); -if (x_77 == 0) +uint8_t x_79; +x_79 = !lean_is_exclusive(x_72); +if (x_79 == 0) { -lean_object* x_78; lean_object* x_79; -x_78 = lean_ctor_get(x_70, 0); -lean_dec(x_78); -x_79 = lean_ctor_get(x_48, 0); -lean_inc(x_79); -lean_dec(x_48); -lean_ctor_set(x_70, 0, x_79); -return x_70; +lean_object* x_80; lean_object* x_81; +x_80 = lean_ctor_get(x_72, 0); +lean_dec(x_80); +x_81 = lean_ctor_get(x_49, 0); +lean_inc(x_81); +lean_dec(x_49); +lean_ctor_set(x_72, 0, x_81); +return x_72; } else { -lean_object* x_80; lean_object* x_81; lean_object* x_82; -x_80 = lean_ctor_get(x_70, 1); -lean_inc(x_80); -lean_dec(x_70); -x_81 = lean_ctor_get(x_48, 0); -lean_inc(x_81); -lean_dec(x_48); -x_82 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_82, 0, x_81); -lean_ctor_set(x_82, 1, x_80); -return x_82; +lean_object* x_82; lean_object* x_83; lean_object* x_84; +x_82 = lean_ctor_get(x_72, 1); +lean_inc(x_82); +lean_dec(x_72); +x_83 = lean_ctor_get(x_49, 0); +lean_inc(x_83); +lean_dec(x_49); +x_84 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_84, 0, x_83); +lean_ctor_set(x_84, 1, x_82); +return x_84; } } } else { -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; -x_83 = lean_ctor_get_uint8(x_60, sizeof(void*)*2); -x_84 = lean_ctor_get(x_60, 0); -x_85 = lean_ctor_get(x_60, 1); -lean_inc(x_85); -lean_inc(x_84); -lean_dec(x_60); -x_86 = lean_ctor_get(x_57, 1); +uint8_t 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_uint8(x_61, sizeof(void*)*2); +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_57); -x_87 = l_Lean_PersistentArray_append___rarg(x_85, x_86); -x_88 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_88, 0, x_84); -lean_ctor_set(x_88, 1, x_87); -lean_ctor_set_uint8(x_88, sizeof(void*)*2, x_83); -lean_ctor_set(x_50, 7, x_88); -lean_ctor_set(x_50, 6, x_54); -lean_ctor_set(x_50, 3, x_53); -lean_ctor_set(x_50, 1, x_65); -lean_ctor_set(x_50, 0, x_52); -x_89 = lean_st_ref_set(x_3, x_50, x_51); -if (lean_obj_tag(x_48) == 0) +lean_dec(x_61); +x_88 = lean_ctor_get(x_58, 1); +lean_inc(x_88); +lean_dec(x_58); +x_89 = l_Lean_PersistentArray_append___rarg(x_87, x_88); +x_90 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_90, 0, x_86); +lean_ctor_set(x_90, 1, x_89); +lean_ctor_set_uint8(x_90, sizeof(void*)*2, x_85); +x_91 = l_Lean_PersistentArray_mapM___at_Lean_Elab_Command_liftTermElabM___spec__12(x_2, x_56); +lean_ctor_set(x_51, 8, x_91); +lean_ctor_set(x_51, 7, x_90); +lean_ctor_set(x_51, 6, x_55); +lean_ctor_set(x_51, 3, x_54); +lean_ctor_set(x_51, 1, x_66); +lean_ctor_set(x_51, 0, x_53); +x_92 = lean_st_ref_set(x_3, x_51, x_52); +if (lean_obj_tag(x_49) == 0) { -lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; -x_90 = lean_ctor_get(x_89, 1); -lean_inc(x_90); -if (lean_is_exclusive(x_89)) { - lean_ctor_release(x_89, 0); - lean_ctor_release(x_89, 1); - x_91 = x_89; +lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; +x_93 = lean_ctor_get(x_92, 1); +lean_inc(x_93); +if (lean_is_exclusive(x_92)) { + lean_ctor_release(x_92, 0); + lean_ctor_release(x_92, 1); + x_94 = x_92; } else { - lean_dec_ref(x_89); - x_91 = lean_box(0); + lean_dec_ref(x_92); + x_94 = lean_box(0); } -x_92 = lean_ctor_get(x_48, 0); -lean_inc(x_92); -lean_dec(x_48); -if (lean_is_scalar(x_91)) { - x_93 = lean_alloc_ctor(1, 2, 0); +x_95 = lean_ctor_get(x_49, 0); +lean_inc(x_95); +lean_dec(x_49); +if (lean_is_scalar(x_94)) { + x_96 = lean_alloc_ctor(1, 2, 0); } else { - x_93 = x_91; - lean_ctor_set_tag(x_93, 1); + x_96 = x_94; + lean_ctor_set_tag(x_96, 1); } -lean_ctor_set(x_93, 0, x_92); -lean_ctor_set(x_93, 1, x_90); -return x_93; +lean_ctor_set(x_96, 0, x_95); +lean_ctor_set(x_96, 1, x_93); +return x_96; } else { -lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; -x_94 = lean_ctor_get(x_89, 1); -lean_inc(x_94); -if (lean_is_exclusive(x_89)) { - lean_ctor_release(x_89, 0); - lean_ctor_release(x_89, 1); - x_95 = x_89; +lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; +x_97 = lean_ctor_get(x_92, 1); +lean_inc(x_97); +if (lean_is_exclusive(x_92)) { + lean_ctor_release(x_92, 0); + lean_ctor_release(x_92, 1); + x_98 = x_92; } else { - lean_dec_ref(x_89); - x_95 = lean_box(0); + lean_dec_ref(x_92); + x_98 = lean_box(0); } -x_96 = lean_ctor_get(x_48, 0); -lean_inc(x_96); -lean_dec(x_48); -if (lean_is_scalar(x_95)) { - x_97 = lean_alloc_ctor(0, 2, 0); +x_99 = lean_ctor_get(x_49, 0); +lean_inc(x_99); +lean_dec(x_49); +if (lean_is_scalar(x_98)) { + x_100 = lean_alloc_ctor(0, 2, 0); } else { - x_97 = x_95; + x_100 = x_98; } -lean_ctor_set(x_97, 0, x_96); -lean_ctor_set(x_97, 1, x_94); -return x_97; +lean_ctor_set(x_100, 0, x_99); +lean_ctor_set(x_100, 1, x_97); +return x_100; } } } 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_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_98 = lean_ctor_get(x_50, 1); -x_99 = lean_ctor_get(x_50, 2); -x_100 = lean_ctor_get(x_50, 4); -x_101 = lean_ctor_get(x_50, 5); -x_102 = lean_ctor_get(x_50, 7); -x_103 = lean_ctor_get(x_50, 8); +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_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_101 = lean_ctor_get(x_51, 1); +x_102 = lean_ctor_get(x_51, 2); +x_103 = lean_ctor_get(x_51, 4); +x_104 = lean_ctor_get(x_51, 5); +x_105 = lean_ctor_get(x_51, 7); +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_dec(x_50); -x_104 = l_Lean_PersistentArray_append___rarg(x_98, x_56); -x_105 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore(x_2, x_104, x_55); -x_106 = lean_ctor_get_uint8(x_102, sizeof(void*)*2); -x_107 = lean_ctor_get(x_102, 0); -lean_inc(x_107); -x_108 = lean_ctor_get(x_102, 1); +lean_dec(x_51); +x_106 = l_Lean_PersistentArray_append___rarg(x_101, x_57); +x_107 = lean_ctor_get_uint8(x_105, sizeof(void*)*2); +x_108 = lean_ctor_get(x_105, 0); lean_inc(x_108); -if (lean_is_exclusive(x_102)) { - lean_ctor_release(x_102, 0); - lean_ctor_release(x_102, 1); - x_109 = x_102; +x_109 = lean_ctor_get(x_105, 1); +lean_inc(x_109); +if (lean_is_exclusive(x_105)) { + lean_ctor_release(x_105, 0); + lean_ctor_release(x_105, 1); + x_110 = x_105; } else { - lean_dec_ref(x_102); - x_109 = lean_box(0); + lean_dec_ref(x_105); + x_110 = lean_box(0); } -x_110 = lean_ctor_get(x_57, 1); -lean_inc(x_110); -lean_dec(x_57); -x_111 = l_Lean_PersistentArray_append___rarg(x_108, x_110); -if (lean_is_scalar(x_109)) { - x_112 = lean_alloc_ctor(0, 2, 1); -} else { - x_112 = x_109; -} -lean_ctor_set(x_112, 0, x_107); -lean_ctor_set(x_112, 1, x_111); -lean_ctor_set_uint8(x_112, sizeof(void*)*2, x_106); -x_113 = lean_alloc_ctor(0, 9, 0); -lean_ctor_set(x_113, 0, x_52); -lean_ctor_set(x_113, 1, x_105); -lean_ctor_set(x_113, 2, x_99); -lean_ctor_set(x_113, 3, x_53); -lean_ctor_set(x_113, 4, x_100); -lean_ctor_set(x_113, 5, x_101); -lean_ctor_set(x_113, 6, x_54); -lean_ctor_set(x_113, 7, x_112); -lean_ctor_set(x_113, 8, x_103); -x_114 = lean_st_ref_set(x_3, x_113, x_51); -if (lean_obj_tag(x_48) == 0) -{ -lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; -x_115 = lean_ctor_get(x_114, 1); -lean_inc(x_115); -if (lean_is_exclusive(x_114)) { - lean_ctor_release(x_114, 0); - lean_ctor_release(x_114, 1); - x_116 = x_114; +x_111 = lean_ctor_get(x_58, 1); +lean_inc(x_111); +lean_dec(x_58); +x_112 = l_Lean_PersistentArray_append___rarg(x_109, x_111); +if (lean_is_scalar(x_110)) { + x_113 = lean_alloc_ctor(0, 2, 1); } else { - lean_dec_ref(x_114); - x_116 = lean_box(0); -} -x_117 = lean_ctor_get(x_48, 0); + x_113 = x_110; +} +lean_ctor_set(x_113, 0, x_108); +lean_ctor_set(x_113, 1, x_112); +lean_ctor_set_uint8(x_113, sizeof(void*)*2, x_107); +x_114 = l_Lean_PersistentArray_mapM___at_Lean_Elab_Command_liftTermElabM___spec__12(x_2, x_56); +x_115 = lean_alloc_ctor(0, 9, 0); +lean_ctor_set(x_115, 0, x_53); +lean_ctor_set(x_115, 1, x_106); +lean_ctor_set(x_115, 2, x_102); +lean_ctor_set(x_115, 3, x_54); +lean_ctor_set(x_115, 4, x_103); +lean_ctor_set(x_115, 5, x_104); +lean_ctor_set(x_115, 6, x_55); +lean_ctor_set(x_115, 7, x_113); +lean_ctor_set(x_115, 8, x_114); +x_116 = lean_st_ref_set(x_3, x_115, x_52); +if (lean_obj_tag(x_49) == 0) +{ +lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; +x_117 = lean_ctor_get(x_116, 1); lean_inc(x_117); -lean_dec(x_48); -if (lean_is_scalar(x_116)) { - x_118 = lean_alloc_ctor(1, 2, 0); -} else { +if (lean_is_exclusive(x_116)) { + lean_ctor_release(x_116, 0); + lean_ctor_release(x_116, 1); x_118 = x_116; - lean_ctor_set_tag(x_118, 1); +} else { + lean_dec_ref(x_116); + x_118 = lean_box(0); +} +x_119 = lean_ctor_get(x_49, 0); +lean_inc(x_119); +lean_dec(x_49); +if (lean_is_scalar(x_118)) { + x_120 = lean_alloc_ctor(1, 2, 0); +} else { + x_120 = x_118; + lean_ctor_set_tag(x_120, 1); } -lean_ctor_set(x_118, 0, x_117); -lean_ctor_set(x_118, 1, x_115); -return x_118; +lean_ctor_set(x_120, 0, x_119); +lean_ctor_set(x_120, 1, x_117); +return x_120; } else { -lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; -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; +lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; +x_121 = lean_ctor_get(x_116, 1); +lean_inc(x_121); +if (lean_is_exclusive(x_116)) { + lean_ctor_release(x_116, 0); + lean_ctor_release(x_116, 1); + x_122 = x_116; } else { - lean_dec_ref(x_114); - x_120 = lean_box(0); + lean_dec_ref(x_116); + x_122 = lean_box(0); } -x_121 = lean_ctor_get(x_48, 0); -lean_inc(x_121); -lean_dec(x_48); -if (lean_is_scalar(x_120)) { - x_122 = lean_alloc_ctor(0, 2, 0); +x_123 = lean_ctor_get(x_49, 0); +lean_inc(x_123); +lean_dec(x_49); +if (lean_is_scalar(x_122)) { + x_124 = lean_alloc_ctor(0, 2, 0); } else { - x_122 = x_120; + x_124 = x_122; } -lean_ctor_set(x_122, 0, x_121); -lean_ctor_set(x_122, 1, x_119); -return x_122; +lean_ctor_set(x_124, 0, x_123); +lean_ctor_set(x_124, 1, x_121); +return x_124; } } } else { -uint8_t x_123; -lean_dec(x_37); -lean_dec(x_33); -lean_dec(x_2); -x_123 = !lean_is_exclusive(x_40); -if (x_123 == 0) +uint8_t x_125; +lean_dec(x_38); +lean_dec(x_34); +x_125 = !lean_is_exclusive(x_41); +if (x_125 == 0) { -return x_40; +return x_41; } else { -lean_object* x_124; lean_object* x_125; lean_object* x_126; -x_124 = lean_ctor_get(x_40, 0); -x_125 = lean_ctor_get(x_40, 1); -lean_inc(x_125); -lean_inc(x_124); -lean_dec(x_40); -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; +lean_object* x_126; lean_object* x_127; lean_object* x_128; +x_126 = lean_ctor_get(x_41, 0); +x_127 = lean_ctor_get(x_41, 1); +lean_inc(x_127); +lean_inc(x_126); +lean_dec(x_41); +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_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_127 = lean_ctor_get_uint8(x_15, sizeof(void*)*2); +uint8_t 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_129 = lean_ctor_get_uint8(x_15, sizeof(void*)*2); lean_dec(x_15); -x_128 = l_Lean_Elab_Command_State_infoState___default___closed__3; -x_129 = l_Lean_Elab_Command_State_messages___default___closed__3; -x_130 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_130, 0, x_128); -lean_ctor_set(x_130, 1, x_129); -lean_ctor_set_uint8(x_130, sizeof(void*)*2, x_127); -x_131 = l_Lean_Elab_Command_liftCoreM___rarg___closed__1; -x_132 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_132, 0, x_11); -lean_ctor_set(x_132, 1, x_13); -lean_ctor_set(x_132, 2, x_14); -lean_ctor_set(x_132, 3, x_129); -lean_ctor_set(x_132, 4, x_131); -lean_ctor_set(x_132, 5, x_129); -lean_ctor_set(x_132, 6, x_130); -x_133 = lean_st_mk_ref(x_132, x_10); -x_134 = lean_ctor_get(x_133, 0); -lean_inc(x_134); -x_135 = lean_ctor_get(x_133, 1); -lean_inc(x_135); -lean_dec(x_133); -x_136 = l_Lean_Elab_Command_liftTermElabM___rarg___closed__15; -x_137 = lean_st_mk_ref(x_136, x_135); -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 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkMetaContext; -lean_inc(x_134); -lean_inc(x_138); -x_141 = l_Lean_Elab_Term_TermElabM_run___rarg(x_18, x_19, x_23, x_140, x_138, x_24, x_134, x_139); -if (lean_obj_tag(x_141) == 0) +x_130 = l_Lean_Elab_Command_State_infoState___default___closed__3; +x_131 = l_Lean_Elab_Command_State_messages___default___closed__3; +x_132 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_132, 0, x_130); +lean_ctor_set(x_132, 1, x_131); +lean_ctor_set_uint8(x_132, sizeof(void*)*2, x_129); +x_133 = l_Lean_Elab_Command_liftCoreM___rarg___closed__1; +x_134 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_134, 0, x_11); +lean_ctor_set(x_134, 1, x_13); +lean_ctor_set(x_134, 2, x_14); +lean_ctor_set(x_134, 3, x_16); +lean_ctor_set(x_134, 4, x_133); +lean_ctor_set(x_134, 5, x_131); +lean_ctor_set(x_134, 6, x_132); +x_135 = lean_st_mk_ref(x_134, x_10); +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_138 = l_Lean_Elab_Command_liftTermElabM___rarg___closed__15; +x_139 = lean_st_mk_ref(x_138, x_137); +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 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkMetaContext; +lean_inc(x_136); +lean_inc(x_140); +x_143 = l_Lean_Elab_Term_TermElabM_run___rarg(x_19, x_20, x_24, x_142, x_140, x_25, x_136, x_141); +if (lean_obj_tag(x_143) == 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; 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; uint8_t 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_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_138, x_143); -lean_dec(x_138); -x_145 = lean_ctor_get(x_144, 1); +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; uint8_t 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; +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_144); -x_146 = lean_st_ref_get(x_134, x_145); -lean_dec(x_134); -x_147 = lean_ctor_get(x_146, 0); +lean_dec(x_143); +x_146 = lean_st_ref_get(x_140, x_145); +lean_dec(x_140); +x_147 = lean_ctor_get(x_146, 1); lean_inc(x_147); -x_148 = lean_ctor_get(x_146, 1); -lean_inc(x_148); lean_dec(x_146); -x_149 = lean_ctor_get(x_142, 0); +x_148 = lean_st_ref_get(x_136, x_147); +lean_dec(x_136); +x_149 = lean_ctor_get(x_148, 0); lean_inc(x_149); -lean_dec(x_142); -x_150 = lean_st_ref_take(x_3, x_148); -x_151 = lean_ctor_get(x_150, 0); +x_150 = lean_ctor_get(x_148, 1); +lean_inc(x_150); +lean_dec(x_148); +x_151 = lean_ctor_get(x_144, 0); lean_inc(x_151); -x_152 = lean_ctor_get(x_150, 1); -lean_inc(x_152); -lean_dec(x_150); -x_153 = lean_ctor_get(x_147, 0); +lean_dec(x_144); +x_152 = lean_st_ref_take(x_3, x_150); +x_153 = lean_ctor_get(x_152, 0); lean_inc(x_153); -x_154 = lean_ctor_get(x_147, 1); +x_154 = lean_ctor_get(x_152, 1); lean_inc(x_154); -x_155 = lean_ctor_get(x_147, 2); +lean_dec(x_152); +x_155 = lean_ctor_get(x_149, 0); lean_inc(x_155); -x_156 = lean_ctor_get(x_147, 3); +x_156 = lean_ctor_get(x_149, 1); lean_inc(x_156); -x_157 = lean_ctor_get(x_147, 5); +x_157 = lean_ctor_get(x_149, 2); lean_inc(x_157); -x_158 = lean_ctor_get(x_147, 6); +x_158 = lean_ctor_get(x_149, 3); lean_inc(x_158); -lean_dec(x_147); -x_159 = lean_ctor_get(x_151, 1); +x_159 = lean_ctor_get(x_149, 5); lean_inc(x_159); -x_160 = lean_ctor_get(x_151, 2); +x_160 = lean_ctor_get(x_149, 6); lean_inc(x_160); -x_161 = lean_ctor_get(x_151, 4); +lean_dec(x_149); +x_161 = lean_ctor_get(x_153, 1); lean_inc(x_161); -x_162 = lean_ctor_get(x_151, 5); +x_162 = lean_ctor_get(x_153, 2); lean_inc(x_162); -x_163 = lean_ctor_get(x_151, 7); +x_163 = lean_ctor_get(x_153, 4); lean_inc(x_163); -x_164 = lean_ctor_get(x_151, 8); +x_164 = lean_ctor_get(x_153, 5); lean_inc(x_164); -if (lean_is_exclusive(x_151)) { - lean_ctor_release(x_151, 0); - lean_ctor_release(x_151, 1); - lean_ctor_release(x_151, 2); - lean_ctor_release(x_151, 3); - lean_ctor_release(x_151, 4); - lean_ctor_release(x_151, 5); - lean_ctor_release(x_151, 6); - lean_ctor_release(x_151, 7); - lean_ctor_release(x_151, 8); - x_165 = x_151; -} else { - lean_dec_ref(x_151); - x_165 = lean_box(0); -} -x_166 = l_Lean_PersistentArray_append___rarg(x_159, x_157); -x_167 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore(x_2, x_166, x_156); -x_168 = lean_ctor_get_uint8(x_163, sizeof(void*)*2); -x_169 = lean_ctor_get(x_163, 0); +x_165 = lean_ctor_get(x_153, 7); +lean_inc(x_165); +if (lean_is_exclusive(x_153)) { + lean_ctor_release(x_153, 0); + lean_ctor_release(x_153, 1); + lean_ctor_release(x_153, 2); + lean_ctor_release(x_153, 3); + lean_ctor_release(x_153, 4); + lean_ctor_release(x_153, 5); + lean_ctor_release(x_153, 6); + lean_ctor_release(x_153, 7); + lean_ctor_release(x_153, 8); + x_166 = x_153; +} else { + lean_dec_ref(x_153); + x_166 = lean_box(0); +} +x_167 = l_Lean_PersistentArray_append___rarg(x_161, x_159); +x_168 = lean_ctor_get_uint8(x_165, sizeof(void*)*2); +x_169 = lean_ctor_get(x_165, 0); lean_inc(x_169); -x_170 = lean_ctor_get(x_163, 1); +x_170 = lean_ctor_get(x_165, 1); lean_inc(x_170); -if (lean_is_exclusive(x_163)) { - lean_ctor_release(x_163, 0); - lean_ctor_release(x_163, 1); - x_171 = x_163; +if (lean_is_exclusive(x_165)) { + lean_ctor_release(x_165, 0); + lean_ctor_release(x_165, 1); + x_171 = x_165; } else { - lean_dec_ref(x_163); + lean_dec_ref(x_165); x_171 = lean_box(0); } -x_172 = lean_ctor_get(x_158, 1); +x_172 = lean_ctor_get(x_160, 1); lean_inc(x_172); -lean_dec(x_158); +lean_dec(x_160); x_173 = l_Lean_PersistentArray_append___rarg(x_170, x_172); if (lean_is_scalar(x_171)) { x_174 = lean_alloc_ctor(0, 2, 1); @@ -19563,99 +23068,99 @@ if (lean_is_scalar(x_171)) { lean_ctor_set(x_174, 0, x_169); lean_ctor_set(x_174, 1, x_173); lean_ctor_set_uint8(x_174, sizeof(void*)*2, x_168); -if (lean_is_scalar(x_165)) { - x_175 = lean_alloc_ctor(0, 9, 0); -} else { - x_175 = x_165; -} -lean_ctor_set(x_175, 0, x_153); -lean_ctor_set(x_175, 1, x_167); -lean_ctor_set(x_175, 2, x_160); -lean_ctor_set(x_175, 3, x_154); -lean_ctor_set(x_175, 4, x_161); -lean_ctor_set(x_175, 5, x_162); -lean_ctor_set(x_175, 6, x_155); -lean_ctor_set(x_175, 7, x_174); -lean_ctor_set(x_175, 8, x_164); -x_176 = lean_st_ref_set(x_3, x_175, x_152); -if (lean_obj_tag(x_149) == 0) -{ -lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; -x_177 = lean_ctor_get(x_176, 1); -lean_inc(x_177); -if (lean_is_exclusive(x_176)) { - lean_ctor_release(x_176, 0); - lean_ctor_release(x_176, 1); - x_178 = x_176; +x_175 = l_Lean_PersistentArray_mapM___at_Lean_Elab_Command_liftTermElabM___spec__12(x_2, x_158); +if (lean_is_scalar(x_166)) { + x_176 = lean_alloc_ctor(0, 9, 0); +} else { + x_176 = x_166; +} +lean_ctor_set(x_176, 0, x_155); +lean_ctor_set(x_176, 1, x_167); +lean_ctor_set(x_176, 2, x_162); +lean_ctor_set(x_176, 3, x_156); +lean_ctor_set(x_176, 4, x_163); +lean_ctor_set(x_176, 5, x_164); +lean_ctor_set(x_176, 6, x_157); +lean_ctor_set(x_176, 7, x_174); +lean_ctor_set(x_176, 8, x_175); +x_177 = lean_st_ref_set(x_3, x_176, x_154); +if (lean_obj_tag(x_151) == 0) +{ +lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; +x_178 = lean_ctor_get(x_177, 1); +lean_inc(x_178); +if (lean_is_exclusive(x_177)) { + lean_ctor_release(x_177, 0); + lean_ctor_release(x_177, 1); + x_179 = x_177; } else { - lean_dec_ref(x_176); - x_178 = lean_box(0); + lean_dec_ref(x_177); + x_179 = lean_box(0); } -x_179 = lean_ctor_get(x_149, 0); -lean_inc(x_179); -lean_dec(x_149); -if (lean_is_scalar(x_178)) { - x_180 = lean_alloc_ctor(1, 2, 0); +x_180 = lean_ctor_get(x_151, 0); +lean_inc(x_180); +lean_dec(x_151); +if (lean_is_scalar(x_179)) { + x_181 = lean_alloc_ctor(1, 2, 0); } else { - x_180 = x_178; - lean_ctor_set_tag(x_180, 1); + x_181 = x_179; + lean_ctor_set_tag(x_181, 1); } -lean_ctor_set(x_180, 0, x_179); -lean_ctor_set(x_180, 1, x_177); -return x_180; +lean_ctor_set(x_181, 0, x_180); +lean_ctor_set(x_181, 1, x_178); +return x_181; } else { -lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; -x_181 = lean_ctor_get(x_176, 1); -lean_inc(x_181); -if (lean_is_exclusive(x_176)) { - lean_ctor_release(x_176, 0); - lean_ctor_release(x_176, 1); - x_182 = x_176; +lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; +x_182 = lean_ctor_get(x_177, 1); +lean_inc(x_182); +if (lean_is_exclusive(x_177)) { + lean_ctor_release(x_177, 0); + lean_ctor_release(x_177, 1); + x_183 = x_177; } else { - lean_dec_ref(x_176); - x_182 = lean_box(0); -} -x_183 = lean_ctor_get(x_149, 0); -lean_inc(x_183); -lean_dec(x_149); -if (lean_is_scalar(x_182)) { - x_184 = lean_alloc_ctor(0, 2, 0); + lean_dec_ref(x_177); + x_183 = lean_box(0); +} +x_184 = lean_ctor_get(x_151, 0); +lean_inc(x_184); +lean_dec(x_151); +if (lean_is_scalar(x_183)) { + x_185 = lean_alloc_ctor(0, 2, 0); } else { - x_184 = x_182; + x_185 = x_183; } -lean_ctor_set(x_184, 0, x_183); -lean_ctor_set(x_184, 1, x_181); -return x_184; +lean_ctor_set(x_185, 0, x_184); +lean_ctor_set(x_185, 1, x_182); +return x_185; } } else { -lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; -lean_dec(x_138); -lean_dec(x_134); -lean_dec(x_2); -x_185 = lean_ctor_get(x_141, 0); -lean_inc(x_185); -x_186 = lean_ctor_get(x_141, 1); +lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; +lean_dec(x_140); +lean_dec(x_136); +x_186 = lean_ctor_get(x_143, 0); lean_inc(x_186); -if (lean_is_exclusive(x_141)) { - lean_ctor_release(x_141, 0); - lean_ctor_release(x_141, 1); - x_187 = x_141; +x_187 = lean_ctor_get(x_143, 1); +lean_inc(x_187); +if (lean_is_exclusive(x_143)) { + lean_ctor_release(x_143, 0); + lean_ctor_release(x_143, 1); + x_188 = x_143; } else { - lean_dec_ref(x_141); - x_187 = lean_box(0); + lean_dec_ref(x_143); + x_188 = lean_box(0); } -if (lean_is_scalar(x_187)) { - x_188 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_188)) { + x_189 = lean_alloc_ctor(1, 2, 0); } else { - x_188 = x_187; + x_189 = x_188; } -lean_ctor_set(x_188, 0, x_185); -lean_ctor_set(x_188, 1, x_186); -return x_188; +lean_ctor_set(x_189, 0, x_186); +lean_ctor_set(x_189, 1, x_187); +return x_189; } } } @@ -19832,12 +23337,70 @@ lean_dec(x_3); return x_10; } } +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftTermElabM___spec__14___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_Lean_Elab_Command_liftTermElabM___spec__14(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_Command_liftTermElabM___spec__15___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_Lean_Elab_Command_liftTermElabM___spec__15(x_1, x_5, x_6, x_4); +lean_dec(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapMAux___at_Lean_Elab_Command_liftTermElabM___spec__13___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_PersistentArray_mapMAux___at_Lean_Elab_Command_liftTermElabM___spec__13(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftTermElabM___spec__16___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_Lean_Elab_Command_liftTermElabM___spec__16(x_1, x_5, x_6, x_4); +lean_dec(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapM___at_Lean_Elab_Command_liftTermElabM___spec__12___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_PersistentArray_mapM___at_Lean_Elab_Command_liftTermElabM___spec__12(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_Elab_Command_liftTermElabM___rarg___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_Command_liftTermElabM___rarg(x_1, x_2, x_3, x_4); lean_dec(x_3); +lean_dec(x_2); return x_5; } } @@ -20236,6 +23799,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_runTermElabM___rarg___boxed(lean_ob lean_object* x_5; x_5 = l_Lean_Elab_Command_runTermElabM___rarg(x_1, x_2, x_3, x_4); lean_dec(x_3); +lean_dec(x_2); return x_5; } } @@ -20317,6 +23881,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_lift lean_object* x_5; x_5 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_liftAttrM___rarg(x_1, x_2, x_3, x_4); lean_dec(x_3); +lean_dec(x_2); return x_5; } } @@ -20419,7 +23984,7 @@ static lean_object* _init_l_Lean_Elab_Command_modifyScope___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_Elab_Command_modifyScope___closed__1; x_2 = l_Lean_Elab_Command_modifyScope___closed__2; -x_3 = lean_unsigned_to_nat(465u); +x_3 = lean_unsigned_to_nat(470u); x_4 = lean_unsigned_to_nat(16u); x_5 = l_Lean_Elab_Command_modifyScope___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -25250,6 +28815,64 @@ l_Lean_Elab_Command_Context_currMacroScope___default = _init_l_Lean_Elab_Command lean_mark_persistent(l_Lean_Elab_Command_Context_currMacroScope___default); l_Lean_Elab_Command_Context_ref___default = _init_l_Lean_Elab_Command_Context_ref___default(); lean_mark_persistent(l_Lean_Elab_Command_Context_ref___default); +l___auto____x40_Lean_Elab_Command___hyg_448____closed__1 = _init_l___auto____x40_Lean_Elab_Command___hyg_448____closed__1(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_448____closed__1); +l___auto____x40_Lean_Elab_Command___hyg_448____closed__2 = _init_l___auto____x40_Lean_Elab_Command___hyg_448____closed__2(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_448____closed__2); +l___auto____x40_Lean_Elab_Command___hyg_448____closed__3 = _init_l___auto____x40_Lean_Elab_Command___hyg_448____closed__3(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_448____closed__3); +l___auto____x40_Lean_Elab_Command___hyg_448____closed__4 = _init_l___auto____x40_Lean_Elab_Command___hyg_448____closed__4(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_448____closed__4); +l___auto____x40_Lean_Elab_Command___hyg_448____closed__5 = _init_l___auto____x40_Lean_Elab_Command___hyg_448____closed__5(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_448____closed__5); +l___auto____x40_Lean_Elab_Command___hyg_448____closed__6 = _init_l___auto____x40_Lean_Elab_Command___hyg_448____closed__6(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_448____closed__6); +l___auto____x40_Lean_Elab_Command___hyg_448____closed__7 = _init_l___auto____x40_Lean_Elab_Command___hyg_448____closed__7(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_448____closed__7); +l___auto____x40_Lean_Elab_Command___hyg_448____closed__8 = _init_l___auto____x40_Lean_Elab_Command___hyg_448____closed__8(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_448____closed__8); +l___auto____x40_Lean_Elab_Command___hyg_448____closed__9 = _init_l___auto____x40_Lean_Elab_Command___hyg_448____closed__9(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_448____closed__9); +l___auto____x40_Lean_Elab_Command___hyg_448____closed__10 = _init_l___auto____x40_Lean_Elab_Command___hyg_448____closed__10(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_448____closed__10); +l___auto____x40_Lean_Elab_Command___hyg_448____closed__11 = _init_l___auto____x40_Lean_Elab_Command___hyg_448____closed__11(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_448____closed__11); +l___auto____x40_Lean_Elab_Command___hyg_448____closed__12 = _init_l___auto____x40_Lean_Elab_Command___hyg_448____closed__12(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_448____closed__12); +l___auto____x40_Lean_Elab_Command___hyg_448____closed__13 = _init_l___auto____x40_Lean_Elab_Command___hyg_448____closed__13(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_448____closed__13); +l___auto____x40_Lean_Elab_Command___hyg_448____closed__14 = _init_l___auto____x40_Lean_Elab_Command___hyg_448____closed__14(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_448____closed__14); +l___auto____x40_Lean_Elab_Command___hyg_448____closed__15 = _init_l___auto____x40_Lean_Elab_Command___hyg_448____closed__15(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_448____closed__15); +l___auto____x40_Lean_Elab_Command___hyg_448____closed__16 = _init_l___auto____x40_Lean_Elab_Command___hyg_448____closed__16(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_448____closed__16); +l___auto____x40_Lean_Elab_Command___hyg_448____closed__17 = _init_l___auto____x40_Lean_Elab_Command___hyg_448____closed__17(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_448____closed__17); +l___auto____x40_Lean_Elab_Command___hyg_448____closed__18 = _init_l___auto____x40_Lean_Elab_Command___hyg_448____closed__18(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_448____closed__18); +l___auto____x40_Lean_Elab_Command___hyg_448____closed__19 = _init_l___auto____x40_Lean_Elab_Command___hyg_448____closed__19(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_448____closed__19); +l___auto____x40_Lean_Elab_Command___hyg_448____closed__20 = _init_l___auto____x40_Lean_Elab_Command___hyg_448____closed__20(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_448____closed__20); +l___auto____x40_Lean_Elab_Command___hyg_448____closed__21 = _init_l___auto____x40_Lean_Elab_Command___hyg_448____closed__21(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_448____closed__21); +l___auto____x40_Lean_Elab_Command___hyg_448____closed__22 = _init_l___auto____x40_Lean_Elab_Command___hyg_448____closed__22(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_448____closed__22); +l___auto____x40_Lean_Elab_Command___hyg_448____closed__23 = _init_l___auto____x40_Lean_Elab_Command___hyg_448____closed__23(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_448____closed__23); +l___auto____x40_Lean_Elab_Command___hyg_448____closed__24 = _init_l___auto____x40_Lean_Elab_Command___hyg_448____closed__24(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_448____closed__24); +l___auto____x40_Lean_Elab_Command___hyg_448____closed__25 = _init_l___auto____x40_Lean_Elab_Command___hyg_448____closed__25(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_448____closed__25); +l___auto____x40_Lean_Elab_Command___hyg_448____closed__26 = _init_l___auto____x40_Lean_Elab_Command___hyg_448____closed__26(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_448____closed__26); +l___auto____x40_Lean_Elab_Command___hyg_448____closed__27 = _init_l___auto____x40_Lean_Elab_Command___hyg_448____closed__27(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_448____closed__27); +l___auto____x40_Lean_Elab_Command___hyg_448____closed__28 = _init_l___auto____x40_Lean_Elab_Command___hyg_448____closed__28(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_448____closed__28); +l___auto____x40_Lean_Elab_Command___hyg_448_ = _init_l___auto____x40_Lean_Elab_Command___hyg_448_(); +lean_mark_persistent(l___auto____x40_Lean_Elab_Command___hyg_448_); l_Lean_Elab_Command_instMonadCommandElabM___closed__1 = _init_l_Lean_Elab_Command_instMonadCommandElabM___closed__1(); lean_mark_persistent(l_Lean_Elab_Command_instMonadCommandElabM___closed__1); l_Lean_Elab_Command_instMonadCommandElabM___closed__2 = _init_l_Lean_Elab_Command_instMonadCommandElabM___closed__2(); @@ -25268,12 +28891,49 @@ 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_677_(lean_io_mk_world()); +if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_734_(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); lean_dec_ref(res); -}l_Lean_Elab_Command_addLinter___closed__1 = _init_l_Lean_Elab_Command_addLinter___closed__1(); +}l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__1(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__1); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__2(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__2); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__3 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__3(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__3); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__4 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__4(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__4); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__5 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__5(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__5); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__6 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__6(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__6); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__7 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__7(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__7); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__8 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__8(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__8); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__9 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__9(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__9); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__10 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__10(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__10); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__11 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__11(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__11); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__12 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__12(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__12); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__13 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__13(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__13); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__14 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__14(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__14); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__15 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__15(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__15); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__16 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__16(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__16); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__17 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__17(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771____closed__17); +res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_771_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Lean_Elab_Command_addLinter___closed__1 = _init_l_Lean_Elab_Command_addLinter___closed__1(); lean_mark_persistent(l_Lean_Elab_Command_addLinter___closed__1); l_Lean_Elab_Command_instMonadInfoTreeCommandElabM___closed__1 = _init_l_Lean_Elab_Command_instMonadInfoTreeCommandElabM___closed__1(); lean_mark_persistent(l_Lean_Elab_Command_instMonadInfoTreeCommandElabM___closed__1); @@ -25387,10 +29047,33 @@ l_Lean_Elab_logException___at_Lean_Elab_Command_runLinters___spec__1___closed__2 lean_mark_persistent(l_Lean_Elab_logException___at_Lean_Elab_Command_runLinters___spec__1___closed__2); l_Lean_Elab_logException___at_Lean_Elab_Command_runLinters___spec__1___closed__3 = _init_l_Lean_Elab_logException___at_Lean_Elab_Command_runLinters___spec__1___closed__3(); lean_mark_persistent(l_Lean_Elab_logException___at_Lean_Elab_Command_runLinters___spec__1___closed__3); -l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__4___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__4___closed__1(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__4___closed__1); +l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__5___closed__1 = _init_l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__5___closed__1(); +lean_mark_persistent(l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__5___closed__1); +l_Lean_withSeconds___at_Lean_Elab_Command_runLinters___spec__7___closed__1 = _init_l_Lean_withSeconds___at_Lean_Elab_Command_runLinters___spec__7___closed__1(); +l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3___closed__1 = _init_l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3___closed__1); +l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3___closed__2 = _init_l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3___closed__2); +l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3___closed__3 = _init_l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3___closed__3(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3___closed__3); +l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3___closed__4 = _init_l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3___closed__4(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3___closed__4); +l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3___closed__5 = _init_l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3___closed__5(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3___closed__5); +l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__4___closed__1 = _init_l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__4___closed__1); +l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__11___lambda__1___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__11___lambda__1___closed__1(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__11___lambda__1___closed__1); +l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__11___lambda__1___closed__2 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__11___lambda__1___closed__2(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__11___lambda__1___closed__2); +l_Lean_Elab_Command_runLinters___lambda__1___closed__1 = _init_l_Lean_Elab_Command_runLinters___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Elab_Command_runLinters___lambda__1___closed__1); +l_Lean_Elab_Command_runLinters___lambda__1___closed__2 = _init_l_Lean_Elab_Command_runLinters___lambda__1___closed__2(); +lean_mark_persistent(l_Lean_Elab_Command_runLinters___lambda__1___closed__2); l_Lean_Elab_Command_runLinters___closed__1 = _init_l_Lean_Elab_Command_runLinters___closed__1(); lean_mark_persistent(l_Lean_Elab_Command_runLinters___closed__1); +l_Lean_Elab_Command_runLinters___closed__2 = _init_l_Lean_Elab_Command_runLinters___closed__2(); +lean_mark_persistent(l_Lean_Elab_Command_runLinters___closed__2); l_Lean_Elab_Command_instMonadQuotationCommandElabM___closed__1 = _init_l_Lean_Elab_Command_instMonadQuotationCommandElabM___closed__1(); lean_mark_persistent(l_Lean_Elab_Command_instMonadQuotationCommandElabM___closed__1); l_Lean_Elab_Command_instMonadQuotationCommandElabM___closed__2 = _init_l_Lean_Elab_Command_instMonadQuotationCommandElabM___closed__2(); @@ -25417,19 +29100,11 @@ l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__7 = _init_l_Lean_Elab lean_mark_persistent(l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__7); l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__8 = _init_l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__8(); lean_mark_persistent(l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__8); -l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__9 = _init_l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__9(); -lean_mark_persistent(l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__9); -l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__10 = _init_l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__10(); -lean_mark_persistent(l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__10); -l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__11 = _init_l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__11(); -lean_mark_persistent(l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__11); -l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__12 = _init_l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__12(); -lean_mark_persistent(l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__12); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3670____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3670____closed__1(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3670____closed__1); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3670____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3670____closed__2(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3670____closed__2); -if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3670_(lean_io_mk_world()); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3931____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3931____closed__1(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3931____closed__1); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3931____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3931____closed__2(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3931____closed__2); +if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3931_(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); @@ -25462,70 +29137,44 @@ 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_4380____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4380____closed__1(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4380____closed__1); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4380____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4380____closed__2(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4380____closed__2); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4380____closed__3 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4380____closed__3(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4380____closed__3); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4380____closed__4 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4380____closed__4(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4380____closed__4); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4380____closed__5 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4380____closed__5(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4380____closed__5); -if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4380_(lean_io_mk_world()); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4641____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4641____closed__1(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4641____closed__1); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4641____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4641____closed__2(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4641____closed__2); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4641____closed__3 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4641____closed__3(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4641____closed__3); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4641____closed__4 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4641____closed__4(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4641____closed__4); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4641____closed__5 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4641____closed__5(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4641____closed__5); +if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4641_(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); lean_dec_ref(res); -}l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__1(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__1); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__2(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__2); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__3 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__3(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__3); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__4 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__4(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__4); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__5 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__5(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__5); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__6 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__6(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__6); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__7 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__7(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__7); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__8 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__8(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__8); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__9 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__9(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__9); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__10 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__10(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__10); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__11 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__11(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__11); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__12 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__12(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__12); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__13 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__13(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__13); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__14 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__14(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417____closed__14); -res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4417_(lean_io_mk_world()); +}l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4678____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4678____closed__1(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4678____closed__1); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4678____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4678____closed__2(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4678____closed__2); +res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4678_(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(); lean_mark_persistent(l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabCommand___spec__1___closed__1); l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabCommand___spec__1___closed__2 = _init_l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabCommand___spec__1___closed__2(); lean_mark_persistent(l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabCommand___spec__1___closed__2); -l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_elabCommand___spec__7___closed__1 = _init_l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_elabCommand___spec__7___closed__1(); -lean_mark_persistent(l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_elabCommand___spec__7___closed__1); -l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__13___rarg___closed__1 = _init_l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__13___rarg___closed__1(); -lean_mark_persistent(l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__13___rarg___closed__1); -l_Lean_Elab_Command_elabCommand___lambda__3___closed__1 = _init_l_Lean_Elab_Command_elabCommand___lambda__3___closed__1(); -lean_mark_persistent(l_Lean_Elab_Command_elabCommand___lambda__3___closed__1); -l_Lean_Elab_Command_elabCommand___lambda__3___closed__2 = _init_l_Lean_Elab_Command_elabCommand___lambda__3___closed__2(); -lean_mark_persistent(l_Lean_Elab_Command_elabCommand___lambda__3___closed__2); -l_Lean_Elab_Command_elabCommand___lambda__3___closed__3 = _init_l_Lean_Elab_Command_elabCommand___lambda__3___closed__3(); -lean_mark_persistent(l_Lean_Elab_Command_elabCommand___lambda__3___closed__3); -l_Lean_Elab_Command_elabCommand___lambda__3___closed__4 = _init_l_Lean_Elab_Command_elabCommand___lambda__3___closed__4(); -lean_mark_persistent(l_Lean_Elab_Command_elabCommand___lambda__3___closed__4); -l_Lean_Elab_Command_elabCommand___lambda__3___closed__5 = _init_l_Lean_Elab_Command_elabCommand___lambda__3___closed__5(); -lean_mark_persistent(l_Lean_Elab_Command_elabCommand___lambda__3___closed__5); +l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__12___rarg___closed__1 = _init_l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__12___rarg___closed__1(); +lean_mark_persistent(l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__12___rarg___closed__1); +l_Lean_Elab_Command_elabCommand___lambda__4___closed__1 = _init_l_Lean_Elab_Command_elabCommand___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_Elab_Command_elabCommand___lambda__4___closed__1); +l_Lean_Elab_Command_elabCommand___lambda__4___closed__2 = _init_l_Lean_Elab_Command_elabCommand___lambda__4___closed__2(); +lean_mark_persistent(l_Lean_Elab_Command_elabCommand___lambda__4___closed__2); +l_Lean_Elab_Command_elabCommand___lambda__4___closed__3 = _init_l_Lean_Elab_Command_elabCommand___lambda__4___closed__3(); +lean_mark_persistent(l_Lean_Elab_Command_elabCommand___lambda__4___closed__3); +l_Lean_Elab_Command_elabCommand___lambda__4___closed__4 = _init_l_Lean_Elab_Command_elabCommand___lambda__4___closed__4(); +lean_mark_persistent(l_Lean_Elab_Command_elabCommand___lambda__4___closed__4); +l_Lean_Elab_Command_elabCommand___lambda__4___closed__5 = _init_l_Lean_Elab_Command_elabCommand___lambda__4___closed__5(); +lean_mark_persistent(l_Lean_Elab_Command_elabCommand___lambda__4___closed__5); l_Lean_Elab_Command_elabCommand___closed__1 = _init_l_Lean_Elab_Command_elabCommand___closed__1(); lean_mark_persistent(l_Lean_Elab_Command_elabCommand___closed__1); l_Lean_Elab_Command_elabCommand___closed__2 = _init_l_Lean_Elab_Command_elabCommand___closed__2(); @@ -25534,25 +29183,23 @@ 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_4918____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4918____closed__1(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4918____closed__1); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4918____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4918____closed__2(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4918____closed__2); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4918____closed__3 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4918____closed__3(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4918____closed__3); -res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_4918_(lean_io_mk_world()); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_5179____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_5179____closed__1(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_5179____closed__1); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_5179____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_5179____closed__2(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_5179____closed__2); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_5179____closed__3 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_5179____closed__3(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_5179____closed__3); +res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_5179_(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__6___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCommandTopLevel___spec__6___closed__1(); lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCommandTopLevel___spec__6___closed__1); l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCommandTopLevel___spec__6___closed__2 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCommandTopLevel___spec__6___closed__2(); lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCommandTopLevel___spec__6___closed__2); +l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCommandTopLevel___spec__6___closed__3 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCommandTopLevel___spec__6___closed__3(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCommandTopLevel___spec__6___closed__3); l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommandTopLevel___spec__12___lambda__1___closed__1 = _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommandTopLevel___spec__12___lambda__1___closed__1(); lean_mark_persistent(l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommandTopLevel___spec__12___lambda__1___closed__1); l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommandTopLevel___spec__12___lambda__1___closed__2 = _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommandTopLevel___spec__12___lambda__1___closed__2(); @@ -25565,14 +29212,12 @@ l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommandTopLevel___spec__12_ lean_mark_persistent(l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommandTopLevel___spec__12___lambda__1___closed__5); l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommandTopLevel___spec__12___lambda__1___closed__6 = _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommandTopLevel___spec__12___lambda__1___closed__6(); lean_mark_persistent(l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommandTopLevel___spec__12___lambda__1___closed__6); -l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommandTopLevel___spec__12___lambda__1___closed__7 = _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommandTopLevel___spec__12___lambda__1___closed__7(); -lean_mark_persistent(l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommandTopLevel___spec__12___lambda__1___closed__7); l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommandTopLevel___spec__12___closed__1 = _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommandTopLevel___spec__12___closed__1(); lean_mark_persistent(l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommandTopLevel___spec__12___closed__1); l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at_Lean_Elab_Command_elabCommandTopLevel___spec__9___closed__1 = _init_l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at_Lean_Elab_Command_elabCommandTopLevel___spec__9___closed__1(); lean_mark_persistent(l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at_Lean_Elab_Command_elabCommandTopLevel___spec__9___closed__1); -l_Lean_Elab_Command_elabCommandTopLevel___lambda__2___closed__1 = _init_l_Lean_Elab_Command_elabCommandTopLevel___lambda__2___closed__1(); -lean_mark_persistent(l_Lean_Elab_Command_elabCommandTopLevel___lambda__2___closed__1); +l_Lean_Elab_Command_elabCommandTopLevel___closed__1 = _init_l_Lean_Elab_Command_elabCommandTopLevel___closed__1(); +lean_mark_persistent(l_Lean_Elab_Command_elabCommandTopLevel___closed__1); l_Lean_Elab_Command_instInhabitedCommandElabM___closed__1 = _init_l_Lean_Elab_Command_instInhabitedCommandElabM___closed__1(); lean_mark_persistent(l_Lean_Elab_Command_instInhabitedCommandElabM___closed__1); l_Lean_Elab_Command_instInhabitedCommandElabM___closed__2 = _init_l_Lean_Elab_Command_instInhabitedCommandElabM___closed__2(); @@ -25599,8 +29244,6 @@ l_Lean_Elab_Command_getBracketedBinderIds___closed__7 = _init_l_Lean_Elab_Comman lean_mark_persistent(l_Lean_Elab_Command_getBracketedBinderIds___closed__7); l_Lean_Elab_Command_getBracketedBinderIds___closed__8 = _init_l_Lean_Elab_Command_getBracketedBinderIds___closed__8(); lean_mark_persistent(l_Lean_Elab_Command_getBracketedBinderIds___closed__8); -l_Lean_Elab_Command_getBracketedBinderIds___closed__9 = _init_l_Lean_Elab_Command_getBracketedBinderIds___closed__9(); -lean_mark_persistent(l_Lean_Elab_Command_getBracketedBinderIds___closed__9); l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkTermContext___closed__1 = _init_l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkTermContext___closed__1(); lean_mark_persistent(l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkTermContext___closed__1); l_Lean_Elab_Command_liftTermElabM___rarg___closed__1 = _init_l_Lean_Elab_Command_liftTermElabM___rarg___closed__1(); diff --git a/stage0/stdlib/Lean/Elab/ComputedFields.c b/stage0/stdlib/Lean/Elab/ComputedFields.c index e283834464b3..5372f49d9f57 100644 --- a/stage0/stdlib/Lean/Elab/ComputedFields.c +++ b/stage0/stdlib/Lean/Elab/ComputedFields.c @@ -14,7 +14,6 @@ extern "C" { #endif lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); -uint8_t l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_mkCasesOn___at_Lean_Elab_ComputedFields_overrideCasesOn___spec__5___closed__6; lean_object* l___private_Init_Util_0__outOfBounds___rarg(lean_object*); lean_object* l_Lean_getConstInfoInduct___at___private_Lean_Meta_RecursorInfo_0__Lean_Meta_mkRecursorInfoForKernelRec___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -202,6 +201,7 @@ extern lean_object* l_Lean_warningAsError; uint8_t l_Lean_isExtern(lean_object*, lean_object*); static lean_object* l_Lean_setImplementedBy___at_Lean_Elab_ComputedFields_overrideCasesOn___spec__12___closed__1; static lean_object* l_Lean_getConstInfoDefn___at_Lean_Elab_ComputedFields_overrideCasesOn___spec__1___closed__2; +uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); lean_object* l_Lean_Meta_forallTelescope___at___private_Lean_Meta_InferType_0__Lean_Meta_inferForallType___spec__2___rarg(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*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls___at_Lean_Elab_ComputedFields_mkComputedFieldOverrides___spec__4___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -10187,7 +10187,7 @@ else { lean_object* x_263; uint8_t x_264; x_263 = l_Lean_logAt___at_Lean_Elab_ComputedFields_setComputedFields___spec__2___closed__1; -x_264 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_260, x_263); +x_264 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_260, x_263); if (x_264 == 0) { x_10 = x_3; diff --git a/stage0/stdlib/Lean/Elab/Declaration.c b/stage0/stdlib/Lean/Elab/Declaration.c index 219e00cc6e46..c34de821258f 100644 --- a/stage0/stdlib/Lean/Elab/Declaration.c +++ b/stage0/stdlib/Lean/Elab/Declaration.c @@ -443,6 +443,7 @@ static lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_setDe uint8_t l_Lean_isAttribute(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_expandMutualPreamble_declRange___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Command_checkValidCtorModifier___at___private_Lean_Elab_Declaration_0__Lean_Elab_Command_inductiveSyntaxToView___spec__1___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_List_forM___at_Lean_Elab_Command_elabCommand___spec__8(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_expandMutualElement_declRange(lean_object*); static lean_object* l_Lean_Elab_Command_elabMutual___closed__4; lean_object* l_Lean_Syntax_node7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -513,7 +514,6 @@ lean_object* lean_array_uget(lean_object*, size_t); static lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isNamedDef___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_inductiveSyntaxToView(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabDeclaration___spec__4___closed__2; -lean_object* l_List_forM___at_Lean_Elab_Command_elabCommand___spec__9(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_isMutualPreambleCommand___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Declaration_0__Lean_Elab_Command_ensureValidNamespace___boxed(lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_instInhabitedSyntax; @@ -2752,6 +2752,7 @@ lean_closure_set(x_23, 6, x_2); lean_closure_set(x_23, 7, x_7); x_24 = l_Lean_Elab_Command_runTermElabM___rarg(x_23, x_3, x_4, x_22); lean_dec(x_4); +lean_dec(x_3); return x_24; } else @@ -5122,7 +5123,7 @@ x_53 = lean_ctor_get(x_44, 1); lean_inc(x_53); lean_dec(x_44); x_54 = l_List_reverse___rarg(x_53); -x_55 = l_List_forM___at_Lean_Elab_Command_elabCommand___spec__9(x_54, x_2, x_3, x_52); +x_55 = l_List_forM___at_Lean_Elab_Command_elabCommand___spec__8(x_54, x_2, x_3, x_52); lean_dec(x_3); lean_dec(x_2); x_56 = !lean_is_exclusive(x_55); @@ -5184,7 +5185,7 @@ x_71 = lean_ctor_get(x_44, 1); lean_inc(x_71); lean_dec(x_44); x_72 = l_List_reverse___rarg(x_71); -x_73 = l_List_forM___at_Lean_Elab_Command_elabCommand___spec__9(x_72, x_2, x_3, x_70); +x_73 = l_List_forM___at_Lean_Elab_Command_elabCommand___spec__8(x_72, x_2, x_3, x_70); lean_dec(x_3); lean_dec(x_2); x_74 = lean_ctor_get(x_73, 1); @@ -9482,6 +9483,7 @@ lean_ctor_set(x_28, 5, x_26); lean_ctor_set(x_28, 6, x_20); lean_ctor_set(x_28, 7, x_27); x_29 = l_Lean_Elab_Command_liftTermElabM___rarg(x_16, x_28, x_9, x_19); +lean_dec(x_28); if (lean_obj_tag(x_29) == 0) { lean_object* x_30; size_t x_31; size_t x_32; lean_object* x_33; diff --git a/stage0/stdlib/Lean/Elab/DefView.c b/stage0/stdlib/Lean/Elab/DefView.c index a490aeae8754..737cbfef5caf 100644 --- a/stage0/stdlib/Lean/Elab/DefView.c +++ b/stage0/stdlib/Lean/Elab/DefView.c @@ -213,6 +213,7 @@ lean_object* l_Lean_Elab_Term_elabBinders___rarg(lean_object*, lean_object*, lea static lean_object* l_Lean_Elab_Command_mkDefViewOfInstance___closed__8; static lean_object* l_Lean_Elab_Command_mkDefView___closed__2; static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_mkInstanceName___spec__17___closed__2; +lean_object* l_List_forM___at_Lean_Elab_Command_elabCommand___spec__8(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Elab_DefView_0__Lean_Elab_beqDefKind____x40_Lean_Elab_DefView___hyg_23_(uint8_t, uint8_t); static lean_object* l_Lean_Elab_Command_isDefLike___closed__8; lean_object* l_Lean_Syntax_node1(lean_object*, lean_object*, lean_object*); @@ -249,7 +250,6 @@ static lean_object* l_Lean_Elab_Command_mkDefViewOfOpaque___closed__2; lean_object* lean_array_uget(lean_object*, size_t); static lean_object* l_Lean_Elab_Command_mkDefViewOfInstance___closed__5; lean_object* l_Lean_Elab_mkUnusedBaseName(lean_object*, lean_object*, lean_object*); -lean_object* l_List_forM___at_Lean_Elab_Command_elabCommand___spec__9(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkFreshInstanceName(lean_object*); @@ -3653,7 +3653,7 @@ x_53 = lean_ctor_get(x_44, 1); lean_inc(x_53); lean_dec(x_44); x_54 = l_List_reverse___rarg(x_53); -x_55 = l_List_forM___at_Lean_Elab_Command_elabCommand___spec__9(x_54, x_2, x_3, x_52); +x_55 = l_List_forM___at_Lean_Elab_Command_elabCommand___spec__8(x_54, x_2, x_3, x_52); lean_dec(x_3); lean_dec(x_2); x_56 = !lean_is_exclusive(x_55); @@ -3715,7 +3715,7 @@ x_71 = lean_ctor_get(x_44, 1); lean_inc(x_71); lean_dec(x_44); x_72 = l_List_reverse___rarg(x_71); -x_73 = l_List_forM___at_Lean_Elab_Command_elabCommand___spec__9(x_72, x_2, x_3, x_70); +x_73 = l_List_forM___at_Lean_Elab_Command_elabCommand___spec__8(x_72, x_2, x_3, x_70); lean_dec(x_3); lean_dec(x_2); x_74 = lean_ctor_get(x_73, 1); @@ -4235,7 +4235,6 @@ lean_dec(x_6); x_9 = lean_alloc_closure((void*)(l_Lean_Elab_Command_mkInstanceName___lambda__4___boxed), 10, 2); lean_closure_set(x_9, 0, x_2); lean_closure_set(x_9, 1, x_1); -lean_inc(x_3); x_10 = l_Lean_Elab_Command_runTermElabM___rarg(x_9, x_3, x_4, x_8); if (lean_obj_tag(x_10) == 0) { @@ -4748,7 +4747,7 @@ x_53 = lean_ctor_get(x_44, 1); lean_inc(x_53); lean_dec(x_44); x_54 = l_List_reverse___rarg(x_53); -x_55 = l_List_forM___at_Lean_Elab_Command_elabCommand___spec__9(x_54, x_2, x_3, x_52); +x_55 = l_List_forM___at_Lean_Elab_Command_elabCommand___spec__8(x_54, x_2, x_3, x_52); lean_dec(x_3); lean_dec(x_2); x_56 = !lean_is_exclusive(x_55); @@ -4810,7 +4809,7 @@ x_71 = lean_ctor_get(x_44, 1); lean_inc(x_71); lean_dec(x_44); x_72 = l_List_reverse___rarg(x_71); -x_73 = l_List_forM___at_Lean_Elab_Command_elabCommand___spec__9(x_72, x_2, x_3, x_70); +x_73 = l_List_forM___at_Lean_Elab_Command_elabCommand___spec__8(x_72, x_2, x_3, x_70); lean_dec(x_3); lean_dec(x_2); x_74 = lean_ctor_get(x_73, 1); @@ -5157,7 +5156,7 @@ x_53 = lean_ctor_get(x_44, 1); lean_inc(x_53); lean_dec(x_44); x_54 = l_List_reverse___rarg(x_53); -x_55 = l_List_forM___at_Lean_Elab_Command_elabCommand___spec__9(x_54, x_2, x_3, x_52); +x_55 = l_List_forM___at_Lean_Elab_Command_elabCommand___spec__8(x_54, x_2, x_3, x_52); lean_dec(x_3); lean_dec(x_2); x_56 = !lean_is_exclusive(x_55); @@ -5219,7 +5218,7 @@ x_71 = lean_ctor_get(x_44, 1); lean_inc(x_71); lean_dec(x_44); x_72 = l_List_reverse___rarg(x_71); -x_73 = l_List_forM___at_Lean_Elab_Command_elabCommand___spec__9(x_72, x_2, x_3, x_70); +x_73 = l_List_forM___at_Lean_Elab_Command_elabCommand___spec__8(x_72, x_2, x_3, x_70); lean_dec(x_3); lean_dec(x_2); x_74 = lean_ctor_get(x_73, 1); diff --git a/stage0/stdlib/Lean/Elab/Deriving/BEq.c b/stage0/stdlib/Lean/Elab/Deriving/BEq.c index 34dd4ac022fc..2b4b211423cc 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/BEq.c +++ b/stage0/stdlib/Lean/Elab/Deriving/BEq.c @@ -38,6 +38,7 @@ static lean_object* l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__16; LEAN_EXPORT lean_object* l_List_allM___at_Lean_Elab_Deriving_BEq_mkBEqInstanceHandler___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkBEqEnumFun___closed__4; static lean_object* l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed__17; +lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__18(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkBEqInstanceCmds___closed__2; static lean_object* l_Lean_Elab_Deriving_BEq_mkMatch___closed__2; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -176,7 +177,6 @@ lean_object* l_Lean_MessageData_ofExpr(lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___closed__6; static lean_object* l_Lean_Elab_Deriving_BEq_initFn____x40_Lean_Elab_Deriving_BEq___hyg_3716____closed__5; LEAN_EXPORT lean_object* l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkBEqInstanceCmds___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__19(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_BEq_mkMatch_mkElseAlt___closed__8; lean_object* l_Lean_Syntax_node4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkBEqEnumFun___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -5914,7 +5914,6 @@ lean_dec(x_5); lean_dec(x_1); x_97 = lean_alloc_closure((void*)(l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkBEqEnumCmd), 8, 1); lean_closure_set(x_97, 0, x_8); -lean_inc(x_2); x_98 = l_Lean_Elab_Command_liftTermElabM___rarg(x_97, x_2, x_3, x_11); if (lean_obj_tag(x_98) == 0) { @@ -5963,7 +5962,7 @@ x_109 = 0; x_110 = lean_usize_of_nat(x_102); lean_dec(x_102); x_111 = lean_box(0); -x_112 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__19(x_100, x_109, x_110, x_111, x_2, x_3, x_101); +x_112 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__18(x_100, x_109, x_110, x_111, x_2, x_3, x_101); lean_dec(x_100); if (lean_obj_tag(x_112) == 0) { @@ -6067,7 +6066,7 @@ x_136 = 0; x_137 = lean_usize_of_nat(x_127); lean_dec(x_127); x_138 = lean_box(0); -x_139 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__19(x_125, x_136, x_137, x_138, x_2, x_3, x_126); +x_139 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__18(x_125, x_136, x_137, x_138, x_2, x_3, x_126); lean_dec(x_125); if (lean_obj_tag(x_139) == 0) { @@ -6252,7 +6251,6 @@ lean_object* x_22; lean_object* x_23; lean_dec(x_12); x_22 = lean_alloc_closure((void*)(l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkBEqInstanceCmds___boxed), 8, 1); lean_closure_set(x_22, 0, x_1); -lean_inc(x_2); x_23 = l_Lean_Elab_Command_liftTermElabM___rarg(x_22, x_2, x_3, x_15); if (lean_obj_tag(x_23) == 0) { @@ -6301,7 +6299,7 @@ x_34 = 0; x_35 = lean_usize_of_nat(x_27); lean_dec(x_27); x_36 = lean_box(0); -x_37 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__19(x_25, x_34, x_35, x_36, x_2, x_3, x_26); +x_37 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__18(x_25, x_34, x_35, x_36, x_2, x_3, x_26); lean_dec(x_25); if (lean_obj_tag(x_37) == 0) { @@ -6405,7 +6403,7 @@ x_61 = 0; x_62 = lean_usize_of_nat(x_52); lean_dec(x_52); x_63 = lean_box(0); -x_64 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__19(x_50, x_61, x_62, x_63, x_2, x_3, x_51); +x_64 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__18(x_50, x_61, x_62, x_63, x_2, x_3, x_51); lean_dec(x_50); if (lean_obj_tag(x_64) == 0) { diff --git a/stage0/stdlib/Lean/Elab/Deriving/Basic.c b/stage0/stdlib/Lean/Elab/Deriving/Basic.c index 8a426cd4a1d4..1047f6aeddcb 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/Basic.c +++ b/stage0/stdlib/Lean/Elab/Deriving/Basic.c @@ -143,7 +143,6 @@ lean_object* lean_st_mk_ref(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_resolveGlobalConstCore___at_Lean_Elab_elabDeriving___spec__12(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_to_list(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_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); uint8_t lean_is_out_param(lean_object*); lean_object* l_Lean_Meta_synthInstance(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -215,6 +214,7 @@ uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_elabDeriving_declRange___closed__1; static lean_object* l___regBuiltin_Lean_Elab_elabDeriving___closed__5; lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Elab_defaultHandler___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_688_(lean_object*); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); @@ -312,7 +312,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Deriving_Basic_0__Lean_Elab_Term_ { 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_Expr_app___override(x_1, x_2); -x_11 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_10, x_5, x_6, x_7, x_8, x_9); +x_11 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_10, 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); @@ -3193,6 +3193,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Deriving_Basic_0__Lean_Elab_tryAp lean_object* x_6; x_6 = l___private_Lean_Elab_Deriving_Basic_0__Lean_Elab_tryApplyDefHandler(x_1, x_2, x_3, x_4, x_5); lean_dec(x_4); +lean_dec(x_3); return x_6; } } @@ -4527,7 +4528,6 @@ if (x_76 == 0) lean_object* x_77; lean_object* x_78; lean_object* x_79; x_77 = l_Lean_instInhabitedName; x_78 = l___private_Init_Util_0__outOfBounds___rarg(x_77); -lean_inc(x_74); lean_inc(x_43); x_79 = l___private_Lean_Elab_Deriving_Basic_0__Lean_Elab_tryApplyDefHandler(x_43, x_78, x_74, x_7, x_65); if (lean_obj_tag(x_79) == 0) @@ -4620,7 +4620,6 @@ else { lean_object* x_96; lean_object* x_97; x_96 = lean_array_fget(x_1, x_75); -lean_inc(x_74); lean_inc(x_43); x_97 = l___private_Lean_Elab_Deriving_Basic_0__Lean_Elab_tryApplyDefHandler(x_43, x_96, x_74, x_7, x_65); if (lean_obj_tag(x_97) == 0) @@ -4960,7 +4959,6 @@ if (x_174 == 0) lean_object* x_175; lean_object* x_176; lean_object* x_177; x_175 = l_Lean_instInhabitedName; x_176 = l___private_Init_Util_0__outOfBounds___rarg(x_175); -lean_inc(x_172); lean_inc(x_141); x_177 = l___private_Lean_Elab_Deriving_Basic_0__Lean_Elab_tryApplyDefHandler(x_141, x_176, x_172, x_7, x_163); if (lean_obj_tag(x_177) == 0) @@ -5049,7 +5047,6 @@ else { lean_object* x_192; lean_object* x_193; x_192 = lean_array_fget(x_1, x_173); -lean_inc(x_172); lean_inc(x_141); x_193 = l___private_Lean_Elab_Deriving_Basic_0__Lean_Elab_tryApplyDefHandler(x_141, x_192, x_172, x_7, x_163); if (lean_obj_tag(x_193) == 0) diff --git a/stage0/stdlib/Lean/Elab/Deriving/DecEq.c b/stage0/stdlib/Lean/Elab/Deriving/DecEq.c index e72e02216d2f..45297a7c17fb 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/DecEq.c +++ b/stage0/stdlib/Lean/Elab/Deriving/DecEq.c @@ -23,25 +23,25 @@ static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch___closed__5; lean_object* l___private_Init_Util_0__outOfBounds___rarg(lean_object*); static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__32; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__3___closed__14; -static lean_object* l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__15; +static lean_object* l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__2; static lean_object* l_Lean_Elab_Deriving_DecEq_mkEnumOfNatThm___closed__10; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__16; 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_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957_(lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__20; +static lean_object* l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__1; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__1; static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__14; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__8(size_t, size_t, lean_object*); static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__13; LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqInstanceHandler___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__38; -static lean_object* l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__13; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___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_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__3___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__5; -static lean_object* l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__17; static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__4; static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__16; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__44; @@ -49,9 +49,11 @@ static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_DecEq_mkAuxFunction(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*); lean_object* l_Lean_mkAppB(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__16; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__23; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__15; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__30; +lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__18(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isEnumType___at_Lean_Elab_Deriving_DecEq_mkDecEqInstanceHandler___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__33; lean_object* l_Lean_MessageData_ofList(lean_object*); @@ -74,6 +76,7 @@ lean_object* l_Lean_Expr_fvarId_x21(lean_object*); static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__41; static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__9; LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs(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__62; LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_DecEq_mkEnumOfNat_mkDecTree___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__19; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___closed__6; @@ -89,7 +92,6 @@ lean_object* l_Lean_Syntax_node5(lean_object*, lean_object*, lean_object*, lean_ static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__21; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__16; static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__55; -static lean_object* l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__1; static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___lambda__1___closed__3; lean_object* l_Lean_stringToMessageData(lean_object*); static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__44; @@ -103,7 +105,6 @@ static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2 static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___closed__1; static lean_object* l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__3; static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__23; -static lean_object* l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__16; LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_DecEq_mkAuxFunction___boxed(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__19; static lean_object* l_Lean_Elab_Deriving_DecEq_mkEnumOfNat___lambda__1___closed__1; @@ -114,7 +115,6 @@ lean_object* l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRu static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___closed__4; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__24; lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__8; static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqHeader___closed__2; static lean_object* l_Lean_Elab_Deriving_DecEq_mkEnumOfNatThm___lambda__1___closed__1; static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__25; @@ -122,20 +122,22 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqCmds(lean_object*, le static lean_object* l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__12; size_t lean_usize_of_nat(lean_object*); static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__28; +static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__60; static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__49; static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__10; 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*); static lean_object* l_Lean_Elab_Deriving_DecEq_mkEnumOfNat___closed__2; LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Elab_Deriving_DecEq_mkDecEqCmds___spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__7; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__29; lean_object* l_Lean_Elab_registerDerivingHandler(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_DecEq_mkMatch___spec__1(size_t, size_t, lean_object*); lean_object* l_Lean_getConstInfo___at_Lean_Elab_Command_expandDeclId___spec__9(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___closed__1; static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__12; +static lean_object* l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__14; static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__35; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); -lean_object* l_Lean_addTrace___at_Lean_Elab_Command_elabCommand___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__14; lean_object* l_Lean_SourceInfo_fromRef(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_DecEq_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*); @@ -145,6 +147,7 @@ static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2 static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__5___closed__2; lean_object* l_Lean_Syntax_node6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_allM___at_Lean_Elab_Deriving_DecEq_mkDecEqInstanceHandler___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__17; static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__34; lean_object* lean_nat_div(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__18; @@ -166,6 +169,7 @@ lean_object* l_Lean_Elab_Command_getRef(lean_object*, lean_object*, lean_object* static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__48; lean_object* l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__8; +static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__58; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__3___closed__11; lean_object* l_Lean_Elab_getBetterRef(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__8___boxed(lean_object*, lean_object*, lean_object*); @@ -182,16 +186,16 @@ static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkEnumOfNatThm___spec__1(lean_object*, 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*); static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__51; -static lean_object* l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__10; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__19; static lean_object* l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__21; 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_List_isEmpty___rarg(lean_object*); +static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__63; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__3___closed__9; lean_object* lean_array_to_list(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__3; static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__24; static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__31; +lean_object* l_Lean_addTrace___at_Lean_Elab_Command_elabCommand___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__3; static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___closed__4; lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); @@ -225,6 +229,7 @@ lean_object* l_Lean_Core_betaReduce___lambda__1___boxed(lean_object*, lean_objec static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqCmds___closed__4; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__32; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__3; lean_object* l_Lean_Syntax_node2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__9; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__43; @@ -250,11 +255,9 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_DecEq_mk static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__31; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__5; static lean_object* l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__13; -LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942_(lean_object*); LEAN_EXPORT lean_object* l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_DecEq_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_DecEq_mkDecEqEnum___closed__27; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__3___closed__1; -lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_elabCommand___spec__7(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_append___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__54; static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__47; @@ -263,9 +266,8 @@ lean_object* l_Lean_MessageData_ofExpr(lean_object*); static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__13; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__4; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch___closed__4; -lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__19(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__5; -static lean_object* l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__9; +static lean_object* l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__8; +static lean_object* l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__9; static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__45; lean_object* l_Lean_Syntax_node4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__1; @@ -276,15 +278,17 @@ static lean_object* l_Lean_Elab_Deriving_DecEq_mkEnumOfNatThm___closed__5; static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqCmds___closed__1; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__12; lean_object* l_Lean_Core_transform___at_Lean_Core_betaReduce___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__59; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__9; -static lean_object* l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__2; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Deriving_DecEq_mkDecEq___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__10; +static lean_object* l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__12; uint8_t l_Lean_Expr_hasAnyFVar_visit___at_Lean_Expr_containsFVar___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__7; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__5; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__39; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__37; +lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__13(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__14; LEAN_EXPORT lean_object* l_Lean_getConstInfoInduct___at_Lean_Elab_Deriving_DecEq_mkEnumOfNat___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getConstInfoInduct___at_Lean_Elab_Deriving_DecEq_mkEnumOfNat___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -296,6 +300,7 @@ static lean_object* l_Lean_getConstInfoInduct___at_Lean_Elab_Deriving_DecEq_mkDe lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Term_traceAtCmdPos___spec__1(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*); static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__3___closed__3; +static lean_object* l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__6; static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___lambda__1___closed__10; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__35; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Deriving_DecEq_mkEnumOfNat___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -323,12 +328,13 @@ static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__3 static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__28; 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_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__50; +static lean_object* l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__15; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__18; -static lean_object* l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__14; static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___lambda__1___closed__7; +static lean_object* l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__4; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___closed__2; +static lean_object* l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__11; lean_object* lean_nat_sub(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__12; 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_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__27; extern lean_object* l_Lean_levelOne; @@ -337,7 +343,6 @@ uint8_t l_Lean_Expr_isProp(lean_object*); static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__49; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__2; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__11; -static lean_object* l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__6; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___closed__3; lean_object* l_Lean_mkArrow(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___lambda__1___closed__2; @@ -351,6 +356,7 @@ static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqCmds___closed__8; lean_object* l_List_reverse___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_DecEq_mkEnumOfNatThm___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_Deriving_DecEq_mkAuxFunction___closed__8; +static lean_object* l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__13; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__6; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__3___closed__15; static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__21; @@ -385,20 +391,21 @@ lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Command_instAddErrorMessag static lean_object* l_Lean_Elab_Deriving_DecEq_mkEnumOfNatThm___closed__1; lean_object* l_List_redLength___rarg(lean_object*); static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___lambda__1___closed__8; +static lean_object* l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__10; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__22; 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_mkAuxFunction___closed__16; +static lean_object* l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__5; lean_object* lean_array_get_size(lean_object*); LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_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*, lean_object*); static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__40; static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__3; +static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__61; lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); -static lean_object* l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__7; 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_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__52; -static lean_object* l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__4; static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__57; lean_object* l_Lean_addAndCompile(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); @@ -408,7 +415,6 @@ static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__3 LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqHeader(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__2___closed__26; LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_DecEq_mkEnumOfNatThm(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__11; static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__39; static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__42; lean_object* l_String_toSubstring_x27(lean_object*); @@ -1690,7 +1696,7 @@ lean_inc(x_3); x_13 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs(x_1, x_2, x_3, x_4, x_5, x_6, x_10, x_11, x_12); if (lean_obj_tag(x_13) == 0) { -lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t 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_14; lean_object* x_15; lean_object* x_16; uint8_t 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; x_14 = lean_ctor_get(x_13, 0); lean_inc(x_14); x_15 = lean_ctor_get(x_13, 1); @@ -1736,96 +1742,94 @@ x_34 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_34, 0, x_18); lean_ctor_set(x_34, 1, x_32); lean_ctor_set(x_34, 2, x_33); -lean_inc(x_31); -lean_inc(x_18); -x_35 = l_Lean_Syntax_node2(x_18, x_32, x_31, x_34); -x_36 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__9; +x_35 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__9; lean_inc(x_18); -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_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__12; +x_36 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_36, 0, x_18); +lean_ctor_set(x_36, 1, x_35); +x_37 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__12; 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_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__11; +x_38 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_38, 0, x_18); +lean_ctor_set(x_38, 1, x_37); +x_39 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__11; lean_inc(x_18); -x_41 = l_Lean_Syntax_node3(x_18, x_40, x_7, x_39, x_8); -x_42 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__3___closed__8; +x_40 = l_Lean_Syntax_node3(x_18, x_39, x_7, x_38, x_8); +x_41 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__3___closed__8; lean_inc(x_18); -x_43 = l_Lean_Syntax_node2(x_18, x_42, x_37, x_41); +x_42 = l_Lean_Syntax_node2(x_18, x_41, x_36, x_40); lean_inc(x_18); -x_44 = l_Lean_Syntax_node1(x_18, x_32, x_43); -x_45 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__16; +x_43 = l_Lean_Syntax_node1(x_18, x_32, x_42); +x_44 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__16; lean_inc(x_18); -x_46 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_46, 0, x_18); -lean_ctor_set(x_46, 1, x_45); -x_47 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__3___closed__11; -x_48 = l_Lean_addMacroScope(x_24, x_47, x_19); -x_49 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__3___closed__10; -x_50 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__3___closed__15; +x_45 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_45, 0, x_18); +lean_ctor_set(x_45, 1, x_44); +x_46 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__3___closed__11; +x_47 = l_Lean_addMacroScope(x_24, x_46, x_19); +x_48 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__3___closed__10; +x_49 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__3___closed__15; lean_inc(x_18); -x_51 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_51, 0, x_18); -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_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__3___closed__6; +x_50 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_50, 0, x_18); +lean_ctor_set(x_50, 1, x_48); +lean_ctor_set(x_50, 2, x_47); +lean_ctor_set(x_50, 3, x_49); +x_51 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__3___closed__6; +lean_inc(x_31); lean_inc(x_18); -x_53 = l_Lean_Syntax_node4(x_18, x_52, x_35, x_44, x_46, x_51); -x_54 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__3___closed__4; +x_52 = l_Lean_Syntax_node5(x_18, x_51, x_31, x_34, x_43, x_45, x_50); +x_53 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__3___closed__4; lean_inc(x_18); -x_55 = l_Lean_Syntax_node1(x_18, x_54, x_53); -x_56 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__19; +x_54 = l_Lean_Syntax_node1(x_18, x_53, x_52); +x_55 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__19; 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_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__16; +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_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__16; lean_inc(x_18); -x_59 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_59, 0, x_18); -lean_ctor_set(x_59, 1, x_58); -x_60 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__22; +x_58 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_58, 0, x_18); +lean_ctor_set(x_58, 1, x_57); +x_59 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__22; lean_inc(x_18); -x_61 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_61, 0, x_18); -lean_ctor_set(x_61, 1, x_60); +x_60 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_60, 0, x_18); +lean_ctor_set(x_60, 1, x_59); lean_inc(x_18); -x_62 = l_Lean_Syntax_node1(x_18, x_32, x_31); -x_63 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__23; +x_61 = l_Lean_Syntax_node1(x_18, x_32, x_31); +x_62 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__23; lean_inc(x_18); -x_64 = l_Lean_Syntax_node2(x_18, x_63, x_61, x_62); -x_65 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__24; +x_63 = l_Lean_Syntax_node2(x_18, x_62, x_60, x_61); +x_64 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__24; lean_inc(x_18); -x_66 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_66, 0, x_18); -lean_ctor_set(x_66, 1, x_65); -x_67 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__25; +x_65 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_65, 0, x_18); +lean_ctor_set(x_65, 1, x_64); +x_66 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__25; lean_inc(x_18); -x_68 = l_Lean_Syntax_node2(x_18, x_67, x_66, x_14); -lean_inc(x_57); +x_67 = l_Lean_Syntax_node2(x_18, x_66, x_65, x_14); +lean_inc(x_56); lean_inc(x_18); -x_69 = l_Lean_Syntax_node3(x_18, x_32, x_64, x_57, x_68); -x_70 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__21; +x_68 = l_Lean_Syntax_node3(x_18, x_32, x_63, x_56, x_67); +x_69 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__21; lean_inc(x_18); -x_71 = l_Lean_Syntax_node1(x_18, x_70, x_69); -x_72 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__19; +x_70 = l_Lean_Syntax_node1(x_18, x_69, x_68); +x_71 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__19; lean_inc(x_18); -x_73 = l_Lean_Syntax_node1(x_18, x_72, x_71); -x_74 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__15; +x_72 = l_Lean_Syntax_node1(x_18, x_71, x_70); +x_73 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__15; lean_inc(x_18); -x_75 = l_Lean_Syntax_node2(x_18, x_74, x_59, x_73); -x_76 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__3___closed__2; -x_77 = l_Lean_Syntax_node4(x_18, x_76, x_26, x_55, x_57, x_75); -x_78 = lean_apply_8(x_9, x_77, x_3, x_4, x_5, x_6, x_10, x_11, x_22); -return x_78; +x_74 = l_Lean_Syntax_node2(x_18, x_73, x_58, x_72); +x_75 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__3___closed__2; +x_76 = l_Lean_Syntax_node4(x_18, x_75, x_26, x_54, x_56, x_74); +x_77 = lean_apply_8(x_9, x_76, x_3, x_4, x_5, x_6, x_10, x_11, x_22); +return x_77; } else { -uint8_t x_79; +uint8_t x_78; lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -1835,23 +1839,23 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_79 = !lean_is_exclusive(x_13); -if (x_79 == 0) +x_78 = !lean_is_exclusive(x_13); +if (x_78 == 0) { return x_13; } else { -lean_object* x_80; lean_object* x_81; lean_object* x_82; -x_80 = lean_ctor_get(x_13, 0); -x_81 = lean_ctor_get(x_13, 1); -lean_inc(x_81); +lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_79 = lean_ctor_get(x_13, 0); +x_80 = lean_ctor_get(x_13, 1); lean_inc(x_80); +lean_inc(x_79); lean_dec(x_13); -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; +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; } } } @@ -5667,7 +5671,6 @@ lean_inc(x_8); lean_dec(x_5); x_9 = lean_alloc_closure((void*)(l_Lean_Elab_Deriving_DecEq_mkDecEqCmds), 8, 1); lean_closure_set(x_9, 0, x_6); -lean_inc(x_2); x_10 = l_Lean_Elab_Command_liftTermElabM___rarg(x_9, x_2, x_3, x_8); if (lean_obj_tag(x_10) == 0) { @@ -5717,7 +5720,7 @@ x_22 = 0; x_23 = lean_usize_of_nat(x_14); lean_dec(x_14); x_24 = lean_box(0); -x_25 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__19(x_12, x_22, x_23, x_24, x_2, x_3, x_13); +x_25 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__18(x_12, x_22, x_23, x_24, x_2, x_3, x_13); lean_dec(x_12); if (lean_obj_tag(x_25) == 0) { @@ -5822,7 +5825,7 @@ x_50 = 0; x_51 = lean_usize_of_nat(x_40); lean_dec(x_40); x_52 = lean_box(0); -x_53 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__19(x_38, x_50, x_51, x_52, x_2, x_3, x_39); +x_53 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__18(x_38, x_50, x_51, x_52, x_2, x_3, x_39); lean_dec(x_38); if (lean_obj_tag(x_53) == 0) { @@ -7267,7 +7270,7 @@ static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__30() _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("tacticHave_", 11); +x_1 = lean_mk_string_from_bytes("tacticRefine_lift_", 18); return x_1; } } @@ -7287,30 +7290,58 @@ static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__32() _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("aux", 3); +x_1 = lean_mk_string_from_bytes("refine_lift", 11); return x_1; } } static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__33() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("let_fun", 7); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__34() { +_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_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__1; +x_2 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__2; +x_3 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__3; +x_4 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__33; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__35() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("aux", 3); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__36() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__32; +x_1 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__35; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__34() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___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_Deriving_DecEq_mkDecEqEnum___closed__32; +x_2 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__35; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__35() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__38() { _start: { lean_object* x_1; @@ -7318,50 +7349,78 @@ x_1 = lean_mk_string_from_bytes("congrArg", 8); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__36() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__39() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__35; +x_1 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__38; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__37() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___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_Deriving_DecEq_mkDecEqEnum___closed__35; +x_2 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__38; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__38() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__41() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__37; +x_2 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__40; 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_DecEq_mkDecEqEnum___closed__39() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___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_Deriving_DecEq_mkDecEqEnum___closed__38; +x_2 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__41; 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_DecEq_mkDecEqEnum___closed__40() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__43() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("syntheticHole", 13); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__44() { +_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_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__1; +x_2 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__2; +x_3 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__3; +x_4 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__43; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__45() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("?", 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__46() { _start: { lean_object* x_1; @@ -7369,19 +7428,19 @@ x_1 = lean_mk_string_from_bytes("rwSeq", 5); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__41() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__47() { _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_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__1; x_2 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__2; x_3 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__17; -x_4 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__40; +x_4 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__46; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__42() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__48() { _start: { lean_object* x_1; @@ -7389,7 +7448,7 @@ x_1 = lean_mk_string_from_bytes("rw", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__43() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__49() { _start: { lean_object* x_1; @@ -7397,19 +7456,19 @@ x_1 = lean_mk_string_from_bytes("rwRuleSeq", 9); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__44() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__50() { _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_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__1; x_2 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__2; x_3 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__17; -x_4 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__43; +x_4 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__49; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__45() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__51() { _start: { lean_object* x_1; @@ -7417,7 +7476,7 @@ x_1 = lean_mk_string_from_bytes("[", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__46() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__52() { _start: { lean_object* x_1; @@ -7425,19 +7484,19 @@ x_1 = lean_mk_string_from_bytes("rwRule", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__47() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__53() { _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_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__1; x_2 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__2; x_3 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__17; -x_4 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__46; +x_4 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__52; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__48() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__54() { _start: { lean_object* x_1; @@ -7445,7 +7504,7 @@ x_1 = lean_mk_string_from_bytes("]", 1); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__49() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__55() { _start: { lean_object* x_1; @@ -7453,19 +7512,19 @@ x_1 = lean_mk_string_from_bytes("location", 8); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__50() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__56() { _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_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__1; x_2 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__2; x_3 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__17; -x_4 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__49; +x_4 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__55; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__51() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__57() { _start: { lean_object* x_1; @@ -7473,7 +7532,7 @@ x_1 = lean_mk_string_from_bytes("at", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__52() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__58() { _start: { lean_object* x_1; @@ -7481,19 +7540,19 @@ x_1 = lean_mk_string_from_bytes("locationHyp", 11); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__53() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__59() { _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_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__1; x_2 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__2; x_3 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__17; -x_4 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__52; +x_4 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__58; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__54() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__60() { _start: { lean_object* x_1; @@ -7501,19 +7560,19 @@ x_1 = lean_mk_string_from_bytes("tacticRfl", 9); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__55() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__61() { _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_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__1; x_2 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__2; x_3 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__17; -x_4 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__54; +x_4 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__60; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__56() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__62() { _start: { lean_object* x_1; @@ -7521,14 +7580,14 @@ x_1 = lean_mk_string_from_bytes("contradiction", 13); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__57() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__63() { _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_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__1; x_2 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__2; x_3 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__17; -x_4 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__56; +x_4 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__62; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } @@ -7540,7 +7599,6 @@ lean_object* x_5; lean_object* x_6; lean_inc(x_1); x_5 = lean_alloc_closure((void*)(l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___lambda__1___boxed), 8, 1); lean_closure_set(x_5, 0, x_1); -lean_inc(x_2); x_6 = l_Lean_Elab_Command_liftTermElabM___rarg(x_5, x_2, x_3, x_4); if (lean_obj_tag(x_6) == 0) { @@ -7551,11 +7609,10 @@ lean_dec(x_6); lean_inc(x_1); x_8 = lean_alloc_closure((void*)(l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___lambda__2___boxed), 8, 1); lean_closure_set(x_8, 0, x_1); -lean_inc(x_2); x_9 = l_Lean_Elab_Command_liftTermElabM___rarg(x_8, x_2, x_3, x_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; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t 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; uint8_t x_220; +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; 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; uint8_t x_231; x_10 = lean_ctor_get(x_9, 1); lean_inc(x_10); lean_dec(x_9); @@ -7765,334 +7822,356 @@ lean_inc(x_21); x_102 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_102, 0, x_21); lean_ctor_set(x_102, 1, x_101); -x_103 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__3___closed__1; +x_103 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__32; lean_inc(x_21); x_104 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_104, 0, x_21); lean_ctor_set(x_104, 1, x_103); -x_105 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__34; +x_105 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__33; +lean_inc(x_21); +x_106 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_106, 0, x_21); +lean_ctor_set(x_106, 1, x_105); +x_107 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__37; lean_inc(x_23); lean_inc(x_26); -x_106 = l_Lean_addMacroScope(x_26, x_105, x_23); -x_107 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__33; -lean_inc(x_21); -x_108 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_108, 0, x_21); -lean_ctor_set(x_108, 1, x_107); -lean_ctor_set(x_108, 2, x_106); -lean_ctor_set(x_108, 3, x_41); -lean_inc(x_30); -lean_inc(x_108); +x_108 = l_Lean_addMacroScope(x_26, x_107, x_23); +x_109 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__36; lean_inc(x_21); -x_109 = l_Lean_Syntax_node2(x_21, x_28, x_108, x_30); -x_110 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__37; +x_110 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_110, 0, x_21); +lean_ctor_set(x_110, 1, x_109); +lean_ctor_set(x_110, 2, x_108); +lean_ctor_set(x_110, 3, x_41); +x_111 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__40; lean_inc(x_23); lean_inc(x_26); -x_111 = l_Lean_addMacroScope(x_26, x_110, x_23); -x_112 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__36; +x_112 = l_Lean_addMacroScope(x_26, x_111, x_23); x_113 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__39; +x_114 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__42; lean_inc(x_21); -x_114 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_114, 0, x_21); -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 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_115, 0, x_21); +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_73); lean_inc(x_21); -x_115 = l_Lean_Syntax_node2(x_21, x_28, x_13, x_73); +x_116 = l_Lean_Syntax_node2(x_21, x_28, x_13, x_73); lean_inc(x_21); -x_116 = l_Lean_Syntax_node2(x_21, x_47, x_114, x_115); -x_117 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__3___closed__6; +x_117 = l_Lean_Syntax_node2(x_21, x_47, x_115, x_116); +x_118 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__9; lean_inc(x_54); -lean_inc(x_30); -lean_inc(x_21); -x_118 = l_Lean_Syntax_node4(x_21, x_117, x_109, x_30, x_54, x_116); -x_119 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__3___closed__4; +lean_inc_n(x_30, 2); +lean_inc(x_110); lean_inc(x_21); -x_120 = l_Lean_Syntax_node1(x_21, x_119, x_118); -x_121 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__31; +x_119 = l_Lean_Syntax_node5(x_21, x_118, x_110, x_30, x_30, x_54, x_117); +x_120 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__7; lean_inc(x_21); -x_122 = l_Lean_Syntax_node2(x_21, x_121, x_104, x_120); -x_123 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__19; +x_121 = l_Lean_Syntax_node1(x_21, x_120, x_119); +x_122 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__19; lean_inc(x_21); -x_124 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_124, 0, x_21); -lean_ctor_set(x_124, 1, x_123); -x_125 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__42; +x_123 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_123, 0, x_21); +lean_ctor_set(x_123, 1, x_122); +x_124 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__45; lean_inc(x_21); -x_126 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_126, 0, x_21); -lean_ctor_set(x_126, 1, x_125); -x_127 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__45; +x_125 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_125, 0, x_21); +lean_ctor_set(x_125, 1, x_124); +x_126 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__48; lean_inc(x_21); -x_128 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_128, 0, x_21); -lean_ctor_set(x_128, 1, x_127); -x_129 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__47; -lean_inc(x_30); +x_127 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_127, 0, x_21); +lean_ctor_set(x_127, 1, x_126); +x_128 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__47; lean_inc(x_21); -x_130 = l_Lean_Syntax_node2(x_21, x_129, x_30, x_16); -x_131 = l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___lambda__1___closed__9; +x_129 = l_Lean_Syntax_node1(x_21, x_128, x_127); +x_130 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__44; lean_inc(x_21); -x_132 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_132, 0, x_21); -lean_ctor_set(x_132, 1, x_131); -lean_inc(x_130); +x_131 = l_Lean_Syntax_node2(x_21, x_130, x_125, x_129); +x_132 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__34; +lean_inc(x_123); lean_inc(x_21); -x_133 = l_Lean_Syntax_node3(x_21, x_28, x_130, x_132, x_130); -x_134 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__48; +x_133 = l_Lean_Syntax_node4(x_21, x_132, x_106, x_121, x_123, x_131); +x_134 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__31; lean_inc(x_21); -x_135 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_135, 0, x_21); -lean_ctor_set(x_135, 1, x_134); -x_136 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__44; +x_135 = l_Lean_Syntax_node2(x_21, x_134, x_104, x_133); +x_136 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__48; lean_inc(x_21); -x_137 = l_Lean_Syntax_node3(x_21, x_136, x_128, x_133, x_135); +x_137 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_137, 0, x_21); +lean_ctor_set(x_137, 1, x_136); x_138 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__51; lean_inc(x_21); x_139 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_139, 0, x_21); lean_ctor_set(x_139, 1, x_138); -lean_inc(x_21); -x_140 = l_Lean_Syntax_node1(x_21, x_28, x_108); -x_141 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__53; +x_140 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__53; lean_inc(x_30); lean_inc(x_21); -x_142 = l_Lean_Syntax_node2(x_21, x_141, x_140, x_30); -x_143 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__50; +x_141 = l_Lean_Syntax_node2(x_21, x_140, x_30, x_16); +x_142 = l_List_forIn_loop___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__6___lambda__1___closed__9; lean_inc(x_21); -x_144 = l_Lean_Syntax_node2(x_21, x_143, x_139, x_142); +x_143 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_143, 0, x_21); +lean_ctor_set(x_143, 1, x_142); +lean_inc(x_141); lean_inc(x_21); -x_145 = l_Lean_Syntax_node1(x_21, x_28, x_144); -x_146 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__41; -lean_inc(x_30); +x_144 = l_Lean_Syntax_node3(x_21, x_28, x_141, x_143, x_141); +x_145 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__54; lean_inc(x_21); -x_147 = l_Lean_Syntax_node4(x_21, x_146, x_126, x_30, x_137, x_145); -x_148 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__49; +x_146 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_146, 0, x_21); +lean_ctor_set(x_146, 1, x_145); +x_147 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__50; lean_inc(x_21); -x_149 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_149, 0, x_21); -lean_ctor_set(x_149, 1, x_148); -x_150 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__50; +x_148 = l_Lean_Syntax_node3(x_21, x_147, x_139, x_144, x_146); +x_149 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__57; lean_inc(x_21); -x_151 = l_Lean_Syntax_node1(x_21, x_150, x_149); -lean_inc_n(x_124, 2); +x_150 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_150, 0, x_21); +lean_ctor_set(x_150, 1, x_149); lean_inc(x_21); -x_152 = l_Lean_Syntax_node5(x_21, x_28, x_122, x_124, x_147, x_124, x_151); -x_153 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__21; +x_151 = l_Lean_Syntax_node1(x_21, x_28, x_110); +x_152 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__59; +lean_inc(x_30); lean_inc(x_21); -x_154 = l_Lean_Syntax_node1(x_21, x_153, x_152); -x_155 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__19; +x_153 = l_Lean_Syntax_node2(x_21, x_152, x_151, x_30); +x_154 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__56; lean_inc(x_21); -x_156 = l_Lean_Syntax_node1(x_21, x_155, x_154); -x_157 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__29; -lean_inc(x_102); +x_155 = l_Lean_Syntax_node2(x_21, x_154, x_150, x_153); +lean_inc(x_21); +x_156 = l_Lean_Syntax_node1(x_21, x_28, x_155); +x_157 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__47; +lean_inc(x_30); lean_inc(x_21); -x_158 = l_Lean_Syntax_node2(x_21, x_157, x_102, x_156); -x_159 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__3___closed__9; +x_158 = l_Lean_Syntax_node4(x_21, x_157, x_137, x_30, x_148, x_156); +x_159 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__49; lean_inc(x_21); x_160 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_160, 0, x_21); lean_ctor_set(x_160, 1, x_159); -x_161 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__55; +x_161 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__50; lean_inc(x_21); x_162 = l_Lean_Syntax_node1(x_21, x_161, x_160); +lean_inc_n(x_123, 2); lean_inc(x_21); -x_163 = l_Lean_Syntax_node1(x_21, x_28, x_162); +x_163 = l_Lean_Syntax_node5(x_21, x_28, x_135, x_123, x_158, x_123, x_162); +x_164 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__21; lean_inc(x_21); -x_164 = l_Lean_Syntax_node1(x_21, x_153, x_163); +x_165 = l_Lean_Syntax_node1(x_21, x_164, x_163); +x_166 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__19; lean_inc(x_21); -x_165 = l_Lean_Syntax_node1(x_21, x_155, x_164); +x_167 = l_Lean_Syntax_node1(x_21, x_166, x_165); +x_168 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__29; +lean_inc(x_102); lean_inc(x_21); -x_166 = l_Lean_Syntax_node2(x_21, x_157, x_102, x_165); +x_169 = l_Lean_Syntax_node2(x_21, x_168, x_102, x_167); +x_170 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__3___closed__9; lean_inc(x_21); -x_167 = l_Lean_Syntax_node2(x_21, x_28, x_158, x_166); -x_168 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__27; +x_171 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_171, 0, x_21); +lean_ctor_set(x_171, 1, x_170); +x_172 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__61; lean_inc(x_21); -x_169 = l_Lean_Syntax_node2(x_21, x_168, x_100, x_167); +x_173 = l_Lean_Syntax_node1(x_21, x_172, x_171); lean_inc(x_21); -x_170 = l_Lean_Syntax_node1(x_21, x_28, x_169); +x_174 = l_Lean_Syntax_node1(x_21, x_28, x_173); lean_inc(x_21); -x_171 = l_Lean_Syntax_node1(x_21, x_153, x_170); +x_175 = l_Lean_Syntax_node1(x_21, x_164, x_174); lean_inc(x_21); -x_172 = l_Lean_Syntax_node1(x_21, x_155, x_171); -x_173 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__15; -lean_inc(x_98); +x_176 = l_Lean_Syntax_node1(x_21, x_166, x_175); lean_inc(x_21); -x_174 = l_Lean_Syntax_node2(x_21, x_173, x_98, x_172); -x_175 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__51; +x_177 = l_Lean_Syntax_node2(x_21, x_168, x_102, x_176); lean_inc(x_21); -x_176 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_176, 0, x_21); -lean_ctor_set(x_176, 1, x_175); -x_177 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__35; +x_178 = l_Lean_Syntax_node2(x_21, x_28, x_169, x_177); +x_179 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__27; lean_inc(x_21); -x_178 = l_Lean_Syntax_node3(x_21, x_177, x_96, x_174, x_176); +x_180 = l_Lean_Syntax_node2(x_21, x_179, x_100, x_178); lean_inc(x_21); -x_179 = l_Lean_Syntax_node1(x_21, x_28, x_178); +x_181 = l_Lean_Syntax_node1(x_21, x_28, x_180); lean_inc(x_21); -x_180 = l_Lean_Syntax_node2(x_21, x_47, x_94, x_179); -x_181 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__26; +x_182 = l_Lean_Syntax_node1(x_21, x_164, x_181); lean_inc(x_21); -x_182 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_182, 0, x_21); -lean_ctor_set(x_182, 1, x_181); -x_183 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__29; -x_184 = l_Lean_addMacroScope(x_26, x_183, x_23); -x_185 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__28; -x_186 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__33; +x_183 = l_Lean_Syntax_node1(x_21, x_166, x_182); +x_184 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__15; +lean_inc(x_98); +lean_inc(x_21); +x_185 = l_Lean_Syntax_node2(x_21, x_184, x_98, x_183); +x_186 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__51; lean_inc(x_21); -x_187 = lean_alloc_ctor(3, 4, 0); +x_187 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_187, 0, x_21); -lean_ctor_set(x_187, 1, x_185); -lean_ctor_set(x_187, 2, x_184); -lean_ctor_set(x_187, 3, x_186); +lean_ctor_set(x_187, 1, x_186); +x_188 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__35; +lean_inc(x_21); +x_189 = l_Lean_Syntax_node3(x_21, x_188, x_96, x_185, x_187); +lean_inc(x_21); +x_190 = l_Lean_Syntax_node1(x_21, x_28, x_189); +lean_inc(x_21); +x_191 = l_Lean_Syntax_node2(x_21, x_47, x_94, x_190); +x_192 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__26; +lean_inc(x_21); +x_193 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_193, 0, x_21); +lean_ctor_set(x_193, 1, x_192); +x_194 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__29; +x_195 = l_Lean_addMacroScope(x_26, x_194, x_23); +x_196 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__28; +x_197 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__33; lean_inc(x_21); -x_188 = l_Lean_Syntax_node1(x_21, x_28, x_73); -x_189 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__22; +x_198 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_198, 0, x_21); +lean_ctor_set(x_198, 1, x_196); +lean_ctor_set(x_198, 2, x_195); +lean_ctor_set(x_198, 3, x_197); lean_inc(x_21); -x_190 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_190, 0, x_21); -lean_ctor_set(x_190, 1, x_189); -x_191 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__23; -lean_inc(x_188); +x_199 = l_Lean_Syntax_node1(x_21, x_28, x_73); +x_200 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__22; lean_inc(x_21); -x_192 = l_Lean_Syntax_node2(x_21, x_191, x_190, x_188); -x_193 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__56; +x_201 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_201, 0, x_21); +lean_ctor_set(x_201, 1, x_200); +x_202 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__23; +lean_inc(x_199); lean_inc(x_21); -x_194 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_194, 0, x_21); -lean_ctor_set(x_194, 1, x_193); -x_195 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__57; +x_203 = l_Lean_Syntax_node2(x_21, x_202, x_201, x_199); +x_204 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__62; lean_inc(x_21); -x_196 = l_Lean_Syntax_node1(x_21, x_195, x_194); +x_205 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_205, 0, x_21); +lean_ctor_set(x_205, 1, x_204); +x_206 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__63; lean_inc(x_21); -x_197 = l_Lean_Syntax_node3(x_21, x_28, x_192, x_124, x_196); +x_207 = l_Lean_Syntax_node1(x_21, x_206, x_205); lean_inc(x_21); -x_198 = l_Lean_Syntax_node1(x_21, x_153, x_197); +x_208 = l_Lean_Syntax_node3(x_21, x_28, x_203, x_123, x_207); lean_inc(x_21); -x_199 = l_Lean_Syntax_node1(x_21, x_155, x_198); +x_209 = l_Lean_Syntax_node1(x_21, x_164, x_208); lean_inc(x_21); -x_200 = l_Lean_Syntax_node2(x_21, x_173, x_98, x_199); -x_201 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__13; +x_210 = l_Lean_Syntax_node1(x_21, x_166, x_209); +lean_inc(x_21); +x_211 = l_Lean_Syntax_node2(x_21, x_184, x_98, x_210); +x_212 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__13; lean_inc(x_67); lean_inc(x_30); lean_inc(x_21); -x_202 = l_Lean_Syntax_node4(x_21, x_201, x_188, x_30, x_67, x_200); -x_203 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__11; +x_213 = l_Lean_Syntax_node4(x_21, x_212, x_199, x_30, x_67, x_211); +x_214 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__11; lean_inc(x_56); lean_inc(x_21); -x_204 = l_Lean_Syntax_node2(x_21, x_203, x_56, x_202); +x_215 = l_Lean_Syntax_node2(x_21, x_214, x_56, x_213); lean_inc(x_21); -x_205 = l_Lean_Syntax_node1(x_21, x_28, x_204); +x_216 = l_Lean_Syntax_node1(x_21, x_28, x_215); lean_inc(x_21); -x_206 = l_Lean_Syntax_node2(x_21, x_47, x_187, x_205); -x_207 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__2; +x_217 = l_Lean_Syntax_node2(x_21, x_47, x_198, x_216); +x_218 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__2; lean_inc(x_21); -x_208 = l_Lean_Syntax_node8(x_21, x_207, x_69, x_75, x_38, x_87, x_89, x_180, x_182, x_206); +x_219 = l_Lean_Syntax_node8(x_21, x_218, x_69, x_75, x_38, x_87, x_89, x_191, x_193, x_217); lean_inc(x_30); lean_inc(x_21); -x_209 = l_Lean_Syntax_node4(x_21, x_201, x_65, x_30, x_67, x_208); +x_220 = l_Lean_Syntax_node4(x_21, x_212, x_65, x_30, x_67, x_219); lean_inc(x_21); -x_210 = l_Lean_Syntax_node2(x_21, x_203, x_56, x_209); -x_211 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__15; +x_221 = l_Lean_Syntax_node2(x_21, x_214, x_56, x_220); +x_222 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__15; lean_inc(x_30); lean_inc(x_21); -x_212 = l_Lean_Syntax_node3(x_21, x_211, x_54, x_210, x_30); -x_213 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__2; +x_223 = l_Lean_Syntax_node3(x_21, x_222, x_54, x_221, x_30); +x_224 = l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__2; lean_inc_n(x_30, 3); lean_inc(x_21); -x_214 = l_Lean_Syntax_node8(x_21, x_213, x_34, x_36, x_30, x_30, x_52, x_212, x_30, x_30); -x_215 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__3; -x_216 = l_Lean_Syntax_node2(x_21, x_215, x_32, x_214); -x_217 = l_Lean_Elab_Deriving_DecEq_mkDecEqCmds___closed__5; -x_218 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_elabCommand___spec__7(x_217, x_2, x_3, x_27); -x_219 = lean_ctor_get(x_218, 0); -lean_inc(x_219); -x_220 = lean_unbox(x_219); -lean_dec(x_219); -if (x_220 == 0) -{ -lean_object* x_221; lean_object* x_222; -x_221 = lean_ctor_get(x_218, 1); -lean_inc(x_221); -lean_dec(x_218); -x_222 = l_Lean_Elab_Command_elabCommand(x_216, x_2, x_3, x_221); -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_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; -x_223 = lean_ctor_get(x_218, 1); -lean_inc(x_223); -lean_dec(x_218); -lean_inc(x_216); -x_224 = l_Lean_MessageData_ofSyntax(x_216); -x_225 = l_Lean_Elab_Deriving_DecEq_mkDecEqCmds___closed__7; -x_226 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_226, 0, x_225); -lean_ctor_set(x_226, 1, x_224); -x_227 = l_Lean_Elab_Deriving_DecEq_mkDecEqCmds___closed__9; -x_228 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_228, 0, x_226); -lean_ctor_set(x_228, 1, x_227); -x_229 = l_Lean_addTrace___at_Lean_Elab_Command_elabCommand___spec__8(x_217, x_228, x_2, x_3, x_223); -x_230 = lean_ctor_get(x_229, 1); +x_225 = l_Lean_Syntax_node8(x_21, x_224, x_34, x_36, x_30, x_30, x_52, x_223, x_30, x_30); +x_226 = l_Lean_Elab_Deriving_DecEq_mkAuxFunction___closed__3; +x_227 = l_Lean_Syntax_node2(x_21, x_226, x_32, x_225); +x_228 = l_Lean_Elab_Deriving_DecEq_mkDecEqCmds___closed__5; +x_229 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__13(x_228, x_2, x_3, x_27); +x_230 = lean_ctor_get(x_229, 0); lean_inc(x_230); +x_231 = lean_unbox(x_230); +lean_dec(x_230); +if (x_231 == 0) +{ +lean_object* x_232; lean_object* x_233; +x_232 = lean_ctor_get(x_229, 1); +lean_inc(x_232); +lean_dec(x_229); +x_233 = l_Lean_Elab_Command_elabCommand(x_227, x_2, x_3, x_232); +return x_233; +} +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; +x_234 = lean_ctor_get(x_229, 1); +lean_inc(x_234); lean_dec(x_229); -x_231 = l_Lean_Elab_Command_elabCommand(x_216, x_2, x_3, x_230); -return x_231; +lean_inc(x_227); +x_235 = l_Lean_MessageData_ofSyntax(x_227); +x_236 = l_Lean_Elab_Deriving_DecEq_mkDecEqCmds___closed__7; +x_237 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_237, 0, x_236); +lean_ctor_set(x_237, 1, x_235); +x_238 = l_Lean_Elab_Deriving_DecEq_mkDecEqCmds___closed__9; +x_239 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_239, 0, x_237); +lean_ctor_set(x_239, 1, x_238); +x_240 = l_Lean_addTrace___at_Lean_Elab_Command_elabCommand___spec__7(x_228, x_239, x_2, x_3, x_234); +x_241 = lean_ctor_get(x_240, 1); +lean_inc(x_241); +lean_dec(x_240); +x_242 = l_Lean_Elab_Command_elabCommand(x_227, x_2, x_3, x_241); +return x_242; } } else { -uint8_t x_232; +uint8_t x_243; lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_232 = !lean_is_exclusive(x_9); -if (x_232 == 0) +x_243 = !lean_is_exclusive(x_9); +if (x_243 == 0) { return x_9; } else { -lean_object* x_233; lean_object* x_234; lean_object* x_235; -x_233 = lean_ctor_get(x_9, 0); -x_234 = lean_ctor_get(x_9, 1); -lean_inc(x_234); -lean_inc(x_233); +lean_object* x_244; lean_object* x_245; lean_object* x_246; +x_244 = lean_ctor_get(x_9, 0); +x_245 = lean_ctor_get(x_9, 1); +lean_inc(x_245); +lean_inc(x_244); lean_dec(x_9); -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; +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; } } } else { -uint8_t x_236; +uint8_t x_247; lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_236 = !lean_is_exclusive(x_6); -if (x_236 == 0) +x_247 = !lean_is_exclusive(x_6); +if (x_247 == 0) { return x_6; } else { -lean_object* x_237; lean_object* x_238; lean_object* x_239; -x_237 = lean_ctor_get(x_6, 0); -x_238 = lean_ctor_get(x_6, 1); -lean_inc(x_238); -lean_inc(x_237); +lean_object* x_248; lean_object* x_249; lean_object* x_250; +x_248 = lean_ctor_get(x_6, 0); +x_249 = lean_ctor_get(x_6, 1); +lean_inc(x_249); +lean_inc(x_248); lean_dec(x_6); -x_239 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_239, 0, x_237); -lean_ctor_set(x_239, 1, x_238); -return x_239; +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; } } } @@ -8911,7 +8990,7 @@ lean_dec(x_1); return x_5; } } -static lean_object* _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__1() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__1() { _start: { lean_object* x_1; @@ -8919,7 +8998,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Deriving_DecEq_mkDecEqInstanceHandl return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__2() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -8929,27 +9008,27 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__3() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__2; +x_1 = l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__2; x_2 = l_Lean_Elab_Deriving_DecEq_mkDecEqCmds___closed__3; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__4() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__3; +x_1 = l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__3; x_2 = l_Lean_Elab_Deriving_DecEq_mkDecEqCmds___closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__5() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__5() { _start: { lean_object* x_1; @@ -8957,17 +9036,17 @@ x_1 = lean_mk_string_from_bytes("DecEq", 5); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__6() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__4; -x_2 = l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__5; +x_1 = l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__4; +x_2 = l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__7() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__7() { _start: { lean_object* x_1; @@ -8975,17 +9054,17 @@ x_1 = lean_mk_string_from_bytes("initFn", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__8() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__6; -x_2 = l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__7; +x_1 = l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__6; +x_2 = l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__7; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__9() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__9() { _start: { lean_object* x_1; @@ -8993,57 +9072,57 @@ x_1 = lean_mk_string_from_bytes("_@", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__10() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__8; -x_2 = l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__9; +x_1 = l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__8; +x_2 = l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__9; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__11() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__10; +x_1 = l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__10; x_2 = l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__12() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__11; +x_1 = l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__11; x_2 = l_Lean_Elab_Deriving_DecEq_mkDecEqCmds___closed__3; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__13() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__12; +x_1 = l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__12; x_2 = l_Lean_Elab_Deriving_DecEq_mkDecEqCmds___closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__14() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__13; -x_2 = l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__5; +x_1 = l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__13; +x_2 = l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__15() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__15() { _start: { lean_object* x_1; @@ -9051,32 +9130,32 @@ x_1 = lean_mk_string_from_bytes("_hyg", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__16() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__14; -x_2 = l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__15; +x_1 = l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__14; +x_2 = l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__15; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__17() { +static lean_object* _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__16; -x_2 = lean_unsigned_to_nat(4942u); +x_1 = l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__16; +x_2 = lean_unsigned_to_nat(4957u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957_(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_4942____closed__1; +x_3 = l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__1; x_4 = l_Lean_Elab_registerDerivingHandler(x_2, x_3, x_1); if (lean_obj_tag(x_4) == 0) { @@ -9086,7 +9165,7 @@ lean_inc(x_5); lean_dec(x_4); x_6 = l_Lean_Elab_Deriving_DecEq_mkDecEqCmds___closed__5; x_7 = 0; -x_8 = l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__17; +x_8 = l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__17; x_9 = l_Lean_registerTraceClass(x_6, x_7, x_8, x_5); return x_9; } @@ -9609,41 +9688,53 @@ l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__56 = _init_l_Lean_Elab_Deriving lean_mark_persistent(l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__56); l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__57 = _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__57(); lean_mark_persistent(l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__57); -l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__1 = _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__1(); -lean_mark_persistent(l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__1); -l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__2 = _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__2(); -lean_mark_persistent(l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__2); -l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__3 = _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__3(); -lean_mark_persistent(l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__3); -l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__4 = _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__4(); -lean_mark_persistent(l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__4); -l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__5 = _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__5(); -lean_mark_persistent(l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__5); -l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__6 = _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__6(); -lean_mark_persistent(l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__6); -l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__7 = _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__7(); -lean_mark_persistent(l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__7); -l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__8 = _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__8(); -lean_mark_persistent(l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__8); -l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__9 = _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__9(); -lean_mark_persistent(l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__9); -l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__10 = _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__10(); -lean_mark_persistent(l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__10); -l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__11 = _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__11(); -lean_mark_persistent(l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__11); -l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__12 = _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__12(); -lean_mark_persistent(l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__12); -l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__13 = _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__13(); -lean_mark_persistent(l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__13); -l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__14 = _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__14(); -lean_mark_persistent(l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__14); -l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__15 = _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__15(); -lean_mark_persistent(l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__15); -l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__16 = _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__16(); -lean_mark_persistent(l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__16); -l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__17 = _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__17(); -lean_mark_persistent(l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942____closed__17); -res = l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4942_(lean_io_mk_world()); +l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__58 = _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__58(); +lean_mark_persistent(l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__58); +l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__59 = _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__59(); +lean_mark_persistent(l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__59); +l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__60 = _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__60(); +lean_mark_persistent(l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__60); +l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__61 = _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__61(); +lean_mark_persistent(l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__61); +l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__62 = _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__62(); +lean_mark_persistent(l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__62); +l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__63 = _init_l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__63(); +lean_mark_persistent(l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__63); +l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__1 = _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__1(); +lean_mark_persistent(l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__1); +l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__2 = _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__2(); +lean_mark_persistent(l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__2); +l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__3 = _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__3(); +lean_mark_persistent(l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__3); +l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__4 = _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__4(); +lean_mark_persistent(l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__4); +l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__5 = _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__5(); +lean_mark_persistent(l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__5); +l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__6 = _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__6(); +lean_mark_persistent(l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__6); +l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__7 = _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__7(); +lean_mark_persistent(l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__7); +l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__8 = _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__8(); +lean_mark_persistent(l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__8); +l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__9 = _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__9(); +lean_mark_persistent(l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__9); +l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__10 = _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__10(); +lean_mark_persistent(l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__10); +l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__11 = _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__11(); +lean_mark_persistent(l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__11); +l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__12 = _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__12(); +lean_mark_persistent(l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__12); +l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__13 = _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__13(); +lean_mark_persistent(l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__13); +l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__14 = _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__14(); +lean_mark_persistent(l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__14); +l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__15 = _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__15(); +lean_mark_persistent(l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__15); +l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__16 = _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__16(); +lean_mark_persistent(l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__16); +l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__17 = _init_l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__17(); +lean_mark_persistent(l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957____closed__17); +res = l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4957_(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 b659480564b8..8c1b4586bdbb 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/FromToJson.c +++ b/stage0/stdlib/Lean/Elab/Deriving/FromToJson.c @@ -64,6 +64,7 @@ static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___ 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*); static lean_object* l_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__3___closed__2; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__4___closed__1; +lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__18(lean_object*, size_t, size_t, 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*, 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_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__3___closed__15; lean_object* lean_array_push(lean_object*, lean_object*); @@ -300,7 +301,6 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJs static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler_mkAlts___spec__7___lambda__2___closed__8; LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_FromToJson_mkJsonField___lambda__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler_mkAlts___spec__4___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_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__19(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___spec__3___closed__15; 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*); LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -5904,7 +5904,6 @@ x_20 = lean_alloc_closure((void*)(l_Lean_Elab_Deriving_FromToJson_mkToJsonInstan lean_closure_set(x_20, 0, x_15); lean_closure_set(x_20, 1, x_18); lean_closure_set(x_20, 2, x_1); -lean_inc(x_2); x_21 = l_Lean_Elab_Command_liftTermElabM___rarg(x_20, x_2, x_3, x_19); if (lean_obj_tag(x_21) == 0) { @@ -5954,7 +5953,7 @@ x_33 = 0; x_34 = lean_usize_of_nat(x_25); lean_dec(x_25); x_35 = lean_box(0); -x_36 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__19(x_23, x_33, x_34, x_35, x_2, x_3, x_24); +x_36 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__18(x_23, x_33, x_34, x_35, x_2, x_3, x_24); lean_dec(x_23); if (lean_obj_tag(x_36) == 0) { @@ -6059,7 +6058,7 @@ x_61 = 0; x_62 = lean_usize_of_nat(x_51); lean_dec(x_51); x_63 = lean_box(0); -x_64 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__19(x_49, x_61, x_62, x_63, x_2, x_3, x_50); +x_64 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__18(x_49, x_61, x_62, x_63, x_2, x_3, x_50); lean_dec(x_49); if (lean_obj_tag(x_64) == 0) { @@ -6171,7 +6170,6 @@ lean_object* x_82; lean_object* x_83; x_82 = lean_alloc_closure((void*)(l_Lean_Elab_Deriving_FromToJson_mkToJsonInstanceHandler___lambda__4___boxed), 9, 2); lean_closure_set(x_82, 0, x_15); lean_closure_set(x_82, 1, x_1); -lean_inc(x_2); x_83 = l_Lean_Elab_Command_liftTermElabM___rarg(x_82, x_2, x_3, x_13); if (lean_obj_tag(x_83) == 0) { @@ -6221,7 +6219,7 @@ x_95 = 0; x_96 = lean_usize_of_nat(x_87); lean_dec(x_87); x_97 = lean_box(0); -x_98 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__19(x_85, x_95, x_96, x_97, x_2, x_3, x_86); +x_98 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__18(x_85, x_95, x_96, x_97, x_2, x_3, x_86); lean_dec(x_85); if (lean_obj_tag(x_98) == 0) { @@ -6326,7 +6324,7 @@ x_123 = 0; x_124 = lean_usize_of_nat(x_113); lean_dec(x_113); x_125 = lean_box(0); -x_126 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__19(x_111, x_123, x_124, x_125, x_2, x_3, x_112); +x_126 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__18(x_111, x_123, x_124, x_125, x_2, x_3, x_112); lean_dec(x_111); if (lean_obj_tag(x_126) == 0) { @@ -11901,7 +11899,6 @@ x_22 = lean_alloc_closure((void*)(l_Lean_Elab_Deriving_FromToJson_mkFromJsonInst lean_closure_set(x_22, 0, x_14); lean_closure_set(x_22, 1, x_20); lean_closure_set(x_22, 2, x_1); -lean_inc(x_2); x_23 = l_Lean_Elab_Command_liftTermElabM___rarg(x_22, x_2, x_3, x_21); if (lean_obj_tag(x_23) == 0) { @@ -11950,7 +11947,7 @@ x_34 = 0; x_35 = lean_usize_of_nat(x_27); lean_dec(x_27); x_36 = lean_box(0); -x_37 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__19(x_25, x_34, x_35, x_36, x_2, x_3, x_26); +x_37 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__18(x_25, x_34, x_35, x_36, x_2, x_3, x_26); lean_dec(x_25); if (lean_obj_tag(x_37) == 0) { @@ -12054,7 +12051,7 @@ x_61 = 0; x_62 = lean_usize_of_nat(x_52); lean_dec(x_52); x_63 = lean_box(0); -x_64 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__19(x_50, x_61, x_62, x_63, x_2, x_3, x_51); +x_64 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__18(x_50, x_61, x_62, x_63, x_2, x_3, x_51); lean_dec(x_50); if (lean_obj_tag(x_64) == 0) { @@ -12166,7 +12163,6 @@ lean_object* x_82; lean_object* x_83; x_82 = lean_alloc_closure((void*)(l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___lambda__3___boxed), 9, 2); lean_closure_set(x_82, 0, x_14); lean_closure_set(x_82, 1, x_1); -lean_inc(x_2); x_83 = l_Lean_Elab_Command_liftTermElabM___rarg(x_82, x_2, x_3, x_16); if (lean_obj_tag(x_83) == 0) { @@ -12215,7 +12211,7 @@ x_94 = 0; x_95 = lean_usize_of_nat(x_87); lean_dec(x_87); x_96 = lean_box(0); -x_97 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__19(x_85, x_94, x_95, x_96, x_2, x_3, x_86); +x_97 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__18(x_85, x_94, x_95, x_96, x_2, x_3, x_86); lean_dec(x_85); if (lean_obj_tag(x_97) == 0) { @@ -12319,7 +12315,7 @@ x_121 = 0; x_122 = lean_usize_of_nat(x_112); lean_dec(x_112); x_123 = lean_box(0); -x_124 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__19(x_110, x_121, x_122, x_123, x_2, x_3, x_111); +x_124 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__18(x_110, x_121, x_122, x_123, x_2, x_3, x_111); lean_dec(x_110); if (lean_obj_tag(x_124) == 0) { diff --git a/stage0/stdlib/Lean/Elab/Deriving/Hashable.c b/stage0/stdlib/Lean/Elab/Deriving/Hashable.c index c79a081cfa55..badbf80ed8ac 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/Hashable.c +++ b/stage0/stdlib/Lean/Elab/Deriving/Hashable.c @@ -34,6 +34,7 @@ static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMa static lean_object* l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__1___closed__1; static lean_object* l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__21; 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_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__18(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Deriving_Hashable_0__Lean_Elab_Deriving_Hashable_mkHashableInstanceCmds___closed__9; static lean_object* l___private_Lean_Elab_Deriving_Hashable_0__Lean_Elab_Deriving_Hashable_mkHashableInstanceCmds___closed__8; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -142,7 +143,6 @@ static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMa lean_object* l_Lean_MessageData_ofExpr(lean_object*); static lean_object* l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__3; static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__2; -lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__19(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_Hashable_mkMatch_mkAlts___spec__4___closed__14; lean_object* l_Lean_Syntax_node4(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_Hashable_mkMatch_mkAlts___spec__4___lambda__1___boxed(lean_object*, lean_object*); @@ -4398,7 +4398,6 @@ else lean_object* x_16; lean_object* x_17; x_16 = lean_alloc_closure((void*)(l___private_Lean_Elab_Deriving_Hashable_0__Lean_Elab_Deriving_Hashable_mkHashableInstanceCmds___boxed), 8, 1); lean_closure_set(x_16, 0, x_1); -lean_inc(x_2); x_17 = l_Lean_Elab_Command_liftTermElabM___rarg(x_16, x_2, x_3, x_7); if (lean_obj_tag(x_17) == 0) { @@ -4447,7 +4446,7 @@ x_28 = 0; x_29 = lean_usize_of_nat(x_21); lean_dec(x_21); x_30 = lean_box(0); -x_31 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__19(x_19, x_28, x_29, x_30, x_2, x_3, x_20); +x_31 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__18(x_19, x_28, x_29, x_30, x_2, x_3, x_20); lean_dec(x_19); if (lean_obj_tag(x_31) == 0) { @@ -4551,7 +4550,7 @@ x_55 = 0; x_56 = lean_usize_of_nat(x_46); lean_dec(x_46); x_57 = lean_box(0); -x_58 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__19(x_44, x_55, x_56, x_57, x_2, x_3, x_45); +x_58 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__18(x_44, x_55, x_56, x_57, x_2, x_3, x_45); lean_dec(x_44); if (lean_obj_tag(x_58) == 0) { diff --git a/stage0/stdlib/Lean/Elab/Deriving/Inhabited.c b/stage0/stdlib/Lean/Elab/Deriving/Inhabited.c index 939c0b30e874..7b336ab8c7b5 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/Inhabited.c +++ b/stage0/stdlib/Lean/Elab/Deriving/Inhabited.c @@ -8364,7 +8364,6 @@ x_8 = lean_alloc_closure((void*)(l___private_Lean_Elab_Deriving_Inhabited_0__Lea lean_closure_set(x_8, 0, x_1); lean_closure_set(x_8, 1, x_2); lean_closure_set(x_8, 2, x_7); -lean_inc(x_4); x_9 = l_Lean_Elab_Command_liftTermElabM___rarg(x_8, x_4, x_5, x_6); if (lean_obj_tag(x_9) == 0) { diff --git a/stage0/stdlib/Lean/Elab/Deriving/Nonempty.c b/stage0/stdlib/Lean/Elab/Deriving/Nonempty.c index 317d6ddf223b..bfd91b0d05f7 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/Nonempty.c +++ b/stage0/stdlib/Lean/Elab/Deriving/Nonempty.c @@ -2520,7 +2520,6 @@ lean_dec(x_4); x_10 = lean_array_uget(x_1, x_3); x_11 = lean_alloc_closure((void*)(l___private_Lean_Elab_Deriving_Nonempty_0__Lean_Elab_mkNonemptyInstance), 8, 1); lean_closure_set(x_11, 0, x_10); -lean_inc(x_5); x_12 = l_Lean_Elab_Command_liftTermElabM___rarg(x_11, x_5, x_6, x_7); if (lean_obj_tag(x_12) == 0) { diff --git a/stage0/stdlib/Lean/Elab/Deriving/Ord.c b/stage0/stdlib/Lean/Elab/Deriving/Ord.c index 7d644734af66..26f5bdc78951 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/Ord.c +++ b/stage0/stdlib/Lean/Elab/Deriving/Ord.c @@ -39,6 +39,7 @@ static lean_object* l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed_ static lean_object* l___private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkOrdInstanceCmds___closed__7; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__9; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__32; +lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__18(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__44; static lean_object* l_Lean_Elab_Deriving_Ord_mkMutualBlock___closed__2; lean_object* l_Lean_MessageData_ofList(lean_object*); @@ -164,7 +165,6 @@ lean_object* l_Array_append___rarg(lean_object*, lean_object*); static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__7___lambda__2___closed__9; lean_object* l_Lean_MessageData_ofExpr(lean_object*); static lean_object* l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__8; -lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__19(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__18; lean_object* l_Lean_Syntax_node4(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; @@ -5080,7 +5080,6 @@ else lean_object* x_16; lean_object* x_17; x_16 = lean_alloc_closure((void*)(l___private_Lean_Elab_Deriving_Ord_0__Lean_Elab_Deriving_Ord_mkOrdInstanceCmds___boxed), 8, 1); lean_closure_set(x_16, 0, x_1); -lean_inc(x_2); x_17 = l_Lean_Elab_Command_liftTermElabM___rarg(x_16, x_2, x_3, x_7); if (lean_obj_tag(x_17) == 0) { @@ -5129,7 +5128,7 @@ x_28 = 0; x_29 = lean_usize_of_nat(x_21); lean_dec(x_21); x_30 = lean_box(0); -x_31 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__19(x_19, x_28, x_29, x_30, x_2, x_3, x_20); +x_31 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__18(x_19, x_28, x_29, x_30, x_2, x_3, x_20); lean_dec(x_19); if (lean_obj_tag(x_31) == 0) { @@ -5233,7 +5232,7 @@ x_55 = 0; x_56 = lean_usize_of_nat(x_46); lean_dec(x_46); x_57 = lean_box(0); -x_58 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__19(x_44, x_55, x_56, x_57, x_2, x_3, x_45); +x_58 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__18(x_44, x_55, x_56, x_57, x_2, x_3, x_45); lean_dec(x_44); if (lean_obj_tag(x_58) == 0) { diff --git a/stage0/stdlib/Lean/Elab/Deriving/Repr.c b/stage0/stdlib/Lean/Elab/Deriving/Repr.c index c55970ff38a7..5735c8c8b546 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/Repr.c +++ b/stage0/stdlib/Lean/Elab/Deriving/Repr.c @@ -47,6 +47,7 @@ static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyFo LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_Repr_mkReprInstanceHandler(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__30; lean_object* l_Lean_Name_toString(lean_object*, uint8_t); +lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__18(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__6; lean_object* l_Lean_MessageData_ofList(lean_object*); static lean_object* l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__1___closed__4; @@ -233,7 +234,6 @@ lean_object* l_Lean_MessageData_ofExpr(lean_object*); static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__7; static lean_object* l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3780____closed__9; 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_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__19(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__9; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__21; lean_object* l_Lean_Syntax_node4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -6977,7 +6977,6 @@ else lean_object* x_16; lean_object* x_17; x_16 = lean_alloc_closure((void*)(l___private_Lean_Elab_Deriving_Repr_0__Lean_Elab_Deriving_Repr_mkReprInstanceCmds___boxed), 8, 1); lean_closure_set(x_16, 0, x_1); -lean_inc(x_2); x_17 = l_Lean_Elab_Command_liftTermElabM___rarg(x_16, x_2, x_3, x_7); if (lean_obj_tag(x_17) == 0) { @@ -7026,7 +7025,7 @@ x_28 = 0; x_29 = lean_usize_of_nat(x_21); lean_dec(x_21); x_30 = lean_box(0); -x_31 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__19(x_19, x_28, x_29, x_30, x_2, x_3, x_20); +x_31 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__18(x_19, x_28, x_29, x_30, x_2, x_3, x_20); lean_dec(x_19); if (lean_obj_tag(x_31) == 0) { @@ -7130,7 +7129,7 @@ x_55 = 0; x_56 = lean_usize_of_nat(x_46); lean_dec(x_46); x_57 = lean_box(0); -x_58 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__19(x_44, x_55, x_56, x_57, x_2, x_3, x_45); +x_58 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__18(x_44, x_55, x_56, x_57, x_2, x_3, x_45); lean_dec(x_44); if (lean_obj_tag(x_58) == 0) { diff --git a/stage0/stdlib/Lean/Elab/Deriving/SizeOf.c b/stage0/stdlib/Lean/Elab/Deriving/SizeOf.c index 10b52b4cd020..ba0de5df4ea9 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/SizeOf.c +++ b/stage0/stdlib/Lean/Elab/Deriving/SizeOf.c @@ -283,7 +283,6 @@ if (x_6 == 0) { uint8_t x_8; lean_object* x_9; lean_object* x_10; lean_dec(x_5); -lean_dec(x_2); x_8 = 0; x_9 = lean_box(x_8); x_10 = lean_alloc_ctor(0, 2, 0); @@ -300,7 +299,6 @@ lean_dec(x_5); if (x_12 == 0) { uint8_t x_13; lean_object* x_14; lean_object* x_15; -lean_dec(x_2); x_13 = 0; x_14 = lean_box(x_13); x_15 = lean_alloc_ctor(0, 2, 0); @@ -411,6 +409,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_SizeOf_mkSizeOfHandler___boxed(lea lean_object* x_5; x_5 = l_Lean_Elab_Deriving_SizeOf_mkSizeOfHandler(x_1, x_2, x_3, x_4); lean_dec(x_3); +lean_dec(x_2); lean_dec(x_1); return x_5; } diff --git a/stage0/stdlib/Lean/Elab/Do.c b/stage0/stdlib/Lean/Elab/Do.c index 07e4e667c38f..dccc137a095b 100644 --- a/stage0/stdlib/Lean/Elab/Do.c +++ b/stage0/stdlib/Lean/Elab/Do.c @@ -118,7 +118,6 @@ static lean_object* l_Lean_Elab_Term_Do_ToTerm_continueToTerm___closed__6; static lean_object* l___regBuiltin_Lean_Elab_Term_expandTermReturn_declRange___closed__1; static lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_checkLetArrowRHS___closed__1; static lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__27; -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__5___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___closed__21; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_hasExitPointPred_loop(lean_object*, lean_object*); @@ -149,11 +148,11 @@ static lean_object* l_Lean_Elab_Term_Do_ToTerm_reassignToTerm___closed__2; static lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_mkPureUnit___closed__8; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_getDoLetVars___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_Lean_Elab_Term_Do_mkJmp___spec__1___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__9; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Do_ToTerm_mkJoinPoint___spec__2(size_t, size_t, lean_object*); lean_object* l_Lean_MessageData_joinSep(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Do_ToTerm_returnToTerm___closed__45; static lean_object* l_Lean_Elab_Term_Do_ToTerm_returnToTerm___closed__4; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__11; static lean_object* l_Lean_Elab_Term_Do_pullExitPointsAux___lambda__1___closed__2; static lean_object* l_Lean_Elab_Term_Do_instInhabitedCode___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_ToCodeBlock_expandLiftMethodAux___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*); @@ -163,7 +162,6 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_expandTermTry_declRange___clos static lean_object* l_Lean_Elab_Term_Do_ToTerm_returnToTerm___closed__25; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_getDoLetArrowVars(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Do_getDoLetArrowVars___closed__3; -LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899_(lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_Do_ToCodeBlock_checkReassignable___spec__2___closed__4; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Do_ToCodeBlock_doLetElseToCode___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_Do_ToTerm_continueToTerm___closed__10; @@ -228,9 +226,9 @@ lean_object* l_instInhabitedReaderT___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_withFor___rarg(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_Do_mkJmp___spec__8___rarg(lean_object*); static lean_object* l_Lean_Elab_Term_Do_hasReturn___closed__1; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__10; static lean_object* l_Lean_Elab_Term_Do_ToTerm_returnToTerm___closed__23; LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__6; 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_Lean_isTracingEnabledFor___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*); static lean_object* l_Lean_Elab_Term_Do_getDoLetArrowVars___closed__1; @@ -275,6 +273,7 @@ LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_Do_ToCode LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_getLetPatDeclVars___boxed(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_doSeqToCode___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_forInUnsafe_loop___at_Lean_Elab_Term_Do_getDoLetRecVars___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_Do___hyg_36009____closed__13; LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at_Lean_Elab_Term_Do_insertVars___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_ToCodeBlock_expandLiftMethodAux___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_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*); @@ -305,7 +304,6 @@ static lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_doSeqToCode___closed__10; static lean_object* l_Lean_Elab_Term_Do_ToTerm_breakToTerm___closed__19; static lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__7; 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_EXPORT lean_object* l_Lean_Elab_Term_Do_getLetIdDeclVar___boxed(lean_object*); static lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_mkPureUnit___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_isMutableLet___boxed(lean_object*); static lean_object* l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__23; @@ -318,6 +316,7 @@ static lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___cl static lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_mkPureUnit___closed__9; static lean_object* l_Lean_Elab_Term_Do_ToTerm_matchNestedTermResult___closed__6; LEAN_EXPORT lean_object* l_List_forM___at_Lean_Elab_Term_Do_ToCodeBlock_mkForInBody___spec__4(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_Do___hyg_36009____closed__8; static lean_object* l_Lean_Elab_Term_Do_mkAuxDeclFor___rarg___lambda__5___closed__2; static lean_object* l_Lean_Elab_Term_Do_CodeBlocl_toMessageData_loop___closed__16; lean_object* l_Lean_Syntax_node5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -337,7 +336,6 @@ static lean_object* l_panic___at_Lean_Elab_Term_Do_ToTerm_returnToTerm___spec__1 static lean_object* l_Lean_Elab_Term_Do_ToTerm_breakToTerm___closed__15; static lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_doReturnToCode___closed__1; static lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__8; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__9; static lean_object* l_panic___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_destructTuple_destruct___spec__1___closed__2; lean_object* l_Lean_RBNode_balRight___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Do_ToCodeBlock_doTryToCode___spec__2___closed__1; @@ -392,7 +390,6 @@ static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean static lean_object* l_Lean_Elab_Term_Do_ToTerm_continueToTerm___closed__20; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___spec__1(lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_doTryToCode___lambda__2___closed__7; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__7; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_expandTermFor_declRange(lean_object*); static lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___closed__16; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_hasReturn___lambda__1___boxed(lean_object*); @@ -411,7 +408,6 @@ static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Do_ToTerm_mkJoinP static lean_object* l_Lean_Elab_Term_Do_pullExitPointsAux___lambda__1___closed__5; static lean_object* l_Lean_Elab_Term_Do_ToTerm_declToTerm___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_expandTermTry(lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_expandLiftMethod(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_doSeqToCode___closed__2; static lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2___closed__6; @@ -568,6 +564,7 @@ lean_object* l_Lean_Elab_Term_exprToSyntax(lean_object*, lean_object*, lean_obje lean_object* l_Lean_Syntax_mkStrLit(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_letDeclArgHasBinders___closed__2; lean_object* l_Lean_Meta_SavedState_restore(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__11; static lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_extractBind___closed__2; static lean_object* l_Lean_Elab_Term_Do_ToTerm_matchNestedTermResult___closed__12; lean_object* l_Lean_Elab_Term_expandOptType(lean_object*, lean_object*); @@ -625,7 +622,6 @@ LEAN_EXPORT lean_object* l_panic___at___private_Lean_Elab_Do_0__Lean_Elab_Term_D static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Do_ToTerm_mkJoinPoint___spec__4___closed__2; static lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_mkUnit___closed__2; 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*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__3; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Do_mkSimpleJmp___spec__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Do_ToTerm_breakToTerm___closed__10; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_checkReassignable(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -637,6 +633,7 @@ LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Elab_Term_Do_attach LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_Do_mkJmp___spec__7(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_expandTermFor___closed__2; static lean_object* l_Lean_Elab_Term_Do_ToTerm_breakToTerm___closed__6; +LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_getLetEqnsDeclVars(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_expandTermTry_declRange___closed__6; 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*); static lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_liftMethodDelimiter___closed__10; @@ -700,6 +697,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_toDoElem(lea static lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_doTryToCode___lambda__3___closed__2; static lean_object* l_Lean_Elab_Term_Do_ToTerm_continueToTerm___closed__9; static lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_letDeclArgHasBinders___closed__6; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__4; static lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_mkUnit___closed__9; static lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__2___closed__4; static lean_object* l_Lean_Elab_Term_Do_ToTerm_breakToTerm___closed__9; @@ -765,7 +763,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_CodeBlocl_toMessageData_loop(lean_o LEAN_EXPORT uint8_t l___private_Lean_Elab_Do_0__Lean_Elab_Term_liftMethodDelimiter(lean_object*); static lean_object* l_Lean_Elab_Term_Do_ToTerm_breakToTerm___closed__3; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__2; static lean_object* l_Lean_Elab_Term_Do_ToTerm_mkNestedKind___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Term_Do_elabDo___closed__2; static lean_object* l_Lean_Elab_Term_Do_ToTerm_breakToTerm___closed__13; @@ -820,10 +817,11 @@ static lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f uint8_t l_Lean_Syntax_isIdent(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_panic___at_Lean_Elab_Term_Do_ToTerm_matchNestedTermResult___spec__1___closed__2; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__7; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__2; static lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_mkUnit___closed__16; static lean_object* l_Lean_Elab_Term_Do_ToTerm_continueToTerm___closed__34; static lean_object* l___regBuiltin_Lean_Elab_Term_expandTermUnless_declRange___closed__7; -LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_getLetIdDeclVar(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_Do_mkMatch___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_instInhabitedAlt(lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Do_ToCodeBlock_doTryToCode___spec__2___closed__7; @@ -832,7 +830,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_doReturnToCode___lambda static lean_object* l_Lean_Elab_Term_Do_ToTerm_returnToTerm___closed__33; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_getTryCatchUpdatedVars___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_mkAuxDeclFor___rarg___lambda__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__13; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_ToTerm_matchNestedTermResult(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Do_ToTerm_seqToTerm___closed__7; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Do_convertTerminalActionIntoJmp_loop___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -886,6 +883,7 @@ LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Term_Do_mkJmp static lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_doTryToCode___lambda__2___closed__5; LEAN_EXPORT uint8_t l___private_Lean_Elab_Do_0__Lean_Elab_Term_letDeclHasBinders(lean_object*); static lean_object* l_Lean_Elab_Term_Do_ToTerm_breakToTerm___closed__24; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__3; static lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_doTryToCode___lambda__3___closed__1; lean_object* l_Lean_Meta_whnfEasyCases___at_Lean_Meta_whnfCore_go___spec__3(uint8_t, uint8_t, 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_destructTuple_destruct___closed__13; @@ -908,6 +906,7 @@ static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_Do_ToCodeBlock LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_Do_ToCodeBlock_checkReassignable___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_EXPORT lean_object* l_Lean_Elab_Term_Do_annotate(lean_object*); 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*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_getLetIdDeclVars(lean_object*); static lean_object* l_Lean_Elab_Term_Do_annotate___rarg___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_mkAuxDeclFor(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_homogenize(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -925,7 +924,6 @@ static lean_object* l_Lean_Elab_Term_Do_mkAuxDeclFor___rarg___lambda__2___closed static lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___closed__3; LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__2___boxed(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__5___lambda__2___closed__15; -LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_getLetEqnsDeclVar___boxed(lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__5___lambda__2___closed__19; static lean_object* l___regBuiltin_Lean_Elab_Term_elabLiftMethod___closed__3; static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_Do_ToCodeBlock_doTryToCode___spec__10___closed__11; @@ -958,7 +956,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_mkSimpleJmp___boxed(lean_object*, l static lean_object* l_Lean_Elab_Term_Do_CodeBlocl_toMessageData_loop___closed__23; lean_object* lean_simp_macro_scopes(lean_object*); LEAN_EXPORT lean_object* l_Array_sequenceMap___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___spec__1(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__1; static lean_object* l_Lean_Elab_Term_Do_ToTerm_mkJoinPoint___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_hasReturn(lean_object*); static lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_doTryToCode___lambda__2___closed__3; @@ -978,6 +975,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_mkAuxDeclFor___rarg___lambda__15___ static lean_object* l_Lean_Elab_Term_Do_ToTerm_seqToTerm___closed__6; static lean_object* l_Lean_Elab_Term_Do_ToTerm_breakToTerm___closed__23; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_ToCodeBlock_expandLiftMethodAux___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_EXPORT lean_object* l_Lean_Elab_Term_Do_getLetEqnsDeclVars___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_expandDoIf_x3f___lambda__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_RBNode_erase___at_Lean_Elab_Term_Do_eraseVars___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Do_ToTerm_matchNestedTermResult___closed__22; @@ -1018,6 +1016,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_doIfToCode(lean_object* static lean_object* l_Lean_Elab_Term_Do_CodeBlocl_toMessageData_loop___closed__8; static lean_object* l_Lean_Elab_Term_Do_ToTerm_returnToTerm___closed__43; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_getLetDeclVars(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_Do___hyg_36009____closed__1; static lean_object* l_Lean_Elab_Term_Do_mkSimpleJmp___closed__8; static lean_object* l_Lean_Elab_Term_Do_ToTerm_continueToTerm___closed__3; static lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_ensureEOS___closed__2; @@ -1057,6 +1056,7 @@ static lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Do_0__Lea static lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__19; static lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_doTryToCode___lambda__4___closed__3; static lean_object* l_Lean_Elab_Term_Do_mkUnless___closed__1; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__10; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_Do_ToCodeBlock_doTryToCode___spec__8(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__22; static lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___closed__20; @@ -1095,6 +1095,7 @@ static lean_object* l_Lean_Elab_Term_Do_getDoHaveVars___closed__1; static lean_object* l_Lean_Elab_Term_Do_ToTerm_mkJoinPoint___closed__4; lean_object* l_Lean_Expr_getAppFn(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_ToCodeBlock_expandLiftMethodAux___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*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__5; static lean_object* l_Lean_Elab_Term_Do_ToTerm_declToTerm___closed__6; uint8_t l_Lean_RBNode_isBlack___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_attachJP(lean_object*, lean_object*); @@ -1113,7 +1114,6 @@ LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_Do_ToCode static lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_destructTuple_destruct___closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandTermFor(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Do_getDoHaveVars___closed__6; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__12; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_mkAuxDeclFor___rarg___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_doTryToCode___lambda__5(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_Lean_Elab_Term_Do_ToCodeBlock_doSeqToCode___closed__14; @@ -1153,7 +1153,6 @@ static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Do_0__Lean LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabLiftMethod(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_mkBreak(lean_object*); lean_object* l_Array_back___rarg(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__6; static lean_object* l_Lean_Elab_Term_Do_ToTerm_continueToTerm___closed__16; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_Do_ToCodeBlock_checkReassignable___spec__2___closed__5; lean_object* l_List_reverse___rarg(lean_object*); @@ -1179,6 +1178,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_ToTerm_mkUVarTuple(lean_object*, le LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_ToCodeBlock_expandLiftMethodAux___spec__10(lean_object*, uint8_t, 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*); LEAN_EXPORT lean_object* l_List_beq___at___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_ToCodeBlock_expandLiftMethodAux___spec__12___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_ToTerm_reassignToTerm(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__12; static lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___closed__11; static lean_object* l_Lean_Elab_Term_Do_ToTerm_mkIte___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Term_elabLiftMethod_declRange___closed__6; @@ -1223,7 +1223,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_extendUpdatedVarsAux(lean_object*, LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Do_ToCodeBlock_doTryToCode___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_Lean_Elab_Term_Do_getDoLetArrowVars___closed__4; static lean_object* l_Lean_Elab_Term_Do_ToTerm_toTerm_go___closed__1; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__4; static lean_object* l_Lean_Elab_Term_Do_ToTerm_breakToTerm___closed__16; LEAN_EXPORT lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_ToCodeBlock_expandLiftMethodAux___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___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_ToCodeBlock_expandLiftMethodAux(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*); @@ -1350,7 +1349,6 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabLiftMethod_declRange___clo LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_doReturnToCode___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_Do_ToTerm_continueToTerm___closed__35; static lean_object* l___private_Lean_Elab_Do_0__Lean_Elab_Term_Do_mkUnit___closed__12; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__8; static lean_object* l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__26; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_convertTerminalActionIntoJmp_loop(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_Do_mkMatch___spec__2(lean_object*, size_t, size_t, lean_object*); @@ -1383,6 +1381,7 @@ static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_Do_ToCodeBlock_ static lean_object* l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__9; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Do_pullExitPointsAux___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_Lean_Elab_Term_Do_ToCodeBlock_doForToCode___lambda__3___closed__26; +LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009_(lean_object*); LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Elab_Do_0__Lean_Elab_Term_getDoSeqElems___spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Do_ToTerm_returnToTerm___closed__21; static lean_object* l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___closed__12; @@ -1395,6 +1394,7 @@ static lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Do_0__Lea static lean_object* l_Lean_Elab_Term_Do_mkAuxDeclFor___rarg___lambda__8___closed__5; static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_Do_CodeBlocl_toMessageData_loop___spec__2___closed__2; lean_object* l_List_toArrayAux___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_getLetIdDeclVars___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_ToCodeBlock_withFor(lean_object*); lean_object* l_ReaderT_pure___at_Lean_Elab_liftMacroM___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_Code_getRef_x3f(lean_object*); @@ -12034,20 +12034,34 @@ lean_dec(x_1); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_getLetIdDeclVar(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_getLetIdDeclVars(lean_object* x_1) { _start: { -lean_object* x_2; lean_object* x_3; +lean_object* x_2; lean_object* x_3; uint8_t x_4; x_2 = lean_unsigned_to_nat(0u); x_3 = l_Lean_Syntax_getArg(x_1, x_2); -return x_3; +x_4 = l_Lean_Syntax_isIdent(x_3); +if (x_4 == 0) +{ +lean_object* x_5; +lean_dec(x_3); +x_5 = l_Lean_Elab_Term_Do_instInhabitedAlt___rarg___closed__1; +return x_5; } +else +{ +lean_object* x_6; lean_object* x_7; +x_6 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_extractBind___lambda__1___closed__3; +x_7 = lean_array_push(x_6, x_3); +return x_7; } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_getLetIdDeclVar___boxed(lean_object* x_1) { +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_getLetIdDeclVars___boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Elab_Term_Do_getLetIdDeclVar(x_1); +x_2 = l_Lean_Elab_Term_Do_getLetIdDeclVars(x_1); lean_dec(x_1); return x_2; } @@ -12170,20 +12184,34 @@ lean_dec(x_1); return x_9; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_getLetEqnsDeclVar(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_getLetEqnsDeclVars(lean_object* x_1) { _start: { -lean_object* x_2; lean_object* x_3; +lean_object* x_2; lean_object* x_3; uint8_t x_4; x_2 = lean_unsigned_to_nat(0u); x_3 = l_Lean_Syntax_getArg(x_1, x_2); -return x_3; +x_4 = l_Lean_Syntax_isIdent(x_3); +if (x_4 == 0) +{ +lean_object* x_5; +lean_dec(x_3); +x_5 = l_Lean_Elab_Term_Do_instInhabitedAlt___rarg___closed__1; +return x_5; +} +else +{ +lean_object* x_6; lean_object* x_7; +x_6 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_extractBind___lambda__1___closed__3; +x_7 = lean_array_push(x_6, x_3); +return x_7; +} } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_getLetEqnsDeclVar___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_getLetEqnsDeclVars___boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Elab_Term_Do_getLetEqnsDeclVar(x_1); +x_2 = l_Lean_Elab_Term_Do_getLetEqnsDeclVars(x_1); lean_dec(x_1); return x_2; } @@ -12273,35 +12301,33 @@ return x_19; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +lean_object* x_20; lean_object* x_21; 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 = l_Lean_Syntax_getArg(x_10, x_9); +x_20 = l_Lean_Elab_Term_Do_getLetEqnsDeclVars(x_10); lean_dec(x_10); -x_21 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_extractBind___lambda__1___closed__3; -x_22 = lean_array_push(x_21, 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_8); -return x_23; +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_8); +return x_21; } } else { -lean_object* x_24; +lean_object* x_22; lean_dec(x_11); -x_24 = l_Lean_Elab_Term_Do_getLetPatDeclVars(x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_22 = l_Lean_Elab_Term_Do_getLetPatDeclVars(x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_10); -return x_24; +return x_22; } } else { -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_dec(x_11); lean_dec(x_7); lean_dec(x_6); @@ -12309,14 +12335,12 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_25 = l_Lean_Syntax_getArg(x_10, x_9); +x_23 = l_Lean_Elab_Term_Do_getLetIdDeclVars(x_10); lean_dec(x_10); -x_26 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_extractBind___lambda__1___closed__3; -x_27 = lean_array_push(x_26, 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_8); -return x_28; +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_8); +return x_24; } } } @@ -12356,26 +12380,27 @@ static lean_object* _init_l_Lean_Elab_Term_Do_getHaveIdLhsVar___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("termThis", 8); +x_1 = lean_mk_string_from_bytes("this", 4); return x_1; } } static lean_object* _init_l_Lean_Elab_Term_Do_getHaveIdLhsVar___closed__2() { _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__1; -x_2 = l_Lean_Elab_Term_Do_getHaveIdLhsVar___closed__1; -x_3 = l_Lean_Name_mkStr2(x_1, x_2); -return x_3; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_Term_Do_getHaveIdLhsVar___closed__1; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; } } static lean_object* _init_l_Lean_Elab_Term_Do_getHaveIdLhsVar___closed__3() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("this", 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_Do_getHaveIdLhsVar___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } } LEAN_EXPORT lean_object* l_Lean_Elab_Term_Do_getHaveIdLhsVar(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { @@ -12396,27 +12421,37 @@ return x_12; } else { -lean_object* x_13; uint8_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_13; uint8_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; x_13 = lean_ctor_get(x_6, 5); lean_inc(x_13); -lean_dec(x_6); x_14 = 0; x_15 = l_Lean_SourceInfo_fromRef(x_13, x_14); -x_16 = lean_st_ref_get(x_7, x_8); -x_17 = lean_ctor_get(x_16, 1); -lean_inc(x_17); -lean_dec(x_16); -x_18 = l_Lean_Elab_Term_Do_getHaveIdLhsVar___closed__3; -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_Elab_Term_Do_getHaveIdLhsVar___closed__2; -x_21 = l_Lean_Syntax_node1(x_15, x_20, 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_17); -return x_22; +x_16 = lean_ctor_get(x_6, 10); +lean_inc(x_16); +lean_dec(x_6); +x_17 = lean_st_ref_get(x_7, x_8); +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_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_Term_Do_getHaveIdLhsVar___closed__3; +x_23 = l_Lean_addMacroScope(x_21, x_22, x_16); +x_24 = lean_box(0); +x_25 = l_Lean_Elab_Term_Do_getHaveIdLhsVar___closed__2; +x_26 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_26, 0, x_15); +lean_ctor_set(x_26, 1, x_25); +lean_ctor_set(x_26, 2, x_23); +lean_ctor_set(x_26, 3, x_24); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_19); +return x_27; } } } @@ -12976,7 +13011,7 @@ return x_18; } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +lean_object* x_19; lean_object* x_20; lean_dec(x_11); lean_dec(x_7); lean_dec(x_6); @@ -12984,14 +13019,12 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_19 = l_Lean_Syntax_getArg(x_10, x_9); +x_19 = l_Lean_Elab_Term_Do_getLetIdDeclVars(x_10); lean_dec(x_10); -x_20 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_extractBind___lambda__1___closed__3; -x_21 = lean_array_push(x_20, 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_8); -return x_22; +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_8); +return x_20; } } } @@ -16034,7 +16067,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(776u); +x_3 = lean_unsigned_to_nat(782u); 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); @@ -17223,7 +17256,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(920u); +x_3 = lean_unsigned_to_nat(926u); x_4 = lean_unsigned_to_nat(24u); 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); @@ -18010,7 +18043,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(929u); +x_3 = lean_unsigned_to_nat(935u); x_4 = lean_unsigned_to_nat(24u); 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); @@ -18254,7 +18287,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(933u); +x_3 = lean_unsigned_to_nat(939u); x_4 = lean_unsigned_to_nat(24u); 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); @@ -18820,7 +18853,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(941u); +x_3 = lean_unsigned_to_nat(947u); x_4 = lean_unsigned_to_nat(24u); 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); @@ -18914,7 +18947,7 @@ static lean_object* _init_l_Lean_Elab_Term_Do_ToTerm_breakToTerm___closed__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; 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(945u); +x_3 = lean_unsigned_to_nat(951u); x_4 = lean_unsigned_to_nat(24u); 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); @@ -19583,7 +19616,7 @@ static lean_object* _init_l_Lean_Elab_Term_Do_ToTerm_actionTerminalToTerm___clos 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__11; -x_3 = lean_unsigned_to_nat(956u); +x_3 = lean_unsigned_to_nat(962u); x_4 = lean_unsigned_to_nat(24u); 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); @@ -23033,7 +23066,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(1096u); +x_3 = lean_unsigned_to_nat(1102u); 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); @@ -23202,7 +23235,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(1152u); +x_3 = lean_unsigned_to_nat(1158u); 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); @@ -40100,7 +40133,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(1670u); +x_1 = lean_unsigned_to_nat(1676u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -40112,7 +40145,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(1678u); +x_1 = lean_unsigned_to_nat(1684u); x_2 = lean_unsigned_to_nat(84u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -40140,7 +40173,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(1670u); +x_1 = lean_unsigned_to_nat(1676u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -40152,7 +40185,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(1670u); +x_1 = lean_unsigned_to_nat(1676u); x_2 = lean_unsigned_to_nat(36u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -40198,7 +40231,7 @@ x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__1() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -40208,27 +40241,27 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__2() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__1; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__1; x_2 = l___regBuiltin_Lean_Elab_Term_elabLiftMethod___closed__3; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__3() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____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_Do___hyg_35899____closed__2; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__2; x_2 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_getDoSeqElems___closed__3; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__4() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__4() { _start: { lean_object* x_1; @@ -40236,17 +40269,17 @@ x_1 = lean_mk_string_from_bytes("initFn", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__5() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__3; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__4; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__3; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__6() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__6() { _start: { lean_object* x_1; @@ -40254,47 +40287,47 @@ x_1 = lean_mk_string_from_bytes("_@", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__7() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__5; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__6; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__5; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__8() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__7; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__7; x_2 = l___private_Lean_Elab_Do_0__Lean_Elab_Term_getDoSeqElems___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__9() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__8; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__8; x_2 = l___regBuiltin_Lean_Elab_Term_elabLiftMethod___closed__3; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__10() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__9; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__9; x_2 = l___regBuiltin_Lean_Elab_Term_Do_elabDo___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__11() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__11() { _start: { lean_object* x_1; @@ -40302,33 +40335,33 @@ x_1 = lean_mk_string_from_bytes("_hyg", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__12() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__10; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__11; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__10; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__11; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__13() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__12; -x_2 = lean_unsigned_to_nat(35899u); +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__12; +x_2 = lean_unsigned_to_nat(36009u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; x_2 = l_Lean_Elab_Term_Do_elabDo___closed__1; x_3 = 0; -x_4 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__13; +x_4 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__13; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } @@ -40449,7 +40482,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(1689u); +x_1 = lean_unsigned_to_nat(1695u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -40461,7 +40494,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(1689u); +x_1 = lean_unsigned_to_nat(1695u); x_2 = lean_unsigned_to_nat(57u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -40489,7 +40522,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(1689u); +x_1 = lean_unsigned_to_nat(1695u); x_2 = lean_unsigned_to_nat(4u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -40501,7 +40534,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(1689u); +x_1 = lean_unsigned_to_nat(1695u); x_2 = lean_unsigned_to_nat(17u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -40609,7 +40642,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(1692u); +x_1 = lean_unsigned_to_nat(1698u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -40621,7 +40654,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(1692u); +x_1 = lean_unsigned_to_nat(1698u); x_2 = lean_unsigned_to_nat(57u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -40649,7 +40682,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(1692u); +x_1 = lean_unsigned_to_nat(1698u); x_2 = lean_unsigned_to_nat(4u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -40661,7 +40694,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(1692u); +x_1 = lean_unsigned_to_nat(1698u); x_2 = lean_unsigned_to_nat(17u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -40769,7 +40802,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(1695u); +x_1 = lean_unsigned_to_nat(1701u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -40781,7 +40814,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(1695u); +x_1 = lean_unsigned_to_nat(1701u); x_2 = lean_unsigned_to_nat(63u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -40809,7 +40842,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(1695u); +x_1 = lean_unsigned_to_nat(1701u); x_2 = lean_unsigned_to_nat(4u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -40821,7 +40854,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(1695u); +x_1 = lean_unsigned_to_nat(1701u); x_2 = lean_unsigned_to_nat(20u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -40929,7 +40962,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(1698u); +x_1 = lean_unsigned_to_nat(1704u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -40941,7 +40974,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(1698u); +x_1 = lean_unsigned_to_nat(1704u); x_2 = lean_unsigned_to_nat(63u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -40969,7 +41002,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(1698u); +x_1 = lean_unsigned_to_nat(1704u); x_2 = lean_unsigned_to_nat(4u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -40981,7 +41014,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(1698u); +x_1 = lean_unsigned_to_nat(1704u); x_2 = lean_unsigned_to_nat(20u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -42356,33 +42389,33 @@ 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); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__1 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__1(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__1); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__2 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__2(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__2); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__3 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__3(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__3); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__4 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__4(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__4); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__5 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__5(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__5); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__6 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__6(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__6); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__7 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__7(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__7); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__8 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__8(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__8); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__9 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__9(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__9); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__10 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__10(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__10); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__11 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__11(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__11); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__12 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__12(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__12); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__13 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__13(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899____closed__13); -res = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_35899_(lean_io_mk_world()); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__1 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__1(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__1); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__2 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__2(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__2); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__3 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__3(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__3); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__4 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__4(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__4); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__5 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__5(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__5); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__6 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__6(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__6); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__7 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__7(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__7); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__8 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__8(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__8); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__9 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__9(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__9); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__10 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__10(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__10); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__11 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__11(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__11); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__12 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__12(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__12); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__13 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__13(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009____closed__13); +res = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Do___hyg_36009_(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 7e930318f8d2..bdc0d1196996 100644 --- a/stage0/stdlib/Lean/Elab/ElabRules.c +++ b/stage0/stdlib/Lean/Elab/ElabRules.c @@ -112,6 +112,7 @@ static lean_object* l_Lean_Elab_Command_elabElab___lambda__1___closed__4; lean_object* l_Lean_Name_mkStr3(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*); static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabElabRulesAux___spec__1___rarg___closed__2; +lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabCommand___spec__14(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_of_nat(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*); static lean_object* l_Lean_Elab_Command_elabElab___lambda__1___closed__9; @@ -119,6 +120,7 @@ static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesA static lean_object* l_Lean_Elab_Command_elabElab___lambda__1___closed__10; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__73; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__2___closed__10; +lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__17___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*); static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__16; 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*); @@ -145,7 +147,6 @@ lean_object* l_Lean_Elab_getBetterRef(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__31; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__90; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__82; -lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__18___rarg(lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabElab(lean_object*); static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__47; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__11; @@ -204,6 +205,7 @@ lean_object* l_Lean_evalOptPrio(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_node4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__36; static lean_object* l___regBuiltin_Lean_Elab_Command_elabElab_declRange___closed__6; +lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabCommand___spec__15(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_setArg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__40; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__78; @@ -211,7 +213,6 @@ static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__3 static lean_object* l_Lean_Elab_Command_elabElabRules___lambda__8___closed__2; uint8_t l_Lean_Syntax_isQuot(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5___lambda__1(lean_object*); -lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabCommand___spec__16(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabElabRulesAux___spec__1___rarg___closed__1; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__21; 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*); @@ -239,7 +240,6 @@ lean_object* l_Lean_mkSepArray(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_withExpectedType___closed__2; lean_object* lean_array_set(lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabElabRules_declRange___closed__2; -lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabCommand___spec__15(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabElab_declRange___closed__2; static lean_object* l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__69; uint8_t l_Lean_Syntax_isNone(lean_object*); @@ -2293,7 +2293,7 @@ x_208 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabElabRulesAux___spec__5 x_209 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_209, 0, x_207); lean_ctor_set(x_209, 1, x_208); -x_210 = l_Lean_throwError___at_Lean_Elab_Command_elabCommand___spec__16(x_209, x_8, x_9, x_10); +x_210 = l_Lean_throwError___at_Lean_Elab_Command_elabCommand___spec__15(x_209, x_8, x_9, x_10); lean_dec(x_9); return x_210; } @@ -3683,7 +3683,7 @@ x_621 = l_Lean_Elab_Command_elabElabRulesAux___lambda__1___closed__89; x_622 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_622, 0, x_620); lean_ctor_set(x_622, 1, x_621); -x_623 = l_Lean_throwErrorAt___at_Lean_Elab_Command_elabCommand___spec__15(x_615, x_622, x_8, x_9, x_10); +x_623 = l_Lean_throwErrorAt___at_Lean_Elab_Command_elabCommand___spec__14(x_615, x_622, x_8, x_9, x_10); return x_623; } else @@ -4503,7 +4503,7 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_16 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__18___rarg(x_11); +x_16 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__17___rarg(x_11); return x_16; } else @@ -4586,7 +4586,7 @@ lean_dec(x_7); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_16 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__18___rarg(x_10); +x_16 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__17___rarg(x_10); return x_16; } else @@ -4951,7 +4951,7 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_16 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__18___rarg(x_11); +x_16 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__17___rarg(x_11); return x_16; } else @@ -5044,7 +5044,7 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_16 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__18___rarg(x_10); +x_16 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__17___rarg(x_10); return x_16; } else @@ -5110,7 +5110,7 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); -x_13 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__18___rarg(x_8); +x_13 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__17___rarg(x_8); return x_13; } else @@ -5137,7 +5137,7 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_2); -x_20 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__18___rarg(x_8); +x_20 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__17___rarg(x_8); return x_20; } else @@ -5163,7 +5163,7 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_2); -x_26 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__18___rarg(x_8); +x_26 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__17___rarg(x_8); return x_26; } else @@ -5214,7 +5214,7 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); -x_39 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__18___rarg(x_8); +x_39 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__17___rarg(x_8); return x_39; } else @@ -5264,7 +5264,7 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); -x_12 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__18___rarg(x_7); +x_12 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__17___rarg(x_7); return x_12; } else @@ -5284,7 +5284,7 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); -x_17 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__18___rarg(x_7); +x_17 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__17___rarg(x_7); return x_17; } else @@ -5358,7 +5358,7 @@ lean_object* x_7; lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_7 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__18___rarg(x_4); +x_7 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__17___rarg(x_4); return x_7; } else @@ -5380,7 +5380,7 @@ lean_dec(x_9); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_13 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__18___rarg(x_4); +x_13 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__17___rarg(x_4); return x_13; } else @@ -5398,7 +5398,7 @@ lean_dec(x_14); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_17 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__18___rarg(x_4); +x_17 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__17___rarg(x_4); return x_17; } else diff --git a/stage0/stdlib/Lean/Elab/Frontend.c b/stage0/stdlib/Lean/Elab/Frontend.c index 53228741fd4d..b780b91028f2 100644 --- a/stage0/stdlib/Lean/Elab/Frontend.c +++ b/stage0/stdlib/Lean/Elab/Frontend.c @@ -32,7 +32,6 @@ lean_object* lean_array_push(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_getParserState___rarg(lean_object*, lean_object*); lean_object* l_Lean_Elab_Command_mkState(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_runFrontend___lambda__1___boxed(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*); LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_setCommandState___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_processCommands(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Frontend___hyg_1025____closed__5; @@ -63,6 +62,7 @@ static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Frontend___hyg_1025____c LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_elabCommandAtFrontend(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Frontend___hyg_1025____closed__6; static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Frontend___hyg_1025____closed__3; +lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(lean_object*, lean_object*); lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean_object*); lean_object* l_Lean_MessageLog_toList(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Frontend_runCommandElabM___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -2517,7 +2517,7 @@ lean_dec(x_30); x_31 = lean_ctor_get(x_22, 0); lean_dec(x_31); x_32 = l_Lean_Elab_runFrontend___closed__2; -x_33 = l_Lean_Option_get___at_Std_Format_pretty_x27___spec__1(x_2, x_32); +x_33 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_2, x_32); x_34 = !lean_is_exclusive(x_26); if (x_34 == 0) { @@ -2574,7 +2574,7 @@ lean_inc(x_49); lean_inc(x_48); lean_dec(x_22); x_52 = l_Lean_Elab_runFrontend___closed__2; -x_53 = l_Lean_Option_get___at_Std_Format_pretty_x27___spec__1(x_2, x_52); +x_53 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_2, x_52); x_54 = lean_ctor_get(x_50, 0); lean_inc(x_54); x_55 = lean_ctor_get(x_50, 1); diff --git a/stage0/stdlib/Lean/Elab/GenInjective.c b/stage0/stdlib/Lean/Elab/GenInjective.c index 6d2d1187441f..a77b026097e1 100644 --- a/stage0/stdlib/Lean/Elab/GenInjective.c +++ b/stage0/stdlib/Lean/Elab/GenInjective.c @@ -957,6 +957,7 @@ x_11 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabGenInjectiveTheorems__ lean_closure_set(x_11, 0, x_9); x_12 = l_Lean_Elab_Command_liftTermElabM___rarg(x_11, x_2, x_3, x_10); lean_dec(x_3); +lean_dec(x_2); return x_12; } else diff --git a/stage0/stdlib/Lean/Elab/Inductive.c b/stage0/stdlib/Lean/Elab/Inductive.c index 1421e7e2ba11..5555f74f03df 100644 --- a/stage0/stdlib/Lean/Elab/Inductive.c +++ b/stage0/stdlib/Lean/Elab/Inductive.c @@ -18,7 +18,6 @@ static lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyCo LEAN_EXPORT lean_object* l_List_forM___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_checkResultingUniverses___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*); LEAN_EXPORT lean_object* l_Lean_assignLevelMVar___at_Lean_Elab_Command_accLevelAtCtor___spec__2(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_const___override(lean_object*, lean_object*); -uint8_t l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_isDomainDefEq___boxed(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_applyComputedFields___lambda__1___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*); @@ -415,7 +414,6 @@ LEAN_EXPORT lean_object* l_List_forM___at___private_Lean_Elab_Inductive_0__Lean_ lean_object* lean_array_to_list(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_updateResultingUniverse(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_accLevelAtCtor___lambda__2___closed__2; -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_checkParamsAndResultType___lambda__1___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___lambda__2___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_Inductive_0__Lean_Elab_Command_mkAuxConstructions___spec__11(uint8_t, uint8_t, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -495,6 +493,7 @@ static lean_object* l_Lean_Elab_Command_checkResultingUniverse___closed__3; 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 lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_mkAuxConstructions___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkRecOn___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_mkAuxConstructions___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Inductive___hyg_8086____closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_replaceArrowBinderNames_go(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_updateResultingUniverse___closed__1; @@ -652,6 +651,7 @@ LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_Inducti LEAN_EXPORT lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_mkInductiveDecl___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_Array_forInUnsafe_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_applyComputedFields___spec__1___lambda__1___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_fixedIndicesToParams(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_Elab_Command_accLevelAtCtor___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofLevel(lean_object*); static lean_object* l_Lean_mkRecOn___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_mkAuxConstructions___spec__1___closed__11; @@ -6547,7 +6547,7 @@ lean_inc(x_37); x_38 = lean_ctor_get(x_36, 1); lean_inc(x_38); lean_dec(x_36); -x_39 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_37, x_4, x_5, x_6, x_7, x_38); +x_39 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_37, x_4, x_5, x_6, x_7, x_38); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); @@ -6707,7 +6707,7 @@ lean_inc(x_77); x_78 = lean_ctor_get(x_76, 1); lean_inc(x_78); lean_dec(x_76); -x_79 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_77, x_4, x_5, x_6, x_7, x_78); +x_79 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_77, x_4, x_5, x_6, x_7, x_78); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); @@ -15059,7 +15059,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_checkResultingUniverse(lean_object* lean_object* x_9; lean_object* x_10; uint8_t x_11; x_9 = lean_ctor_get(x_6, 2); x_10 = l_Lean_Elab_Command_checkResultingUniverse___closed__1; -x_11 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_9, x_10); +x_11 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_9, x_10); if (x_11 == 0) { lean_object* x_12; lean_object* x_13; @@ -33422,6 +33422,7 @@ lean_closure_set(x_45, 0, x_44); lean_closure_set(x_45, 1, x_41); x_46 = l_Lean_Elab_Command_liftTermElabM___rarg(x_45, x_3, x_4, x_38); lean_dec(x_4); +lean_dec(x_3); return x_46; } else @@ -33436,6 +33437,7 @@ lean_closure_set(x_49, 0, x_48); lean_closure_set(x_49, 1, x_41); x_50 = l_Lean_Elab_Command_liftTermElabM___rarg(x_49, x_3, x_4, x_38); lean_dec(x_4); +lean_dec(x_3); return x_50; } } @@ -34507,7 +34509,6 @@ x_7 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabInductiveViews___lambda lean_closure_set(x_7, 0, x_5); lean_closure_set(x_7, 1, x_6); lean_closure_set(x_7, 2, x_1); -lean_inc(x_2); x_8 = l_Lean_Elab_Command_runTermElabM___rarg(x_7, x_2, x_3, x_4); if (lean_obj_tag(x_8) == 0) { @@ -34539,6 +34540,7 @@ lean_closure_set(x_14, 1, x_1); lean_closure_set(x_14, 2, x_6); x_15 = l_Lean_Elab_Command_runTermElabM___rarg(x_14, x_2, x_3, x_13); lean_dec(x_3); +lean_dec(x_2); return x_15; } else diff --git a/stage0/stdlib/Lean/Elab/InheritDoc.c b/stage0/stdlib/Lean/Elab/InheritDoc.c index 5540e0cd8be6..7f3fb34ef382 100644 --- a/stage0/stdlib/Lean/Elab/InheritDoc.c +++ b/stage0/stdlib/Lean/Elab/InheritDoc.c @@ -16,7 +16,6 @@ extern "C" { LEAN_EXPORT lean_object* l_Lean_Elab_resolveGlobalConstNoOverloadWithInfo___at_Lean_initFn____x40_Lean_Elab_InheritDoc___hyg_4____spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_initFn____x40_Lean_Elab_InheritDoc___hyg_4____spec__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Elab_InheritDoc___hyg_4____lambda__1___closed__2; static lean_object* l_Lean_initFn____x40_Lean_Elab_InheritDoc___hyg_4____closed__8; LEAN_EXPORT lean_object* l_Lean_Elab_addConstInfo___at_Lean_initFn____x40_Lean_Elab_InheritDoc___hyg_4____spec__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -106,6 +105,7 @@ static lean_object* l_Lean_initFn____x40_Lean_Elab_InheritDoc___hyg_4____closed_ static lean_object* l_Lean_addDocString___at_Lean_initFn____x40_Lean_Elab_InheritDoc___hyg_4____spec__18___closed__5; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); extern lean_object* l_Lean_warningAsError; +uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Elab_InheritDoc___hyg_4____lambda__4___closed__2; uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); @@ -1578,7 +1578,7 @@ else { lean_object* x_261; uint8_t x_262; x_261 = l_Lean_logAt___at_Lean_initFn____x40_Lean_Elab_InheritDoc___hyg_4____spec__17___closed__1; -x_262 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_258, x_261); +x_262 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_258, x_261); if (x_262 == 0) { x_8 = x_3; @@ -2574,7 +2574,7 @@ lean_object* x_5; lean_object* x_6; uint8_t x_7; x_5 = lean_ctor_get(x_2, 2); lean_inc(x_5); x_6 = l_Lean_logAt___at_Lean_initFn____x40_Lean_Elab_InheritDoc___hyg_4____spec__17___closed__1; -x_7 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_5, x_6); +x_7 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_5, x_6); lean_dec(x_5); if (x_7 == 0) { diff --git a/stage0/stdlib/Lean/Elab/LetRec.c b/stage0/stdlib/Lean/Elab/LetRec.c index 1f597db882c7..db121289974a 100644 --- a/stage0/stdlib/Lean/Elab/LetRec.c +++ b/stage0/stdlib/Lean/Elab/LetRec.c @@ -20,8 +20,10 @@ lean_object* l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(lean_object*, lean_obj lean_object* l___private_Init_Util_0__outOfBounds___rarg(lean_object*); LEAN_EXPORT uint8_t l_List_foldr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_registerLetRecsToLift___spec__1(lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___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_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___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_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___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_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__17___rarg(lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__6; +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__3; lean_object* l_Lean_MapDeclarationExtension_insert___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__12___lambda__1___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__13___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); @@ -40,7 +42,6 @@ extern lean_object* l_Lean_maxRecDepthErrorMessage; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__17(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_indentD(lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabLetRec(lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3___closed__3; lean_object* l_Lean_FileMap_toPosition(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__9___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_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__13___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -70,6 +71,7 @@ lean_object* l_Lean_Syntax_getTailPos_x3f(lean_object*, uint8_t); lean_object* l_Lean_Expr_fvarId_x21(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*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabLetRec___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_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__8; lean_object* l_Lean_addBuiltinDeclarationRanges(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); @@ -85,9 +87,9 @@ lean_object* l_instBEqProd___rarg(lean_object*, lean_object*, lean_object*, lean lean_object* lean_string_utf8_byte_size(lean_object*); static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_elabDeclAttrs___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__10(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_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3___closed__4; 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*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_registerLetRecsToLift___spec__2___closed__4; +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__13___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__6___closed__12; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -128,6 +130,7 @@ lean_object* l_Lean_Elab_getBetterRef(lean_object*, lean_object*); static lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__12___closed__6; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_elabLetRecDeclValues___spec__2(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_expandOptDocComment_x3f___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__1___closed__4; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___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_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__12___lambda__1___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_LetRec_0__Lean_Elab_Term_elabLetRecDeclValues(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_getDeclarationRange___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -140,6 +143,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetRec_declRange___closed_ lean_object* l_List_forM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getOptional_x3f(lean_object*); lean_object* l_Lean_Elab_Term_expandMatchAltsIntoMatch___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__7; static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetRec_declRange___closed__7; lean_object* lean_array_to_list(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__12(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -152,10 +156,10 @@ extern lean_object* l_Lean_Expr_instHashableExpr; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__18(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_LetRec_0__Lean_Elab_Term_registerLetRecsToLift___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_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__5; static lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__6___closed__9; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___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_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__12___lambda__2___closed__2; -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3___closed__1; lean_object* l_Lean_Syntax_getSepArgs(lean_object*); lean_object* l_Lean_instHashableLocalInstance___boxed(lean_object*); lean_object* l_instHashableProd___rarg___boxed(lean_object*, lean_object*, lean_object*); @@ -170,16 +174,16 @@ lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_registerLetRecsToLift___spec__1___boxed(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_Std_Range_forIn_loop___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_elabLetRecDeclValues___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*); +uint8_t l_Lean_Syntax_isIdent(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*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_elabLetRecDeclValues___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*); +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__9; LEAN_EXPORT lean_object* l_Lean_Elab_elabAttrs___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_expandOptDocComment_x3f___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__13___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3___closed__6; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__14(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_LetRec_0__Lean_Elab_Term_registerLetRecsToLift(lean_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_LetRec_0__Lean_Elab_Term_withAuxLocalDecls_loop___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_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3___closed__7; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_registerLetRecsToLift___spec__2(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_addDocString_x27___at_Lean_Elab_Term_expandDeclId___spec__8(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_elabLetRec_declRange___closed__3; @@ -190,6 +194,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRec LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__16___boxed(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___private_Lean_Elab_LetRec_0__Lean_Elab_Term_registerLetRecsToLift___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_Lean_Elab_liftMacroM___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__13___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__4; static lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__6___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_expandOptDocComment_x3f___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___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_elabLetRec___closed__2; @@ -203,7 +208,6 @@ uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__6___closed__1; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__18___boxed(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_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__3(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_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3___closed__2; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); lean_object* lean_environment_main_module(lean_object*); LEAN_EXPORT lean_object* l_Array_mapIdxM_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_registerLetRecsToLift___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -211,7 +215,6 @@ static lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean lean_object* l_Lean_Meta_getLocalInstances(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView(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_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___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_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3___closed__5; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__19(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_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*); @@ -241,10 +244,10 @@ size_t lean_usize_add(size_t, size_t); static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetRec___closed__6; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__15(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); -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3___closed__8; static lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__6___closed__13; extern lean_object* l_Lean_instInhabitedSyntax; lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4(lean_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_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__12___lambda__2(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_elabLetRec(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); @@ -254,6 +257,7 @@ LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_LetRec_0__ lean_object* l_Lean_getAttributeImpl(lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__1___closed__2; lean_object* lean_array_get_size(lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__10; uint8_t lean_nat_dec_le(lean_object*, lean_object*); extern lean_object* l_Lean_Elab_unsupportedSyntaxExceptionId; uint8_t lean_usize_dec_lt(size_t, size_t); @@ -269,6 +273,7 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabLetRec_declRange(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__12___lambda__1___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_LetRec_0__Lean_Elab_Term_withAuxLocalDecls_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*); +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__2; static lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__12___lambda__2___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetRec___closed__7; lean_object* l_Lean_Elab_Term_elabType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -276,6 +281,7 @@ static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_LetRec_0__ LEAN_EXPORT lean_object* l___private_Lean_Elab_LetRec_0__Lean_Elab_Term_registerLetRecsToLift___boxed(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_liftMacroM___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_getDeclarationRange___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBinderViews_loop___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_withDeclName___rarg(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_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___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: @@ -2886,7 +2892,329 @@ lean_ctor_set(x_17, 1, x_15); return x_17; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3___closed__1() { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3(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, 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_dec(x_6); +x_14 = l_Lean_Syntax_getId(x_1); +x_15 = l_Lean_Elab_Term_getDeclName_x3f(x_7, x_8, x_9, x_10, 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_inc(x_17); +lean_dec(x_15); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_78; +x_78 = lean_box(0); +x_18 = x_78; +goto block_77; +} +else +{ +lean_object* x_79; +x_79 = lean_ctor_get(x_16, 0); +lean_inc(x_79); +lean_dec(x_16); +x_18 = x_79; +goto block_77; +} +block_77: +{ +lean_object* x_19; lean_object* x_20; +lean_inc(x_14); +x_19 = l_Lean_Name_append(x_18, x_14); +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_19); +x_20 = l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Term_expandDeclId___spec__6(x_19, x_7, x_8, x_9, x_10, x_11, x_12, x_17); +if (lean_obj_tag(x_20) == 0) +{ +lean_object* x_21; uint8_t x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_20, 1); +lean_inc(x_21); +lean_dec(x_20); +x_22 = 2; +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_2); +lean_inc(x_19); +x_23 = l_Lean_Elab_Term_applyAttributesAt(x_19, x_2, x_22, x_7, x_8, x_9, x_10, x_11, x_12, x_21); +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; +x_24 = lean_ctor_get(x_23, 1); +lean_inc(x_24); +lean_dec(x_23); +lean_inc(x_19); +x_25 = l_Lean_addDocString_x27___at_Lean_Elab_Term_expandDeclId___spec__8(x_19, x_3, x_7, x_8, x_9, x_10, x_11, x_12, x_24); +x_26 = lean_ctor_get(x_25, 1); +lean_inc(x_26); +lean_dec(x_25); +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_1); +lean_inc(x_4); +lean_inc(x_19); +x_27 = l_Lean_Elab_addAuxDeclarationRanges___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__4(x_19, x_4, x_1, x_7, x_8, x_9, x_10, x_11, x_12, x_26); +x_28 = lean_ctor_get(x_27, 1); +lean_inc(x_28); +lean_dec(x_27); +x_29 = lean_unsigned_to_nat(1u); +x_30 = l_Lean_Syntax_getArg(x_4, x_29); +x_31 = l_Lean_Syntax_getArgs(x_30); +lean_dec(x_30); +x_32 = lean_unsigned_to_nat(2u); +x_33 = l_Lean_Syntax_getArg(x_4, x_32); +lean_inc(x_1); +x_34 = l_Lean_Elab_Term_expandOptType(x_1, x_33); +lean_dec(x_33); +x_35 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__1___boxed), 9, 1); +lean_closure_set(x_35, 0, x_34); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +x_36 = l_Lean_Elab_Term_elabBindersEx___rarg(x_31, x_35, x_7, x_8, x_9, x_10, x_11, x_12, x_28); +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; uint8_t x_42; lean_object* x_43; lean_object* x_44; +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_37, 0); +lean_inc(x_39); +x_40 = lean_ctor_get(x_37, 1); +lean_inc(x_40); +lean_dec(x_37); +lean_inc(x_39); +x_41 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_41, 0, x_39); +x_42 = 2; +x_43 = lean_box(0); +lean_inc(x_9); +x_44 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkFreshExprMVarImpl(x_41, x_42, x_43, x_9, x_10, x_11, x_12, x_38); +if (x_5 == 0) +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; +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 = lean_unsigned_to_nat(3u); +x_48 = l_Lean_Syntax_getArg(x_4, x_47); +x_49 = 1; +x_50 = lean_box(x_49); +x_51 = lean_alloc_closure((void*)(l_Lean_Elab_Term_expandMatchAltsIntoMatch___boxed), 5, 3); +lean_closure_set(x_51, 0, x_4); +lean_closure_set(x_51, 1, x_48); +lean_closure_set(x_51, 2, x_50); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +x_52 = l_Lean_Elab_liftMacroM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__9(x_51, x_7, x_8, x_9, x_10, x_11, x_12, x_46); +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, 0); +lean_inc(x_53); +x_54 = lean_ctor_get(x_52, 1); +lean_inc(x_54); +lean_dec(x_52); +x_55 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__2(x_1, x_2, x_14, x_19, x_40, x_39, x_45, x_53, x_7, x_8, x_9, x_10, x_11, x_12, x_54); +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_55; +} +else +{ +uint8_t x_56; +lean_dec(x_45); +lean_dec(x_40); +lean_dec(x_39); +lean_dec(x_19); +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_2); +lean_dec(x_1); +x_56 = !lean_is_exclusive(x_52); +if (x_56 == 0) +{ +return x_52; +} +else +{ +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_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_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +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_62 = lean_unsigned_to_nat(4u); +x_63 = l_Lean_Syntax_getArg(x_4, x_62); +lean_dec(x_4); +x_64 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__2(x_1, x_2, x_14, x_19, x_40, x_39, x_60, x_63, x_7, x_8, x_9, x_10, x_11, x_12, x_61); +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_64; +} +} +else +{ +uint8_t x_65; +lean_dec(x_19); +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_4); +lean_dec(x_2); +lean_dec(x_1); +x_65 = !lean_is_exclusive(x_36); +if (x_65 == 0) +{ +return x_36; +} +else +{ +lean_object* x_66; lean_object* x_67; lean_object* x_68; +x_66 = lean_ctor_get(x_36, 0); +x_67 = lean_ctor_get(x_36, 1); +lean_inc(x_67); +lean_inc(x_66); +lean_dec(x_36); +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 +{ +uint8_t x_69; +lean_dec(x_19); +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_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_69 = !lean_is_exclusive(x_23); +if (x_69 == 0) +{ +return x_23; +} +else +{ +lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_70 = lean_ctor_get(x_23, 0); +x_71 = lean_ctor_get(x_23, 1); +lean_inc(x_71); +lean_inc(x_70); +lean_dec(x_23); +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_73; +lean_dec(x_19); +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_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_73 = !lean_is_exclusive(x_20); +if (x_73 == 0) +{ +return x_20; +} +else +{ +lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_74 = lean_ctor_get(x_20, 0); +x_75 = lean_ctor_get(x_20, 1); +lean_inc(x_75); +lean_inc(x_74); +lean_dec(x_20); +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; +} +} +} +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__1() { _start: { lean_object* x_1; @@ -2894,19 +3222,19 @@ x_1 = lean_mk_string_from_bytes("letPatDecl", 10); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3___closed__2() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__2() { _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_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__12___lambda__1___closed__1; x_2 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__12___lambda__1___closed__2; x_3 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__12___lambda__1___closed__3; -x_4 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3___closed__1; +x_4 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3___closed__3() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__3() { _start: { lean_object* x_1; @@ -2914,19 +3242,36 @@ x_1 = lean_mk_string_from_bytes("letIdDecl", 9); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3___closed__4() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___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_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__12___lambda__1___closed__1; x_2 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__12___lambda__1___closed__2; x_3 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__12___lambda__1___closed__3; -x_4 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3___closed__3; +x_4 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__3; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3___closed__5() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("'let rec' expressions must be named", 35); +return x_1; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__5; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__7() { _start: { lean_object* x_1; @@ -2934,19 +3279,19 @@ x_1 = lean_mk_string_from_bytes("letEqnsDecl", 11); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3___closed__6() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__8() { _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_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__12___lambda__1___closed__1; x_2 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__12___lambda__1___closed__2; x_3 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__12___lambda__1___closed__3; -x_4 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3___closed__5; +x_4 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__7; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3___closed__7() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__9() { _start: { lean_object* x_1; @@ -2954,16 +3299,16 @@ x_1 = lean_mk_string_from_bytes("patterns are not allowed in 'let rec' expressio return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3___closed__8() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__10() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3___closed__7; +x_1 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__9; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___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_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___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) { _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; @@ -2973,24 +3318,24 @@ lean_dec(x_1); x_13 = lean_unsigned_to_nat(0u); x_14 = l_Lean_Syntax_getArg(x_12, x_13); lean_dec(x_12); -x_15 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3___closed__2; +x_15 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__2; 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; lean_object* x_19; -x_17 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3___closed__4; +x_17 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__4; lean_inc(x_14); x_18 = l_Lean_Syntax_isOfKind(x_14, x_17); if (x_18 == 0) { -lean_object* x_87; uint8_t x_88; -x_87 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3___closed__6; +lean_object* x_31; uint8_t x_32; +x_31 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__8; lean_inc(x_14); -x_88 = l_Lean_Syntax_isOfKind(x_14, x_87); -if (x_88 == 0) +x_32 = l_Lean_Syntax_isOfKind(x_14, x_31); +if (x_32 == 0) { -lean_object* x_89; +lean_object* x_33; lean_dec(x_14); lean_dec(x_9); lean_dec(x_8); @@ -3000,353 +3345,79 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_89 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___rarg(x_10); -return x_89; +x_33 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__7___rarg(x_10); +return x_33; } else { -lean_object* x_90; -x_90 = lean_box(0); -x_19 = x_90; -goto block_86; +lean_object* x_34; +x_34 = lean_box(0); +x_19 = x_34; +goto block_30; } } else { -lean_object* x_91; -x_91 = lean_box(0); -x_19 = x_91; -goto block_86; +lean_object* x_35; +x_35 = lean_box(0); +x_19 = x_35; +goto block_30; } -block_86: +block_30: { -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; uint8_t x_21; lean_dec(x_19); x_20 = l_Lean_Syntax_getArg(x_14, x_13); -x_21 = l_Lean_Syntax_getId(x_20); -x_22 = l_Lean_Elab_Term_getDeclName_x3f(x_4, x_5, x_6, x_7, x_8, x_9, x_10); -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); -if (lean_obj_tag(x_23) == 0) -{ -lean_object* x_84; -x_84 = lean_box(0); -x_25 = x_84; -goto block_83; -} -else -{ -lean_object* x_85; -x_85 = lean_ctor_get(x_23, 0); -lean_inc(x_85); -lean_dec(x_23); -x_25 = x_85; -goto block_83; -} -block_83: -{ -lean_object* x_26; lean_object* x_27; -lean_inc(x_21); -x_26 = l_Lean_Name_append(x_25, x_21); -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_26); -x_27 = l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Term_expandDeclId___spec__6(x_26, x_4, x_5, x_6, x_7, x_8, x_9, x_24); -if (lean_obj_tag(x_27) == 0) -{ -lean_object* x_28; uint8_t x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_27, 1); -lean_inc(x_28); -lean_dec(x_27); -x_29 = 2; -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_26); -x_30 = l_Lean_Elab_Term_applyAttributesAt(x_26, x_3, x_29, x_4, x_5, x_6, x_7, x_8, x_9, x_28); -if (lean_obj_tag(x_30) == 0) +x_21 = l_Lean_Syntax_isIdent(x_20); +if (x_21 == 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; -x_31 = lean_ctor_get(x_30, 1); -lean_inc(x_31); -lean_dec(x_30); -lean_inc(x_26); -x_32 = l_Lean_addDocString_x27___at_Lean_Elab_Term_expandDeclId___spec__8(x_26, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_31); -x_33 = lean_ctor_get(x_32, 1); -lean_inc(x_33); -lean_dec(x_32); -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_20); -lean_inc(x_14); -lean_inc(x_26); -x_34 = l_Lean_Elab_addAuxDeclarationRanges___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__4(x_26, x_14, x_20, x_4, x_5, x_6, x_7, x_8, x_9, x_33); -x_35 = lean_ctor_get(x_34, 1); -lean_inc(x_35); -lean_dec(x_34); -x_36 = lean_unsigned_to_nat(1u); -x_37 = l_Lean_Syntax_getArg(x_14, x_36); -x_38 = l_Lean_Syntax_getArgs(x_37); -lean_dec(x_37); -x_39 = l_Lean_Syntax_getArg(x_14, x_11); -lean_inc(x_20); -x_40 = l_Lean_Elab_Term_expandOptType(x_20, x_39); -lean_dec(x_39); -x_41 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__1___boxed), 9, 1); -lean_closure_set(x_41, 0, x_40); -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_42 = l_Lean_Elab_Term_elabBindersEx___rarg(x_38, x_41, x_4, x_5, x_6, x_7, x_8, x_9, x_35); -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; uint8_t x_48; lean_object* x_49; lean_object* x_50; -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_ctor_get(x_43, 0); -lean_inc(x_45); -x_46 = lean_ctor_get(x_43, 1); -lean_inc(x_46); -lean_dec(x_43); -lean_inc(x_45); -x_47 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_47, 0, x_45); -x_48 = 2; -x_49 = lean_box(0); -lean_inc(x_6); -x_50 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkFreshExprMVarImpl(x_47, x_48, x_49, x_6, x_7, x_8, x_9, x_44); -if (x_18 == 0) -{ -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; lean_object* x_58; -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_unsigned_to_nat(3u); -x_54 = l_Lean_Syntax_getArg(x_14, x_53); -x_55 = 1; -x_56 = lean_box(x_55); -x_57 = lean_alloc_closure((void*)(l_Lean_Elab_Term_expandMatchAltsIntoMatch___boxed), 5, 3); -lean_closure_set(x_57, 0, x_14); -lean_closure_set(x_57, 1, x_54); -lean_closure_set(x_57, 2, x_56); -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_58 = l_Lean_Elab_liftMacroM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__9(x_57, x_4, x_5, x_6, x_7, x_8, x_9, x_52); -if (lean_obj_tag(x_58) == 0) -{ -lean_object* x_59; lean_object* x_60; lean_object* x_61; -x_59 = lean_ctor_get(x_58, 0); -lean_inc(x_59); -x_60 = lean_ctor_get(x_58, 1); -lean_inc(x_60); -lean_dec(x_58); -x_61 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__2(x_20, x_3, x_21, x_26, x_46, x_45, x_51, x_59, x_4, x_5, x_6, x_7, x_8, x_9, x_60); -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_61; -} -else -{ -uint8_t x_62; -lean_dec(x_51); -lean_dec(x_46); -lean_dec(x_45); -lean_dec(x_26); -lean_dec(x_21); -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); -x_62 = !lean_is_exclusive(x_58); -if (x_62 == 0) -{ -return x_58; -} -else -{ -lean_object* x_63; lean_object* x_64; lean_object* x_65; -x_63 = lean_ctor_get(x_58, 0); -x_64 = lean_ctor_get(x_58, 1); -lean_inc(x_64); -lean_inc(x_63); -lean_dec(x_58); -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_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; -x_66 = lean_ctor_get(x_50, 0); -lean_inc(x_66); -x_67 = lean_ctor_get(x_50, 1); -lean_inc(x_67); -lean_dec(x_50); -x_68 = lean_unsigned_to_nat(4u); -x_69 = l_Lean_Syntax_getArg(x_14, x_68); +lean_object* x_22; lean_object* x_23; uint8_t x_24; lean_dec(x_14); -x_70 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__2(x_20, x_3, x_21, x_26, x_46, x_45, x_66, x_69, x_4, x_5, x_6, x_7, x_8, x_9, x_67); -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_70; -} -} -else -{ -uint8_t x_71; -lean_dec(x_26); -lean_dec(x_21); -lean_dec(x_20); -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_71 = !lean_is_exclusive(x_42); -if (x_71 == 0) -{ -return x_42; -} -else -{ -lean_object* x_72; lean_object* x_73; lean_object* x_74; -x_72 = lean_ctor_get(x_42, 0); -x_73 = lean_ctor_get(x_42, 1); -lean_inc(x_73); -lean_inc(x_72); -lean_dec(x_42); -x_74 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_74, 0, x_72); -lean_ctor_set(x_74, 1, x_73); -return x_74; -} -} -} -else -{ -uint8_t x_75; -lean_dec(x_26); -lean_dec(x_21); -lean_dec(x_20); -lean_dec(x_14); +lean_dec(x_2); +x_22 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__6; +x_23 = l_Lean_throwErrorAt___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBinderViews_loop___spec__1(x_20, x_22, 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); -x_75 = !lean_is_exclusive(x_30); -if (x_75 == 0) -{ -return x_30; -} -else +lean_dec(x_20); +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) { -lean_object* x_76; lean_object* x_77; lean_object* x_78; -x_76 = lean_ctor_get(x_30, 0); -x_77 = lean_ctor_get(x_30, 1); -lean_inc(x_77); -lean_inc(x_76); -lean_dec(x_30); -x_78 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_78, 0, x_76); -lean_ctor_set(x_78, 1, x_77); -return x_78; -} -} +return x_23; } else { -uint8_t x_79; -lean_dec(x_26); -lean_dec(x_21); -lean_dec(x_20); -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); -lean_dec(x_2); -x_79 = !lean_is_exclusive(x_27); -if (x_79 == 0) -{ +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(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_80; lean_object* x_81; lean_object* x_82; -x_80 = lean_ctor_get(x_27, 0); -x_81 = lean_ctor_get(x_27, 1); -lean_inc(x_81); -lean_inc(x_80); -lean_dec(x_27); -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_28; lean_object* x_29; +x_28 = lean_box(0); +x_29 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3(x_20, x_3, x_2, x_14, x_18, x_28, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_29; } } } else { -lean_object* x_92; lean_object* x_93; +lean_object* x_36; lean_object* x_37; lean_dec(x_3); lean_dec(x_2); -x_92 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3___closed__8; -x_93 = l_Lean_throwErrorAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__8(x_14, x_92, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_93; +x_36 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__10; +x_37 = l_Lean_throwErrorAt___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__8(x_14, x_36, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_37; } } } @@ -3420,7 +3491,7 @@ lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_27 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3(x_13, x_18, x_25, x_4, x_5, x_6, x_7, x_8, x_9, x_26); +x_27 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4(x_13, x_18, x_25, x_4, x_5, x_6, x_7, x_8, x_9, x_26); if (lean_obj_tag(x_27) == 0) { lean_object* x_28; lean_object* x_29; size_t x_30; size_t x_31; lean_object* x_32; @@ -3510,7 +3581,7 @@ lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_43 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3(x_13, x_18, x_42, x_4, x_5, x_6, x_7, x_8, x_9, x_19); +x_43 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4(x_13, x_18, x_42, x_4, x_5, x_6, x_7, x_8, x_9, x_19); if (lean_obj_tag(x_43) == 0) { lean_object* x_44; lean_object* x_45; size_t x_46; size_t x_47; lean_object* x_48; @@ -3905,6 +3976,16 @@ lean_dec(x_9); return x_16; } } +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___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) { +_start: +{ +uint8_t x_14; lean_object* x_15; +x_14 = lean_unbox(x_5); +lean_dec(x_5); +x_15 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3(x_1, x_2, x_3, x_4, x_14, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_15; +} +} LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, 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: { @@ -5799,7 +5880,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetRec_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(108u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -5811,7 +5892,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetRec_declRange___c _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(117u); x_2 = lean_unsigned_to_nat(52u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -5839,7 +5920,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetRec_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(108u); x_2 = lean_unsigned_to_nat(34u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -5851,7 +5932,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetRec_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(108u); x_2 = lean_unsigned_to_nat(44u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -6018,22 +6099,26 @@ l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRe lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__1___closed__2); l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__1___closed__3 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__1___closed__3(); lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__1___closed__3); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3___closed__1 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3___closed__1(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3___closed__1); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3___closed__2 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3___closed__2(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3___closed__2); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3___closed__3 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3___closed__3(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3___closed__3); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3___closed__4 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3___closed__4(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3___closed__4); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3___closed__5 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3___closed__5(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3___closed__5); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3___closed__6 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3___closed__6(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3___closed__6); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3___closed__7 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3___closed__7(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3___closed__7); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3___closed__8 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3___closed__8(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__3___closed__8); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__1 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__1(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__1); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__2 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__2(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__2); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__3 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__3(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__3); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__4 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__4(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__4); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__5 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__5(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__5); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__6 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__6(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__6); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__7 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__7(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__7); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__8 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__8(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__8); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__9 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__9(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__9); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__10 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__10(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_mkLetRecDeclView___spec__21___lambda__4___closed__10); l_Std_Range_forIn_loop___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_elabLetRecDeclValues___spec__1___closed__1 = _init_l_Std_Range_forIn_loop___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_elabLetRecDeclValues___spec__1___closed__1(); lean_mark_persistent(l_Std_Range_forIn_loop___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_elabLetRecDeclValues___spec__1___closed__1); l_Array_forInUnsafe_loop___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_registerLetRecsToLift___spec__2___closed__1 = _init_l_Array_forInUnsafe_loop___at___private_Lean_Elab_LetRec_0__Lean_Elab_Term_registerLetRecsToLift___spec__2___closed__1(); diff --git a/stage0/stdlib/Lean/Elab/Level.c b/stage0/stdlib/Lean/Elab/Level.c index 6fa4fcab42e2..b86ec333867d 100644 --- a/stage0/stdlib/Lean/Elab/Level.c +++ b/stage0/stdlib/Lean/Elab/Level.c @@ -14,7 +14,6 @@ extern "C" { #endif static lean_object* l_Lean_Elab_Level_elabLevel___closed__18; -uint8_t l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Elab_Level_elabLevel___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Elab_Level_elabLevel___spec__7(lean_object*, size_t, size_t, lean_object*, 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*); @@ -39,7 +38,6 @@ static lean_object* l___private_Lean_Elab_Level_0__Lean_Elab_Level_checkUniverse uint8_t lean_usize_dec_eq(size_t, size_t); lean_object* l_Lean_MetavarContext_addLevelMVarDecl(lean_object*, lean_object*); lean_object* l_Lean_Syntax_getArgs(lean_object*); -lean_object* l_Lean_Option_get___at_Std_Format_pretty_x27___spec__1(lean_object*, lean_object*); lean_object* l_Lean_replaceRef(lean_object*, lean_object*); static lean_object* l_Lean_Elab_throwIllFormedSyntax___at_Lean_Elab_Level_elabLevel___spec__2___closed__1; LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Elab_Level_instMonadOptionsLevelElabM___spec__2(lean_object*, lean_object*); @@ -67,6 +65,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Level_instMonadNameGeneratorLevelElabM___la static lean_object* l_Lean_Elab_Level_elabLevel___closed__17; LEAN_EXPORT lean_object* l_Lean_Elab_Level_initFn____x40_Lean_Elab_Level___hyg_440_(lean_object*); static lean_object* l_Lean_Elab_Level_initFn____x40_Lean_Elab_Level___hyg_440____closed__3; +lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Level_instMonadOptionsLevelElabM___lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Level_instMonadOptionsLevelElabM; lean_object* l_Lean_Syntax_getKind(lean_object*); @@ -89,6 +88,7 @@ uint8_t lean_name_eq(lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Level_instMonadRefLevelElabM___closed__1; static lean_object* l_Lean_Elab_Level_initFn____x40_Lean_Elab_Level___hyg_440____closed__5; +uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Level_instMonadNameGeneratorLevelElabM___lambda__3___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Level_instMonadNameGeneratorLevelElabM___closed__3; lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); @@ -909,7 +909,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Level_0__Lean_Elab_Level_checkUni { lean_object* x_5; lean_object* x_6; uint8_t x_7; x_5 = l___private_Lean_Elab_Level_0__Lean_Elab_Level_checkUniverseOffset___rarg___lambda__1___closed__1; -x_6 = l_Lean_Option_get___at_Std_Format_pretty_x27___spec__1(x_4, x_5); +x_6 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_4, x_5); x_7 = lean_nat_dec_le(x_1, x_6); if (x_7 == 0) { @@ -1042,7 +1042,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Level_0__Lean_Elab_Level_checkUni lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; x_4 = lean_ctor_get(x_2, 0); x_5 = l___private_Lean_Elab_Level_0__Lean_Elab_Level_checkUniverseOffset___rarg___lambda__1___closed__1; -x_6 = l_Lean_Option_get___at_Std_Format_pretty_x27___spec__1(x_4, x_5); +x_6 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_4, x_5); x_7 = lean_nat_dec_le(x_1, x_6); lean_dec(x_1); if (x_7 == 0) @@ -1661,7 +1661,7 @@ else { lean_object* x_66; uint8_t x_67; uint8_t x_68; x_66 = l_Lean_Elab_Level_elabLevel___closed__22; -x_67 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_8, x_66); +x_67 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_8, x_66); lean_dec(x_8); lean_inc(x_51); x_68 = l_Lean_Elab_isValidAutoBoundLevelName(x_51, x_67); @@ -2483,7 +2483,7 @@ else { lean_object* x_269; uint8_t x_270; uint8_t x_271; x_269 = l_Lean_Elab_Level_elabLevel___closed__22; -x_270 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_212, x_269); +x_270 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_212, x_269); lean_dec(x_212); lean_inc(x_254); x_271 = l_Lean_Elab_isValidAutoBoundLevelName(x_254, x_270); diff --git a/stage0/stdlib/Lean/Elab/MacroArgUtil.c b/stage0/stdlib/Lean/Elab/MacroArgUtil.c index c4894545da99..48890382bf48 100644 --- a/stage0/stdlib/Lean/Elab/MacroArgUtil.c +++ b/stage0/stdlib/Lean/Elab/MacroArgUtil.c @@ -73,6 +73,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat(lean_ lean_object* l_Lean_SourceInfo_fromRef(lean_object*, uint8_t); static lean_object* l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__20; static lean_object* l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__2___closed__4; +lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabCommand___spec__13(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Command_getRef(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__2___closed__1; static lean_object* l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___closed__14; @@ -123,7 +124,6 @@ static lean_object* l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_expandMacroArg_mkAntiquotNode___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_Command_expandMacroArg___closed__4; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); -lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabCommand___spec__14(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_expandMacroArg_mkSyntaxAndPat___lambda__2___boxed(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_Command_expandMacroArg___spec__1___rarg___closed__2; @@ -1374,7 +1374,6 @@ lean_dec(x_4); lean_inc(x_1); x_8 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabParserName_x3f), 8, 1); lean_closure_set(x_8, 0, x_1); -lean_inc(x_5); x_9 = l_Lean_Elab_Command_liftTermElabM___rarg(x_8, x_5, x_6, x_7); if (lean_obj_tag(x_9) == 0) { @@ -7410,7 +7409,7 @@ x_1053 = lean_alloc_closure((void*)(l_Lean_Elab_Command_strLitToPattern___boxed) lean_closure_set(x_1053, 0, x_1005); lean_inc(x_5); lean_inc(x_4); -x_1054 = l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabCommand___spec__14(x_1053, x_4, x_5, x_6); +x_1054 = l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabCommand___spec__13(x_1053, x_4, x_5, x_6); if (lean_obj_tag(x_1054) == 0) { 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; @@ -7642,7 +7641,7 @@ x_1121 = lean_alloc_closure((void*)(l_Lean_Elab_Command_strLitToPattern___boxed) lean_closure_set(x_1121, 0, x_1073); lean_inc(x_5); lean_inc(x_4); -x_1122 = l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabCommand___spec__14(x_1121, x_4, x_5, x_6); +x_1122 = l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabCommand___spec__13(x_1121, x_4, x_5, x_6); if (lean_obj_tag(x_1122) == 0) { 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; @@ -7906,7 +7905,7 @@ lean_closure_set(x_6, 0, x_1); lean_closure_set(x_6, 1, x_5); lean_inc(x_3); lean_inc(x_2); -x_7 = l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabCommand___spec__14(x_6, x_2, x_3, x_4); +x_7 = l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabCommand___spec__13(x_6, x_2, x_3, x_4); if (lean_obj_tag(x_7) == 0) { lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; diff --git a/stage0/stdlib/Lean/Elab/MacroRules.c b/stage0/stdlib/Lean/Elab/MacroRules.c index c2f1283c695d..bcbaa9e8398f 100644 --- a/stage0/stdlib/Lean/Elab/MacroRules.c +++ b/stage0/stdlib/Lean/Elab/MacroRules.c @@ -80,6 +80,7 @@ size_t lean_usize_of_nat(lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__15; static lean_object* l___regBuiltin_Lean_Elab_Command_elabMacroRules___closed__2; 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_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__17___rarg(lean_object*); static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__36; extern lean_object* l_Lean_Elab_Command_commandElabAttribute; 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*); @@ -103,7 +104,6 @@ lean_object* l_Lean_Elab_getBetterRef(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabMacroRules___lambda__3___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_elabMacroRulesAux___spec__5___closed__1; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMacroRulesAux___spec__5___lambda__2___closed__6; -lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__18___rarg(lean_object*); static lean_object* l_Lean_Elab_Command_elabMacroRules___lambda__2___closed__4; static lean_object* l_Lean_Elab_Command_elabMacroRulesAux___closed__18; static lean_object* l_Lean_Elab_Command_elabMacroRules___lambda__2___closed__1; @@ -2488,7 +2488,7 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); -x_13 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__18___rarg(x_8); +x_13 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__17___rarg(x_8); return x_13; } else @@ -2518,7 +2518,7 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_2); -x_22 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__18___rarg(x_8); +x_22 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__17___rarg(x_8); return x_22; } else @@ -2541,7 +2541,7 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_2); -x_27 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__18___rarg(x_8); +x_27 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__17___rarg(x_8); return x_27; } else @@ -3977,7 +3977,7 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); -x_475 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__18___rarg(x_8); +x_475 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__17___rarg(x_8); return x_475; } else @@ -4175,7 +4175,7 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); -x_12 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__18___rarg(x_7); +x_12 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__17___rarg(x_7); return x_12; } else @@ -4195,7 +4195,7 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); -x_17 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__18___rarg(x_7); +x_17 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__17___rarg(x_7); return x_17; } else @@ -4269,7 +4269,7 @@ lean_object* x_7; lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_7 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__18___rarg(x_4); +x_7 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__17___rarg(x_4); return x_7; } else @@ -4291,7 +4291,7 @@ lean_dec(x_9); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_13 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__18___rarg(x_4); +x_13 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__17___rarg(x_4); return x_13; } else @@ -4309,7 +4309,7 @@ lean_dec(x_14); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_17 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__18___rarg(x_4); +x_17 = l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__17___rarg(x_4); return x_17; } else diff --git a/stage0/stdlib/Lean/Elab/Match.c b/stage0/stdlib/Lean/Elab/Match.c index 2772227c7bf9..278e18b03dd2 100644 --- a/stage0/stdlib/Lean/Elab/Match.c +++ b/stage0/stdlib/Lean/Elab/Match.c @@ -14,7 +14,7 @@ extern "C" { #endif 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*); -uint8_t l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18388_(lean_object*); static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrs___lambda__1___closed__1; 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_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews_loop___spec__6(lean_object*, size_t, size_t, lean_object*); @@ -44,7 +44,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_Match_0 LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_isMatchUnit_x3f(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_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_Array_forInUnsafe_loop___at_Lean_Elab_Term_ToDepElimPattern_main___spec__5___closed__6; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18386____closed__4; 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___private_Lean_Elab_Match_0__Lean_Elab_Term_withEqs_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_EXPORT lean_object* l_Lean_Elab_throwAbortTerm___at_Lean_Elab_Term_isAtomicDiscr___spec__1___rarg(lean_object*); @@ -62,6 +61,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at___private_Lean_Elab_Match_ 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*); static lean_object* l_Lean_Elab_Term_elabMatch_elabMatchDefault___closed__1; static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___lambda__2___closed__1; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18388____closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrs___spec__7___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* l_Lean_Meta_withLocalInstancesImp___rarg(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_withToClear___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*); @@ -75,6 +75,7 @@ static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Match_0__L lean_object* l_Lean_mkSimpleThunk(lean_object*); lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Offset_0__Lean_Meta_withInstantiatedMVars___spec__1(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___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_initFn____x40_Lean_Elab_Match___hyg_22721_(lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_Term_ToDepElimPattern_main___spec__4(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews_loop___lambda__3___boxed(lean_object**); static lean_object* l_Lean_Elab_Term_elabNoMatch___closed__2; @@ -130,7 +131,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabMatch_elabMatchDefault___lambda__1 static lean_object* l___regBuiltin_Lean_Elab_Term_precheckMatch___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*); uint8_t l_Lean_Expr_isApp(lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__5; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews_collectDeps___spec__4(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* l_Lean_Syntax_getId(lean_object*); 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*); @@ -178,7 +178,6 @@ lean_object* l_Lean_Elab_Term_getPatternsVars(lean_object*, lean_object*, lean_o LEAN_EXPORT lean_object* l_List_forIn_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___spec__9___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_Match_0__Lean_Elab_Term_findDiscrRefinementPath_goType___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_Meta_transform_visit___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrs___spec__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_initFn____x40_Lean_Elab_Match___hyg_18386_(lean_object*); 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___private_Lean_Elab_Match_0__Lean_Elab_Term_getIndexToInclude_x3f_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_ToDepElimPattern_main_pack(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -261,7 +260,6 @@ uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandSimpleMatch___closed__6; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Match_0__Lean_Elab_Term_eraseIndices___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_Term_reportMatcherResultErrors(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_Lean_Meta_transform_visit_visitForall___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrs___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*); @@ -269,6 +267,7 @@ LEAN_EXPORT uint8_t l_Lean_AssocList_contains___at___private_Lean_Elab_Match_0__ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_precheckMatch___spec__3(size_t, size_t, lean_object*); uint8_t l___private_Lean_Util_Trace_0__Lean_checkTraceOption(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabNoMatch(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_18388____closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit_visitPost___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrs___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* l_Lean_mkApp4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_RBNode_findCore___at_Lean_Meta_mkGeneralizationForbiddenSet_visit___spec__1(lean_object*, lean_object*); @@ -298,6 +297,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Elab_Term_ToDepElim static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_isMatchUnit_x3f___closed__7; lean_object* l_Lean_Option_register___at_Lean_Elab_initFn____x40_Lean_Elab_AutoBound___hyg_7____spec__1(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__15; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltView___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_EXPORT lean_object* l_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___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___private_Lean_Elab_Match_0__Lean_Elab_Term_isPatternVar___closed__2; @@ -319,6 +319,7 @@ LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_ LEAN_EXPORT lean_object* l_Lean_mkHashMap___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__1___boxed(lean_object*); lean_object* l_Lean_Meta_Match_Pattern_toExpr_visit(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_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrsWitMatchType___spec__1___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_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__7; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_topSort___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___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchOptMotive___boxed(lean_object*); extern lean_object* l_instInhabitedPUnit; @@ -333,11 +334,12 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElim LEAN_EXPORT lean_object* l_Lean_Elab_Term_ToDepElimPattern_normalize_processVar(lean_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_Match_0__Lean_Elab_Term_findDiscrRefinementPath_goIndex___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_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___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_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__1; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__12; static lean_object* l___regBuiltin_Lean_Elab_Term_elabInaccessible___closed__6; static lean_object* l___regBuiltin_Lean_Elab_Term_elabNoMatch_declRange___closed__2; static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandSimpleMatch___closed__1; size_t lean_ptr_addr(lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__10; lean_object* l_Lean_Meta_isTypeCorrect(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at___private_Lean_Elab_Match_0__Lean_Elab_Term_withToClear___spec__3___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandNonAtomicDiscrs_x3f_loop___lambda__1___closed__4; @@ -360,13 +362,13 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_ToDepElimPat 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*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_Match_0__Lean_Elab_Term_withToClear___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrs___lambda__3___closed__2; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__3; 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___private_Lean_Elab_Match_0__Lean_Elab_Term_generalize___lambda__1___closed__1; lean_object* l_Lean_log___at_Lean_Elab_Term_exceptionToSorry___spec__2(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_checked___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__12(lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrsWitMatchType___spec__1___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_elabInaccessible_declRange___closed__4; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18386____closed__2; static lean_object* l_Lean_Meta_transform___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrs___spec__1___closed__1; lean_object* l_Array_reverse___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit_go___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__10___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -383,6 +385,7 @@ LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_v lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLetDeclImp___rarg(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___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews_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_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_precheckMatch___spec__5(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_initFn____x40_Lean_Elab_Match___hyg_22721____closed__13; lean_object* l_liftExcept___at_Lean_Elab_liftMacroM___spec__1(lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews_collectDeps___closed__2; @@ -517,7 +520,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_ToDepElimPattern_main_unpack_go___rarg lean_object* lean_array_to_list(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabNoMatch___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___spec__5___rarg(lean_object*); -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Elab_Term_termElabAttribute; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_ToDepElimPattern_main___spec__5___closed__3; 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*); @@ -529,6 +531,7 @@ LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Match_0 lean_object* l_Lean_Elab_Term_saveState___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltView___spec__4(lean_object*, lean_object*); lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_withLocalInstances___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltView___spec__3___rarg___boxed(lean_object*, 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_addAutoBoundImplicits_go___spec__50(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_elabMatchAltView___lambda__4___closed__2; @@ -562,6 +565,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_generaliz 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_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabMatch___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews(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_18388____closed__1; lean_object* l_Lean_Meta_whnfD(lean_object*, 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, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltView___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -573,7 +577,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_ToDepElimPattern_main___rarg___lambda_ LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews_collectDeps___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___private_Lean_Elab_Match_0__Lean_Elab_Term_findDiscrRefinementPath_goType___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_Lean_FVarId_getType(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_22719____closed__3; lean_object* l_Lean_Elab_Term_tryPostponeIfMVar(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_waitExpectedType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkFreshBinderName___at___private_Lean_Elab_Match_0__Lean_Elab_Term_mkUserNameFor___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -583,7 +586,6 @@ lean_object* l_List_lengthTRAux___rarg(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_object* l_Lean_Syntax_getSepArgs(lean_object*); lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltView___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_Term_commitIfNoErrors_x3f___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_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*); @@ -604,16 +606,17 @@ lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_isMatchUnit_x3f___closed__3; static lean_object* l_Lean_Elab_Term_reportMatcherResultErrors___closed__2; LEAN_EXPORT lean_object* l_Lean_Expr_ReplaceImpl_replaceUnsafeM_visit___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_applyRefMap___spec__3(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__9; static lean_object* l___regBuiltin_Lean_Elab_Term_elabNoMatch_declRange___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabPatterns(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*); lean_object* l_Lean_mkPatternWithRef(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__13; LEAN_EXPORT lean_object* l_Lean_Elab_Term_ToDepElimPattern_main_unpack_go___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___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_Lean_Elab_Term_ToDepElimPattern_State_patternVars___default; lean_object* l_Lean_Syntax_node2(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(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_synthesizeSyntheticMVarsUsingDefault(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_exprDependsOn___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_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); @@ -654,7 +657,6 @@ static lean_object* l_Lean_Elab_Term_elabMatch_elabMatchDefault___closed__2; LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltView___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_List_elem___at_Lean_Occurrences_contains___spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___lambda__5___closed__1; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__8; LEAN_EXPORT lean_object* l_Lean_Meta_withExistingLocalDecls___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___spec__8___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_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go(lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabNoMatch_declRange___closed__1; @@ -664,15 +666,12 @@ static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltVi lean_object* l_Lean_KVMap_setBool(lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_withEqs(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrs___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_initFn____x40_Lean_Elab_Match___hyg_22719____closed__7; static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_withToClear___rarg___closed__2; 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*); static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandSimpleMatch___closed__13; size_t lean_usize_mod(size_t, size_t); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__10; LEAN_EXPORT lean_object* l_List_forIn_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___spec__10___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_precheckMatch___lambda__2(lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18386____closed__1; static lean_object* l_List_forIn_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___spec__9___lambda__1___closed__4; size_t lean_data_hashmap_mk_idx(lean_object*, uint64_t); lean_object* l_Lean_Meta_instInhabitedMetaM___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -724,6 +723,7 @@ static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux__ static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabPatterns___spec__2___closed__2; LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyM___at___private_Lean_Elab_Match_0__Lean_Elab_Term_withToClear___spec__5(lean_object*, lean_object*); lean_object* l_Lean_Syntax_setArg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__11; LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_withToClear___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___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*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews_loop___lambda__4___boxed(lean_object**); @@ -734,6 +734,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_ToDepElimPattern_normalize___lambda__1 lean_object* l_Lean_Elab_Term_addTermInfo_x27(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* l_Lean_HashSetImp_insert___at_Lean_CollectFVars_visit___spec__3(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Elab_Match_0__Lean_Elab_Term_withToClear___spec__16(lean_object*, lean_object*, size_t, size_t); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__4; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_ToDepElimPattern_main_unpack_go___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_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts___spec__2___closed__2; static lean_object* l_Lean_Elab_Term_ToDepElimPattern_normalize_addVar___closed__2; @@ -742,6 +743,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_Match_0 LEAN_EXPORT lean_object* l_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go(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_ToDepElimPattern_main___rarg___lambda__2(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___private_Lean_Elab_Match_0__Lean_Elab_Term_expandSimpleMatch___closed__14; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__14; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Match_0__Lean_Elab_Term_generalize___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_withDepElimPatterns(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit_visitLambda___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrs___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_object*, lean_object*, lean_object*, lean_object*); @@ -793,10 +795,12 @@ LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at___private_Lean_Elab_Matc LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrs___spec__14(lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); lean_object* lean_environment_main_module(lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__6; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Match_0__Lean_Elab_Term_withToClear___spec__23(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at___private_Lean_Elab_Match_0__Lean_Elab_Term_withToClear___spec__10(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltView___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_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___spec__1___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(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___rarg___boxed(lean_object*, lean_object*); lean_object* l_Lean_Meta_setMVarUserNamesAt(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_precheckMatch(lean_object*); @@ -829,7 +833,6 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabNoMatch_declRange___closed LEAN_EXPORT lean_object* l_Lean_RBTree_toList___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews_collectDeps___spec__2(lean_object*); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Match_0__Lean_Elab_Term_generalize___spec__6___lambda__1(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_node7(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_22719____closed__9; 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*); static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchGeneralizing_x3f___closed__4; LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Match_0__Lean_Elab_Term_generalize___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*); @@ -857,12 +860,14 @@ lean_object* l_Lean_Meta_forallBoundedTelescope___at_Lean_Elab_Term_addAutoBound lean_object* l_Lean_localDeclDependsOn___at_Lean_Elab_Term_addAutoBoundImplicits_go___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_Meta_transform_visit_visitLet___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrs___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_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_initFn____x40_Lean_Elab_Match___hyg_18388____closed__2; 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*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews_collectDeps___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_Array_mapMUnsafe_map___at_Lean_Elab_Term_precheckMatch___spec__4(size_t, size_t, 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_object* l_Lean_Elab_Term_commitIfDidNotPostpone___rarg(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_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18388____closed__5; static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandSimpleMatch___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandMacrosInPatterns(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_isNone(lean_object*); @@ -901,8 +906,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Elab_Matc static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_throwInvalidPattern___rarg___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_elabNoMatch___spec__1___rarg(lean_object*); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Match_0__Lean_Elab_Term_generalize___spec__6(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_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__12; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__11; static lean_object* l___regBuiltin_Lean_Elab_Term_elabNoMatch_declRange___closed__7; 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*); 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*); @@ -939,7 +942,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_isAtomicDiscr___boxed(lean_object*, le LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltView___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_List_mapM_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___spec__11(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_GeneralizeResult_toClear___default; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18386____closed__3; LEAN_EXPORT lean_object* l_Array_sequenceMap___at_Lean_Elab_Term_precheckMatch___spec__1___boxed(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__9; static lean_object* l_Lean_Elab_Term_ToDepElimPattern_main___rarg___closed__4; @@ -950,7 +952,6 @@ size_t lean_usize_add(size_t, size_t); lean_object* l_Lean_Meta_erasePatternRefAnnotations___lambda__1___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_1116____at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___spec__13___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_Match_0__Lean_Elab_Term_withToClear___spec__21(lean_object*, lean_object*, size_t, size_t, 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_22719____closed__14; static lean_object* l___regBuiltin_Lean_Elab_Term_elabNoMatch___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabInaccessible_declRange(lean_object*); lean_object* l_Lean_throwError___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__50___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -998,7 +999,6 @@ lean_object* l_List_redLength___rarg(lean_object*); static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchGeneralizing_x3f___closed__11; LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_checked___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__12___rarg___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltView___lambda__2___closed__2; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__4; static lean_object* l___regBuiltin_Lean_Elab_Term_elabInaccessible_declRange___closed__1; lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Match_0__Lean_Elab_Term_generalize___spec__8(size_t, lean_object*, lean_object*, size_t, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1006,6 +1006,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_ToDepElimPattern_TopSort_State_visited LEAN_EXPORT lean_object* l_Lean_Elab_Term_ToDepElimPattern_main(lean_object*); lean_object* l_Lean_Meta_sortFVarIds(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_setTag(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_22721____closed__8; lean_object* l_Lean_Meta_mkFreshTypeMVar(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_findDiscrRefinementPath_goIndex___spec__1___closed__1; lean_object* lean_string_append(lean_object*, lean_object*); @@ -1031,7 +1032,6 @@ static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Match_0__L LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews_containsFVar___spec__1(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___spec__6___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_ToDepElimPattern_main___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_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrs___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___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_toPattern___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Match_0__Lean_Elab_Term_generalize___spec__3___boxed(lean_object*, lean_object*, lean_object*); @@ -1078,7 +1078,6 @@ uint8_t l_Lean_Expr_isFVar(lean_object*); 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_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_findDiscrRefinementPath_checkCompatibleApps___boxed(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(lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18386____closed__5; 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_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrsWitMatchType___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_List_mapM_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1112,6 +1111,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElim LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_applyRefMap(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__9___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandSimpleMatch___closed__10; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__5; LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at_Lean_Elab_Term_precheckMatch___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Elab_Match_0__Lean_Elab_Term_withToClear___spec__7(lean_object*, lean_object*, size_t, size_t); static lean_object* l_Lean_Elab_Term_mkFreshBinderName___at___private_Lean_Elab_Match_0__Lean_Elab_Term_mkUserNameFor___spec__1___rarg___closed__2; @@ -22509,7 +22509,7 @@ lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean x_11 = lean_array_uget(x_3, x_2); x_12 = lean_unsigned_to_nat(0u); x_13 = lean_array_uset(x_3, x_2, x_12); -x_14 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_11, x_4, x_5, x_6, x_7, x_8); +x_14 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_11, x_4, x_5, 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); @@ -38986,7 +38986,7 @@ lean_dec(x_2); return x_9; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18386____closed__1() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18388____closed__1() { _start: { lean_object* x_1; @@ -38994,17 +38994,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_18386____closed__2() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18388____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrsWitMatchType___spec__1___closed__6; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18386____closed__1; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18388____closed__1; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18386____closed__3() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18388____closed__3() { _start: { lean_object* x_1; @@ -39012,13 +39012,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_18386____closed__4() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18388____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_Array_forInUnsafe_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrsWitMatchType___spec__1___closed__14; -x_3 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18386____closed__3; +x_3 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18388____closed__3; x_4 = lean_box(x_1); x_5 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_5, 0, x_4); @@ -39027,7 +39027,7 @@ lean_ctor_set(x_5, 2, x_3); return x_5; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18386____closed__5() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18388____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; @@ -39035,18 +39035,18 @@ x_1 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandSimpleMatch___closed__ x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrsWitMatchType___spec__1___closed__5; x_3 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandSimpleMatch___closed__3; x_4 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrsWitMatchType___spec__1___closed__6; -x_5 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18386____closed__1; +x_5 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18388____closed__1; x_6 = l_Lean_Name_mkStr5(x_1, x_2, x_3, x_4, x_5); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18386_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18388_(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_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18386____closed__2; -x_3 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18386____closed__4; -x_4 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18386____closed__5; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18388____closed__2; +x_3 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18388____closed__4; +x_4 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18388____closed__5; x_5 = l_Lean_Option_register___at_Lean_Elab_initFn____x40_Lean_Elab_AutoBound___hyg_7____spec__1(x_2, x_3, x_4, x_1); return x_5; } @@ -39380,7 +39380,7 @@ lean_object* x_11; lean_object* x_12; uint8_t x_13; x_11 = lean_ctor_get(x_8, 2); lean_inc(x_11); x_12 = l_Lean_Elab_Term_reportMatcherResultErrors___lambda__1___closed__1; -x_13 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_11, x_12); +x_13 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_11, x_12); lean_dec(x_11); if (x_13 == 0) { @@ -46387,7 +46387,7 @@ x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__1() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -46397,27 +46397,27 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__2() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__2() { _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_22719____closed__1; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__1; x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrsWitMatchType___spec__1___closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__3() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____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_22719____closed__2; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__2; x_2 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandSimpleMatch___closed__3; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__4() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__4() { _start: { lean_object* x_1; @@ -46425,17 +46425,17 @@ x_1 = lean_mk_string_from_bytes("initFn", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__5() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__5() { _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_22719____closed__3; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__4; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__3; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__6() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__6() { _start: { lean_object* x_1; @@ -46443,37 +46443,37 @@ x_1 = lean_mk_string_from_bytes("_@", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__7() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__7() { _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_22719____closed__5; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__6; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__5; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__8() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__8() { _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_22719____closed__7; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__7; x_2 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandSimpleMatch___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__9() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__9() { _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_22719____closed__8; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__8; x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrsWitMatchType___spec__1___closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__10() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__10() { _start: { lean_object* x_1; @@ -46481,17 +46481,17 @@ x_1 = lean_mk_string_from_bytes("Match", 5); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__11() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__11() { _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_22719____closed__9; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__10; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__9; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__10; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__12() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__12() { _start: { lean_object* x_1; @@ -46499,33 +46499,33 @@ x_1 = lean_mk_string_from_bytes("_hyg", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__13() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__13() { _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_22719____closed__11; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__12; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__11; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__12; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__14() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__14() { _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_22719____closed__13; -x_2 = lean_unsigned_to_nat(22719u); +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__13; +x_2 = lean_unsigned_to_nat(22721u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrsWitMatchType___spec__1___closed__7; x_3 = 0; -x_4 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__14; +x_4 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__14; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } @@ -47352,17 +47352,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_18386____closed__1 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18386____closed__1(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18386____closed__1); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18386____closed__2 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18386____closed__2(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18386____closed__2); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18386____closed__3 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18386____closed__3(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18386____closed__3); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18386____closed__4 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18386____closed__4(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18386____closed__4); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18386____closed__5 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18386____closed__5(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18386____closed__5); -if (builtin) {res = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18386_(lean_io_mk_world()); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18388____closed__1 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18388____closed__1(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18388____closed__1); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18388____closed__2 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18388____closed__2(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18388____closed__2); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18388____closed__3 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18388____closed__3(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18388____closed__3); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18388____closed__4 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18388____closed__4(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18388____closed__4); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18388____closed__5 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18388____closed__5(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18388____closed__5); +if (builtin) {res = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18388_(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); @@ -47475,35 +47475,35 @@ 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); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__1 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__1(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__1); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__2 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__2(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__2); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__3 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__3(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__3); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__4 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__4(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__4); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__5 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__5(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__5); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__6 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__6(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__6); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__7 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__7(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__7); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__8 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__8(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__8); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__9 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__9(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__9); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__10 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__10(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__10); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__11 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__11(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__11); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__12 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__12(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__12); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__13 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__13(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__13); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__14 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__14(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719____closed__14); -res = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22719_(lean_io_mk_world()); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__1 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__1(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__1); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__2 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__2(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__2); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__3 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__3(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__3); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__4 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__4(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__4); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__5 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__5(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__5); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__6 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__6(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__6); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__7 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__7(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__7); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__8 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__8(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__8); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__9 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__9(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__9); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__10 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__10(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__10); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__11 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__11(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__11); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__12 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__12(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__12); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__13 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__13(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__13); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__14 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__14(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721____closed__14); +res = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_22721_(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/MutualDef.c b/stage0/stdlib/Lean/Elab/MutualDef.c index 11ccc99d3fa9..4ed50dec36fd 100644 --- a/stage0/stdlib/Lean/Elab/MutualDef.c +++ b/stage0/stdlib/Lean/Elab/MutualDef.c @@ -89,6 +89,7 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutual static lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__5___closed__5; extern lean_object* l_Lean_declRangeExt; static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__2___closed__5; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, 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_mkLetRecClosureFor___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_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__4___closed__2; static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_cleanupOfNat___lambda__2___closed__1; @@ -99,7 +100,7 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Mutua static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__6___lambda__5___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_MutualClosure_ClosureState_newLetDecls___default; 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_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__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_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_FixPoint_isModified___rarg(lean_object*); lean_object* l_Lean_indentD(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_pushLocalDecl___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*); @@ -192,6 +193,7 @@ lean_object* lean_array_fget(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*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_checkForHiddenUnivLevels___spec__2___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapM___at_Lean_Elab_Term_elabMutualDef___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_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___closed__6; lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getTailPos_x3f(lean_object*, uint8_t); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_cleanupOfNat___lambda__3___closed__2; @@ -204,6 +206,7 @@ lean_object* l_Lean_Elab_logException___at_Lean_Elab_Term_exceptionToSorry___spe LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkFreeVarMap___spec__5___rarg___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_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosureFor___spec__5___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosureFor___spec__6(lean_object*, size_t, size_t, lean_object*); uint8_t l_Lean_Elab_DefKind_isDefOrAbbrevOrOpaque(uint8_t); +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2___closed__7; LEAN_EXPORT lean_object* l_Lean_RBNode_fold___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkFreeVarMap___spec__4(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__5___lambda__3___closed__4; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkFreeVarMap___spec__5___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkFreeVarMap___spec__6(lean_object*); @@ -213,7 +216,7 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Mutua LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__4(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabModifiers___at_Lean_Elab_Command_elabMutualDef___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_MutualClosure_main___spec__3(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__5___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, 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__5___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, size_t, 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__5___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef___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*); uint8_t l_Array_contains___at___private_Lean_Meta_Match_Value_0__Lean_Meta_isUIntTypeName___spec__1(lean_object*, lean_object*); @@ -302,7 +305,6 @@ LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at___private_Lean_Elab_MutualDef_ LEAN_EXPORT lean_object* l_Lean_Elab_Term_MutualClosure_FixPoint_run___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_Elab_Term_MutualClosure_pushLetRecs___spec__1___lambda__2___boxed(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_cleanupOfNat___closed__2; -lean_object* l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__3___rarg(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__5___lambda__3___closed__5; 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*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_checkModifiers___closed__2; @@ -311,7 +313,6 @@ static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkClosureForAux___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_instHashableArray___rarg___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__5___closed__2; -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__1___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_removeUnusedVars(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_mapTR_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_isMultiConstant_x3f___spec__1(lean_object*, lean_object*); size_t lean_ptr_addr(lean_object*); @@ -342,14 +343,13 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualD static lean_object* l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__8; static lean_object* l_List_forM___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_checkLetRecsToLiftTypes___spec__1___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef___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_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___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___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___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_Elab_MutualDef_0__Lean_Elab_Term_removeUnusedVars___closed__2; 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_object* l_Array_reverse___rarg(lean_object*); LEAN_EXPORT uint8_t l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_levelMVarToParamHeaders_process___spec__1___lambda__1(lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___lambda__1(lean_object*, size_t, lean_object*); 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_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__1___closed__6; 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**); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkFreeVarMap___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosures___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_EXPORT lean_object* l_Lean_Elab_withSaveInfoContext___at_Lean_Elab_Term_elabMutualDef___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -407,6 +407,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check static lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__5___closed__13; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosureFor___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___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___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_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabCommand___spec__13(lean_object*, 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_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_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef___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*); @@ -420,7 +421,7 @@ lean_object* l_Lean_Elab_toAttributeKind___boxed(lean_object*, lean_object*, lea LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__5___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, size_t, size_t, lean_object*); 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*); LEAN_EXPORT lean_object* l_Lean_RBNode_find___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosures___spec__13___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, size_t, 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__5___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_Elab_Term_MutualClosure_pushLetRecs___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* l_Lean_Syntax_getKind(lean_object*); lean_object* l_Lean_PersistentArray_append___rarg(lean_object*, lean_object*); @@ -458,6 +459,7 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Mutua lean_object* lean_st_ref_get(lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__6___lambda__5___closed__2; static lean_object* l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4___closed__1; +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__4___closed__2; 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*); extern lean_object* l_Lean_pp_universes; uint8_t l_List_isEmpty___rarg(lean_object*); @@ -471,9 +473,7 @@ LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0_ LEAN_EXPORT uint8_t l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__4___rarg___lambda__4___boxed(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*); -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__1___closed__7; LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_Elab_Term_MutualClosure_pushLetRecs___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forM___at_Lean_Elab_Term_checkForHiddenUnivLevels_visit___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_MutualDef_0__Lean_Elab_Term_checkModifiers___closed__1; LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Elab_Term_MutualClosure_main___spec__7(lean_object*, lean_object*, lean_object*); @@ -569,7 +569,7 @@ lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_isTheorem___boxed(lean_object*); uint8_t l_Lean_LocalContext_contains(lean_object*, 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_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__1(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__5___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Elab_Term_checkForHiddenUnivLevels_visit___spec__4(lean_object*); uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMutualDef___spec__13___closed__2; @@ -580,6 +580,7 @@ static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_getPendind uint8_t l_Lean_Syntax_isIdent(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*); 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__5___lambda__4___closed__3; LEAN_EXPORT uint8_t l_Lean_Elab_Term_MutualClosure_FixPoint_State_modified___default; LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkClosureForAux___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_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosures___spec__7___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosures___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*); @@ -639,6 +640,7 @@ LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Elab_Term_MutualClosure_Re LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Elab_Term_checkForHiddenUnivLevels_visit___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_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_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__4___rarg___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__5___lambda__4___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__6(lean_object*, size_t, size_t, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__5___rarg___closed__1; lean_object* l_Lean_Elab_Term_withAutoBoundImplicit___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -654,6 +656,7 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Mutua static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___closed__1; static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_registerFailedToInferDefTypeInfo___closed__2; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_processDeriving___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*); +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___closed__5; LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withUsed___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_List_mapTR_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_getPendindMVarErrorMessage___spec__2(lean_object*, lean_object*); lean_object* l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_withoutModifyingElabMetaStateWithInfo___spec__2___rarg(lean_object*, lean_object*); @@ -668,7 +671,9 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDe LEAN_EXPORT lean_object* l_Lean_resolveGlobalConst___at_Lean_Elab_Term_elabMutualDef_processDeriving___spec__2(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__2(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*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__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_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_object* l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_FindImpl_findUnsafe_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___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_addDeclarationRanges___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__3___closed__4; @@ -711,15 +716,16 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_getFu LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__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_Command_elabMutualDef___spec__13___lambda__1___boxed(lean_object*, lean_object*, lean_object*, 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_elabCommand___spec__14(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_throwError___at_Lean_Elab_Term_synthesizeInstMVarCore___spec__3(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___lambda__3___closed__1; 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_Array_anyMUnsafe_any___at_Lean_Elab_Term_MutualClosure_getModifiersForLetRecs___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__5___lambda__4___closed__4; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkClosureForAux___spec__1___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__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_Lean_Elab_Command_elabMutualDef(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Term_checkForHiddenUnivLevels___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofLevel(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*); @@ -793,7 +799,6 @@ static lean_object* l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___ LEAN_EXPORT lean_object* l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerm___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapM___at_Lean_Elab_Term_elabMutualDef___spec__7(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__5___lambda__1___closed__2; static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___closed__3; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_levelMVarToParamHeaders_process___spec__1___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___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*); @@ -806,10 +811,10 @@ lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_mkDefViewOfInstance__ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_MutualClosure_main___spec__5(lean_object*, size_t, size_t, lean_object*); static lean_object* l_List_forM___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_checkLetRecsToLiftTypes___spec__1___closed__4; uint8_t l_Lean_Expr_isProp(lean_object*); -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__1___closed__5; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef___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_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___closed__6; LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkFreeVarMap___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosures___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___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2___closed__6; lean_object* l_Lean_Meta_InfoCacheKey_instHashableInfoCacheKey___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Expr_ReplaceImpl_replaceUnsafeM_visit___at_Lean_Elab_Term_MutualClosure_Replacement_apply___spec__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_elabMutualDef___spec__13___closed__1; @@ -983,6 +988,7 @@ lean_object* l_Lean_Elab_Term_expandWhereDecls(lean_object*, lean_object*, lean_ lean_object* lean_array_uset(lean_object*, size_t, lean_object*); lean_object* l_Lean_Expr_ReplaceImpl_Cache_store(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofName(lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__4___closed__5; static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___closed__2; lean_object* lean_expr_instantiate1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__3(lean_object*); @@ -997,11 +1003,10 @@ LEAN_EXPORT lean_object* l_Nat_foldTR_loop___at_Lean_Elab_Term_MutualClosure_ins LEAN_EXPORT lean_object* l_Array_getMax_x3f___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_pickMaxFVar_x3f___spec__1___boxed(lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_getLetRecsToLift___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef___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_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___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___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_collectFVars(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_MutualDef_0__Lean_Elab_Term_isTheorem___spec__1(lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Lean_RBNode_fold___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosures___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__5___lambda__1___closed__3; LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Elab_Term_MutualClosure_Replacement_apply___spec__1___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___closed__2; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosures___spec__6(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*); @@ -3320,7 +3325,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_clean lean_object* x_8; lean_object* x_9; uint8_t x_10; lean_dec(x_2); x_8 = l_Lean_Expr_appArg_x21(x_1); -x_9 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_8, x_3, x_4, x_5, x_6, x_7); +x_9 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_8, x_3, x_4, x_5, x_6, x_7); x_10 = !lean_is_exclusive(x_9); if (x_10 == 0) { @@ -6595,23 +6600,40 @@ static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Mutual _start: { lean_object* x_1; +x_1 = lean_mk_string_from_bytes("'_' is not allowed here", 23); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___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; +x_5 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__1___closed__1; +x_6 = l_Lean_Macro_throwErrorAt___rarg(x_1, x_5, x_3, x_4); +return x_6; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2___closed__1() { +_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___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__1___closed__2() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2___closed__2() { _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_Elab_addDeclarationRanges___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__3___closed__1; x_2 = l_Lean_Elab_addDeclarationRanges___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__3___closed__2; x_3 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___closed__1; -x_4 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__1___closed__1; +x_4 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2___closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__1___closed__3() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2___closed__3() { _start: { lean_object* x_1; @@ -6619,19 +6641,19 @@ 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__5___lambda__1___closed__4() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_Elab_addDeclarationRanges___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__3___closed__1; x_2 = l_Lean_Elab_addDeclarationRanges___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__3___closed__2; x_3 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___closed__1; -x_4 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__1___closed__3; +x_4 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2___closed__3; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__1___closed__5() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2___closed__5() { _start: { lean_object* x_1; @@ -6639,17 +6661,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__5___lambda__1___closed__6() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___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_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__1___closed__5; +x_2 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2___closed__5; x_3 = l_Lean_Name_str___override(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__5___lambda__1___closed__7() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2___closed__7() { _start: { lean_object* x_1; @@ -6657,7 +6679,7 @@ 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__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_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___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) { _start: { lean_object* x_7; uint8_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_object* x_18; @@ -6666,21 +6688,21 @@ lean_inc(x_7); lean_dec(x_5); x_8 = 0; x_9 = l_Lean_SourceInfo_fromRef(x_7, x_8); -x_10 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__1___closed__6; +x_10 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2___closed__6; lean_inc(x_9); 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_1); -x_12 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__1___closed__4; +x_12 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2___closed__4; lean_inc(x_9); x_13 = l_Lean_Syntax_node2(x_9, x_12, x_2, x_11); -x_14 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__1___closed__7; +x_14 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2___closed__7; lean_inc(x_9); x_15 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_15, 0, x_9); lean_ctor_set(x_15, 1, x_14); -x_16 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__1___closed__2; +x_16 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2___closed__2; x_17 = l_Lean_Syntax_node3(x_9, x_16, x_13, x_15, x_4); x_18 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_18, 0, x_17); @@ -6688,7 +6710,7 @@ lean_ctor_set(x_18, 1, x_6); return x_18; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2___closed__1() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__3___closed__1() { _start: { lean_object* x_1; @@ -6696,19 +6718,19 @@ 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__5___lambda__2___closed__2() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__3___closed__2() { _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_Elab_addDeclarationRanges___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__3___closed__1; x_2 = l_Lean_Elab_addDeclarationRanges___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__3___closed__2; x_3 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___closed__1; -x_4 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2___closed__1; +x_4 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__3___closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2___closed__3() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__3___closed__3() { _start: { lean_object* x_1; @@ -6716,19 +6738,19 @@ 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__5___lambda__2___closed__4() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__3___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_Elab_addDeclarationRanges___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__3___closed__1; x_2 = l_Lean_Elab_addDeclarationRanges___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__3___closed__2; x_3 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___closed__1; -x_4 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2___closed__3; +x_4 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__3___closed__3; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2___closed__5() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__3___closed__5() { _start: { lean_object* x_1; @@ -6736,7 +6758,7 @@ 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__5___lambda__2(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__5___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) { _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; @@ -6750,7 +6772,7 @@ lean_dec(x_10); lean_dec(x_4); lean_dec(x_3); lean_inc(x_6); -x_13 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__1(x_1, x_2, x_6, x_6, x_8, x_9); +x_13 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2(x_1, x_2, x_6, x_6, x_8, x_9); lean_dec(x_6); return x_13; } @@ -6761,7 +6783,7 @@ x_14 = lean_ctor_get(x_8, 5); lean_inc(x_14); x_15 = 0; x_16 = l_Lean_SourceInfo_fromRef(x_14, x_15); -x_17 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2___closed__1; +x_17 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__3___closed__1; lean_inc(x_16); x_18 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_18, 0, x_16); @@ -6772,7 +6794,7 @@ x_20 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Te lean_dec(x_4); lean_inc(x_1); x_21 = l_Array_append___rarg(x_1, x_20); -x_22 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__1___closed__6; +x_22 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2___closed__6; lean_inc(x_16); x_23 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_23, 0, x_16); @@ -6784,24 +6806,24 @@ x_24 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_24, 0, x_16); lean_ctor_set(x_24, 1, x_22); lean_ctor_set(x_24, 2, x_1); -x_25 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2___closed__5; +x_25 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__3___closed__5; lean_inc(x_16); x_26 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_26, 0, x_16); 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__5___lambda__2___closed__4; +x_27 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__3___closed__4; lean_inc(x_6); lean_inc(x_16); x_28 = l_Lean_Syntax_node4(x_16, x_27, x_23, x_24, x_26, x_6); -x_29 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2___closed__2; +x_29 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__3___closed__2; x_30 = l_Lean_Syntax_node2(x_16, x_29, x_18, x_28); -x_31 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__1(x_1, x_2, x_6, x_30, x_8, x_9); +x_31 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2(x_1, x_2, x_6, x_30, x_8, x_9); lean_dec(x_6); return x_31; } } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__3___closed__1() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__4___closed__1() { _start: { lean_object* x_1; @@ -6809,19 +6831,19 @@ 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__5___lambda__3___closed__2() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__4___closed__2() { _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_Elab_addDeclarationRanges___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__3___closed__1; x_2 = l_Lean_Elab_addDeclarationRanges___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__3___closed__2; x_3 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___closed__1; -x_4 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__3___closed__1; +x_4 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__4___closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__3___closed__3() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__4___closed__3() { _start: { lean_object* x_1; @@ -6829,7 +6851,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__5___lambda__3___closed__4() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__4___closed__4() { _start: { lean_object* x_1; @@ -6837,7 +6859,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__5___lambda__3___closed__5() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__4___closed__5() { _start: { lean_object* x_1; @@ -6845,7 +6867,7 @@ 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__5___lambda__3(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__5___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) { _start: { lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; @@ -6869,7 +6891,7 @@ if (lean_obj_tag(x_8) == 0) lean_object* x_17; lean_object* x_18; lean_dec(x_16); x_17 = lean_box(0); -x_18 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2(x_3, x_4, x_13, x_5, x_6, x_12, x_17, x_9, x_10); +x_18 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__3(x_3, x_4, x_13, x_5, x_6, x_12, x_17, x_9, x_10); return x_18; } else @@ -6880,28 +6902,28 @@ lean_inc(x_19); lean_dec(x_8); x_20 = 0; x_21 = l_Lean_SourceInfo_fromRef(x_16, x_20); -x_22 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__3___closed__3; +x_22 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__4___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_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__3___closed__4; +x_24 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__4___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_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__1___closed__6; +x_26 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2___closed__6; lean_inc(x_21); x_27 = l_Lean_Syntax_node1(x_21, x_26, x_19); -x_28 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__3___closed__5; +x_28 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__4___closed__5; lean_inc(x_21); x_29 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_29, 0, x_21); 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__5___lambda__3___closed__2; +x_30 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__4___closed__2; x_31 = l_Lean_Syntax_node5(x_21, x_30, x_23, x_12, x_25, x_27, x_29); x_32 = lean_box(0); -x_33 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2(x_3, x_4, x_13, x_5, x_6, x_31, x_32, x_9, x_10); +x_33 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__3(x_3, x_4, x_13, x_5, x_6, x_31, x_32, x_9, x_10); return x_33; } } @@ -6937,7 +6959,7 @@ if (lean_obj_tag(x_8) == 0) lean_object* x_42; lean_object* x_43; lean_dec(x_40); x_42 = lean_box(0); -x_43 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2(x_3, x_4, x_13, x_5, x_6, x_12, x_42, x_41, x_10); +x_43 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__3(x_3, x_4, x_13, x_5, x_6, x_12, x_42, x_41, x_10); return x_43; } else @@ -6948,28 +6970,28 @@ lean_inc(x_44); lean_dec(x_8); x_45 = 0; x_46 = l_Lean_SourceInfo_fromRef(x_40, x_45); -x_47 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__3___closed__3; +x_47 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__4___closed__3; lean_inc(x_46); 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_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__3___closed__4; +x_49 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__4___closed__4; lean_inc(x_46); x_50 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_50, 0, x_46); lean_ctor_set(x_50, 1, x_49); -x_51 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__1___closed__6; +x_51 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2___closed__6; lean_inc(x_46); x_52 = l_Lean_Syntax_node1(x_46, x_51, x_44); -x_53 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__3___closed__5; +x_53 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__4___closed__5; lean_inc(x_46); x_54 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_54, 0, x_46); lean_ctor_set(x_54, 1, x_53); -x_55 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__3___closed__2; +x_55 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__4___closed__2; x_56 = l_Lean_Syntax_node5(x_46, x_55, x_48, x_12, x_50, x_52, x_54); x_57 = lean_box(0); -x_58 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2(x_3, x_4, x_13, x_5, x_6, x_56, x_57, x_41, x_10); +x_58 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__3(x_3, x_4, x_13, x_5, x_6, x_56, x_57, x_41, x_10); return x_58; } } @@ -6997,7 +7019,7 @@ static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Mutual _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("typeSpec", 8); +x_1 = lean_mk_string_from_bytes("hole", 4); return x_1; } } @@ -7013,6 +7035,26 @@ x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("typeSpec", 8); +return x_1; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___closed__6() { +_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_Elab_addDeclarationRanges___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__3___closed__1; +x_2 = l_Lean_Elab_addDeclarationRanges___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__3___closed__2; +x_3 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___closed__1; +x_4 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___closed__5; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5(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) { _start: { @@ -7031,156 +7073,260 @@ return x_10; } else { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_21; uint8_t x_22; 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_14 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___closed__9; +x_21 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___closed__9; lean_inc(x_11); -x_15 = l_Lean_Syntax_isOfKind(x_11, x_14); -if (x_15 == 0) +x_22 = l_Lean_Syntax_isOfKind(x_11, x_21); +if (x_22 == 0) { -lean_object* x_16; lean_object* x_17; +lean_object* x_23; lean_object* x_24; lean_dec(x_13); lean_dec(x_11); lean_dec(x_7); 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_8); -return x_17; +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_8); +return x_24; } else { -lean_object* x_18; lean_object* x_19; uint8_t x_20; -x_18 = l_Lean_Syntax_getArg(x_11, x_12); -x_19 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___closed__2; -lean_inc(x_18); -x_20 = l_Lean_Syntax_isOfKind(x_18, x_19); -if (x_20 == 0) +lean_object* x_25; lean_object* x_26; uint8_t x_27; +x_25 = l_Lean_Syntax_getArg(x_11, x_12); +x_26 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___closed__2; +lean_inc(x_25); +x_27 = l_Lean_Syntax_isOfKind(x_25, x_26); +if (x_27 == 0) { -lean_object* x_21; lean_object* x_22; -lean_dec(x_18); +lean_object* x_28; uint8_t x_29; lean_dec(x_13); -lean_dec(x_11); -lean_dec(x_7); 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_8); -return x_22; +x_28 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___closed__4; +x_29 = l_Lean_Syntax_isOfKind(x_25, x_28); +if (x_29 == 0) +{ +lean_object* x_30; lean_object* x_31; +lean_dec(x_11); +lean_dec(x_7); +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_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; -x_23 = lean_unsigned_to_nat(1u); -x_24 = l_Lean_Syntax_getArg(x_11, x_23); -x_25 = lean_unsigned_to_nat(2u); -x_26 = l_Lean_Syntax_getArg(x_11, x_25); -x_27 = l_Lean_Syntax_isNone(x_26); -if (x_27 == 0) +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_11, x_32); +x_34 = l_Lean_Syntax_isNone(x_33); +if (x_34 == 0) { -uint8_t x_28; -lean_inc(x_26); -x_28 = l_Lean_Syntax_matchesNull(x_26, x_23); -if (x_28 == 0) +lean_object* x_35; uint8_t x_36; +x_35 = lean_unsigned_to_nat(1u); +lean_inc(x_33); +x_36 = l_Lean_Syntax_matchesNull(x_33, x_35); +if (x_36 == 0) { -lean_object* x_29; lean_object* x_30; -lean_dec(x_26); -lean_dec(x_24); -lean_dec(x_18); +lean_object* x_37; lean_object* x_38; +lean_dec(x_33); +lean_dec(x_11); +lean_dec(x_7); +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_8); +return x_38; +} +else +{ +lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_39 = l_Lean_Syntax_getArg(x_33, x_12); +lean_dec(x_33); +x_40 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___closed__6; +x_41 = l_Lean_Syntax_isOfKind(x_39, x_40); +if (x_41 == 0) +{ +lean_object* x_42; lean_object* x_43; +lean_dec(x_11); +lean_dec(x_7); +x_42 = lean_box(1); +x_43 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_8); +return x_43; +} +else +{ +lean_object* x_44; lean_object* x_45; uint8_t x_46; +x_44 = lean_mk_string_from_bytes("'_' is not allowed here", 23); +x_45 = l_Lean_Macro_throwErrorAt___rarg(x_11, x_44, x_7, x_8); +lean_dec(x_11); +x_46 = !lean_is_exclusive(x_45); +if (x_46 == 0) +{ +return x_45; +} +else +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_ctor_get(x_45, 0); +x_48 = lean_ctor_get(x_45, 1); +lean_inc(x_48); +lean_inc(x_47); +lean_dec(x_45); +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 +{ +lean_object* x_50; lean_object* x_51; uint8_t x_52; +lean_dec(x_33); +x_50 = lean_mk_string_from_bytes("'_' is not allowed here", 23); +x_51 = l_Lean_Macro_throwErrorAt___rarg(x_11, x_50, x_7, x_8); +lean_dec(x_11); +x_52 = !lean_is_exclusive(x_51); +if (x_52 == 0) +{ +return x_51; +} +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_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; +} +} +} +} +else +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; uint8_t x_60; +x_56 = lean_unsigned_to_nat(1u); +x_57 = l_Lean_Syntax_getArg(x_11, x_56); +x_58 = lean_unsigned_to_nat(2u); +x_59 = l_Lean_Syntax_getArg(x_11, x_58); +x_60 = l_Lean_Syntax_isNone(x_59); +if (x_60 == 0) +{ +uint8_t x_61; +lean_inc(x_59); +x_61 = l_Lean_Syntax_matchesNull(x_59, x_56); +if (x_61 == 0) +{ +lean_object* x_62; lean_object* x_63; +lean_dec(x_59); +lean_dec(x_57); +lean_dec(x_25); lean_dec(x_13); lean_dec(x_11); lean_dec(x_7); lean_dec(x_2); lean_dec(x_1); -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_8); -return x_30; +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_8); +return x_63; } else { -lean_object* x_31; lean_object* x_32; uint8_t x_33; -x_31 = l_Lean_Syntax_getArg(x_26, x_12); -lean_dec(x_26); -x_32 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___closed__4; -lean_inc(x_31); -x_33 = l_Lean_Syntax_isOfKind(x_31, x_32); -if (x_33 == 0) +lean_object* x_64; lean_object* x_65; uint8_t x_66; +x_64 = l_Lean_Syntax_getArg(x_59, x_12); +lean_dec(x_59); +x_65 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___closed__6; +lean_inc(x_64); +x_66 = l_Lean_Syntax_isOfKind(x_64, x_65); +if (x_66 == 0) { -lean_object* x_34; lean_object* x_35; -lean_dec(x_31); -lean_dec(x_24); -lean_dec(x_18); +lean_object* x_67; lean_object* x_68; +lean_dec(x_64); +lean_dec(x_57); +lean_dec(x_25); lean_dec(x_13); lean_dec(x_11); lean_dec(x_7); lean_dec(x_2); lean_dec(x_1); -x_34 = lean_box(1); -x_35 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_35, 0, x_34); -lean_ctor_set(x_35, 1, x_8); -return x_35; +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_8); +return x_68; } 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; size_t x_42; size_t x_43; lean_object* x_44; -x_36 = l_Lean_Syntax_getArg(x_31, x_23); -lean_dec(x_31); -x_37 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_37, 0, x_36); -x_38 = lean_box(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; +x_69 = l_Lean_Syntax_getArg(x_64, x_56); +lean_dec(x_64); +x_70 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_70, 0, x_69); +x_71 = lean_box(0); lean_inc(x_7); lean_inc(x_2); lean_inc(x_1); -x_39 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__3(x_11, x_24, x_1, x_18, x_2, x_3, x_38, x_37, x_7, x_8); -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 = 1; -x_43 = lean_usize_add(x_5, x_42); -x_44 = lean_array_uset(x_13, x_5, x_40); -x_5 = x_43; -x_6 = x_44; -x_8 = x_41; -goto _start; +x_72 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__4(x_11, x_57, x_1, x_25, x_2, x_3, x_71, x_70, x_7, x_8); +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_14 = x_73; +x_15 = x_74; +goto block_20; } } } else { -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_dec(x_26); -x_46 = lean_box(0); -x_47 = lean_box(0); +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; +lean_dec(x_59); +x_75 = lean_box(0); +x_76 = lean_box(0); lean_inc(x_7); lean_inc(x_2); lean_inc(x_1); -x_48 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__3(x_11, x_24, x_1, x_18, x_2, x_3, x_47, x_46, x_7, x_8); -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 = 1; -x_52 = lean_usize_add(x_5, x_51); -x_53 = lean_array_uset(x_13, x_5, x_49); -x_5 = x_52; -x_6 = x_53; -x_8 = x_50; -goto _start; +x_77 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__4(x_11, x_57, x_1, x_25, x_2, x_3, x_76, x_75, x_7, x_8); +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_14 = x_78; +x_15 = x_79; +goto block_20; } } } +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; +} } } } @@ -7321,7 +7467,7 @@ static lean_object* _init_l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expa { 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__5___lambda__1___closed__2; +x_2 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___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); @@ -7431,7 +7577,7 @@ 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_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__1___closed__6; +x_25 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2___closed__6; lean_inc(x_2); lean_inc(x_22); x_26 = lean_alloc_ctor(1, 3, 0); @@ -7498,7 +7644,7 @@ lean_inc(x_47); x_49 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_49, 0, x_47); lean_ctor_set(x_49, 1, x_48); -x_50 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__1___closed__6; +x_50 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2___closed__6; lean_inc(x_2); lean_inc(x_47); x_51 = lean_alloc_ctor(1, 3, 0); @@ -7940,33 +8086,43 @@ lean_dec(x_1); return x_7; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___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_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___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_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__1(x_1, x_2, x_3, x_4); +lean_dec(x_2); +lean_dec(x_1); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___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) { _start: { lean_object* x_7; -x_7 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6); +x_7 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6); lean_dec(x_3); return x_7; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___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_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___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: { size_t x_10; lean_object* x_11; 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__5___lambda__2(x_1, x_2, x_3, x_4, x_10, x_6, x_7, x_8, x_9); +x_11 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__3(x_1, x_2, x_3, x_4, x_10, x_6, x_7, x_8, x_9); lean_dec(x_7); return x_11; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___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_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___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__5___lambda__3(x_1, x_2, x_3, x_4, x_5, x_11, x_7, x_8, x_9, x_10); +x_12 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___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; } } @@ -19159,7 +19315,7 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_29 = l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__3___rarg(x_16, x_17, x_28, x_5, x_6, x_7, x_8, x_9); +x_29 = l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2___rarg(x_16, x_17, x_28, x_5, x_6, x_7, x_8, x_9); if (lean_obj_tag(x_29) == 0) { lean_object* x_30; lean_object* x_31; uint8_t x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; @@ -30415,7 +30571,7 @@ lean_closure_set(x_14, 0, x_12); lean_closure_set(x_14, 1, x_13); lean_inc(x_3); lean_inc(x_2); -x_15 = l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabCommand___spec__14(x_14, x_2, x_3, x_10); +x_15 = l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabCommand___spec__13(x_14, x_2, x_3, x_10); if (lean_obj_tag(x_15) == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; @@ -31802,6 +31958,7 @@ lean_closure_set(x_12, 0, x_10); lean_closure_set(x_12, 1, x_2); x_13 = l_Lean_Elab_Command_runTermElabM___rarg(x_12, x_3, x_4, x_11); lean_dec(x_4); +lean_dec(x_3); return x_13; } else @@ -32221,18 +32378,6 @@ l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expa lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__3___closed__10); l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__1___closed__1 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__1___closed__1(); lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__1___closed__1); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__1___closed__2 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__1___closed__2(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__1___closed__2); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__1___closed__3 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__1___closed__3(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__1___closed__3); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__1___closed__4 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__1___closed__4(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__1___closed__4); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__1___closed__5 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__1___closed__5(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__1___closed__5); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__1___closed__6 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__1___closed__6(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__1___closed__6); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__1___closed__7 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__1___closed__7(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__1___closed__7); l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2___closed__1 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2___closed__1(); lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2___closed__1); l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2___closed__2 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2___closed__2(); @@ -32243,6 +32388,10 @@ l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expa lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2___closed__4); l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2___closed__5 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2___closed__5(); lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2___closed__5); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2___closed__6 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2___closed__6(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2___closed__6); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2___closed__7 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2___closed__7(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2___closed__7); l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__3___closed__1 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__3___closed__1(); lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__3___closed__1); l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__3___closed__2 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__3___closed__2(); @@ -32253,6 +32402,16 @@ l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expa lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__3___closed__4); l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__3___closed__5 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__3___closed__5(); lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__3___closed__5); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__4___closed__1 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__4___closed__1(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__4___closed__1); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__4___closed__2 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__4___closed__2(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__4___closed__2); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__4___closed__3 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__4___closed__3(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__4___closed__3); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__4___closed__4 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__4___closed__4(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__4___closed__4); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__4___closed__5 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__4___closed__5(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__4___closed__5); l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___closed__1 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___closed__1(); lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___closed__1); l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___closed__2 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___closed__2(); @@ -32261,6 +32420,10 @@ l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expa lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___closed__3); l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___closed__4 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___closed__4(); lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___closed__4); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___closed__5 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___closed__5(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___closed__5); +l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___closed__6 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___closed__6(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___closed__6); 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__1___closed__2 = _init_l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__1___closed__2(); diff --git a/stage0/stdlib/Lean/Elab/PreDefinition/Eqns.c b/stage0/stdlib/Lean/Elab/PreDefinition/Eqns.c index b47ba8efee21..03c1477de752 100644 --- a/stage0/stdlib/Lean/Elab/PreDefinition/Eqns.c +++ b/stage0/stdlib/Lean/Elab/PreDefinition/Eqns.c @@ -177,7 +177,6 @@ extern lean_object* l_Lean_Meta_smartUnfolding; extern lean_object* l_instInhabitedPUnit; lean_object* l_Lean_Expr_replaceFVarId(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_HashSetImp_insert___at_Lean_ForEachExprWhere_checked___spec__3(lean_object*, lean_object*); -lean_object* l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__3___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___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___spec__30___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_Expr_hasMVar(lean_object*); lean_object* l_instHashableArray___rarg___boxed(lean_object*, lean_object*); @@ -282,7 +281,6 @@ lean_object* l_Lean_Expr_replaceFVar(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_getUnfoldFor_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Eqns_mkUnfoldProof_go___lambda__1___closed__2; -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withMVarContextImp___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_findMatchToSplit_x3f___lambda__2___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_findMatchToSplit_x3f___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -400,6 +398,7 @@ LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visited___at_Lean_Elab_Eqns_sim LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__19___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__20___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__9___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__10(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_FindImpl_findUnsafe_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_checked___at_Lean_Elab_Eqns_simpEqnType_collect___spec__5(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldrMAux___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -429,6 +428,7 @@ LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_checked___at_Lean_Elab_Eqns_sim lean_object* l_Lean_MVarId_isAssigned___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldrMAux___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Eqns_getUnfoldFor_x3f___closed__5; +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Eqns_getUnfoldFor_x3f___closed__8; LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_mkEqnTypes_go___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_foldrMUnsafe_fold___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___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*); @@ -6338,7 +6338,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_simpEqnType(lean_object* x_1, lean_obj { lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_inc(x_1); -x_7 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_1, x_2, x_3, x_4, x_5, x_6); +x_7 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_1, x_2, x_3, x_4, x_5, x_6); x_8 = lean_ctor_get(x_7, 0); lean_inc(x_8); x_9 = lean_ctor_get(x_7, 1); @@ -6595,7 +6595,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_E { lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; x_9 = l_Lean_LocalDecl_type(x_3); -x_10 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_9, x_4, x_5, x_6, x_7, x_8); +x_10 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_9, x_4, x_5, x_6, x_7, x_8); x_11 = lean_ctor_get(x_10, 0); lean_inc(x_11); x_12 = lean_ctor_get(x_10, 1); @@ -22223,7 +22223,7 @@ lean_inc(x_30); x_31 = lean_ctor_get(x_29, 1); lean_inc(x_31); lean_dec(x_29); -x_32 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_21, x_6, x_7, x_8, x_9, x_31); +x_32 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_21, x_6, x_7, 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); @@ -22525,7 +22525,7 @@ lean_closure_set(x_8, 0, x_1); lean_closure_set(x_8, 1, x_2); x_9 = l_Lean_Elab_Eqns_removeUnusedEqnHypotheses___closed__7; x_10 = l_Lean_RBTree_toArray___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn_collectDeps___spec__1___closed__1; -x_11 = l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__3___rarg(x_9, x_10, x_8, x_3, x_4, x_5, x_6, x_7); +x_11 = l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2___rarg(x_9, x_10, x_8, x_3, x_4, x_5, x_6, x_7); return x_11; } } diff --git a/stage0/stdlib/Lean/Elab/PreDefinition/Structural/Eqns.c b/stage0/stdlib/Lean/Elab/PreDefinition/Structural/Eqns.c index 0d5523cb8740..98f1caf8a626 100644 --- a/stage0/stdlib/Lean/Elab/PreDefinition/Structural/Eqns.c +++ b/stage0/stdlib/Lean/Elab/PreDefinition/Structural/Eqns.c @@ -74,7 +74,6 @@ lean_object* l_Lean_Elab_Eqns_deltaRHS_x3f(lean_object*, lean_object*, lean_obje lean_object* l_Lean_Elab_Eqns_getUnfoldFor_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1715____closed__2; lean_object* l_Lean_Meta_mkEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); lean_object* l_Lean_Name_append(lean_object*, lean_object*); extern lean_object* l_Lean_instInhabitedExpr; @@ -108,6 +107,7 @@ static lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition uint8_t lean_nat_dec_eq(lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1715____closed__8; +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___closed__6; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Structural_mkEqns___spec__1___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Structural_mkEqns___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*); @@ -1340,7 +1340,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__ _start: { lean_object* x_8; -x_8 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_1, x_3, x_4, x_5, x_6, x_7); +x_8 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_1, x_3, x_4, x_5, x_6, x_7); return x_8; } } @@ -1414,7 +1414,7 @@ lean_object* x_25; lean_object* x_26; x_25 = lean_ctor_get(x_24, 1); lean_inc(x_25); lean_dec(x_24); -x_26 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_10, x_3, x_4, x_5, x_6, x_25); +x_26 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_10, x_3, x_4, x_5, x_6, x_25); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); @@ -1486,7 +1486,7 @@ lean_dec(x_2); x_35 = lean_ctor_get(x_17, 1); lean_inc(x_35); lean_dec(x_17); -x_36 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_10, x_3, x_4, x_5, x_6, x_35); +x_36 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_10, x_3, x_4, x_5, x_6, x_35); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); diff --git a/stage0/stdlib/Lean/Elab/PreDefinition/Structural/IndPred.c b/stage0/stdlib/Lean/Elab/PreDefinition/Structural/IndPred.c index 105e0525435e..876a4f7f42d8 100644 --- a/stage0/stdlib/Lean/Elab/PreDefinition/Structural/IndPred.c +++ b/stage0/stdlib/Lean/Elab/PreDefinition/Structural/IndPred.c @@ -43,7 +43,6 @@ lean_object* lean_mk_array(lean_object*, lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); lean_object* l_Lean_Expr_mdata___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallBoundedTelescope___at_Lean_Elab_Structural_mkIndPredBRecOn___spec__1___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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*); lean_object* l_Lean_Expr_proj___override(lean_object*, lean_object*, lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); lean_object* l_Lean_Meta_Match_MatcherInfo_arity(lean_object*); @@ -69,6 +68,7 @@ static lean_object* l_Lean_Elab_Structural_mkIndPredBRecOn___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at___private_Lean_Elab_PreDefinition_Structural_IndPred_0__Lean_Elab_Structural_replaceIndPredRecApps_loop___spec__10___rarg(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___private_Lean_Elab_PreDefinition_Structural_IndPred_0__Lean_Elab_Structural_replaceIndPredRecApps_loop___closed__3; lean_object* l_List_mapTR_loop___at_Lean_MessageData_instCoeListExprMessageData___spec__1(lean_object*, lean_object*); +lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_Structural_IndPred_0__Lean_Elab_Structural_replaceIndPredRecApps_loop___lambda__2___closed__1; uint8_t l_Array_contains___at_Lean_Elab_Term_collectUnassignedMVars_go___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Structural_mkIndPredBRecOn___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*); @@ -3426,7 +3426,7 @@ lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; x_11 = lean_ctor_get(x_8, 2); lean_inc(x_11); x_12 = l___private_Lean_Elab_PreDefinition_Structural_IndPred_0__Lean_Elab_Structural_replaceIndPredRecApps___closed__1; -x_13 = l_Lean_Option_get___at_Std_Format_pretty_x27___spec__1(x_11, x_12); +x_13 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_11, x_12); lean_dec(x_11); x_14 = l___private_Lean_Elab_PreDefinition_Structural_IndPred_0__Lean_Elab_Structural_replaceIndPredRecApps_loop(x_1, x_2, x_3, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10); return x_14; diff --git a/stage0/stdlib/Lean/Elab/PreDefinition/WF/Eqns.c b/stage0/stdlib/Lean/Elab/PreDefinition/WF/Eqns.c index 31db66cfb230..d96c32823e44 100644 --- a/stage0/stdlib/Lean/Elab/PreDefinition/WF/Eqns.c +++ b/stage0/stdlib/Lean/Elab/PreDefinition/WF/Eqns.c @@ -149,7 +149,6 @@ lean_object* l_Lean_Elab_Eqns_getUnfoldFor_x3f(lean_object*, lean_object*, lean_ static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_3809____closed__1; lean_object* lean_array_to_list(lean_object*, lean_object*); lean_object* l_Lean_Meta_mkEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); lean_object* l_Lean_Name_append(lean_object*, lean_object*); extern lean_object* l_Lean_levelZero; @@ -211,6 +210,7 @@ static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_ uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_rwFixEq___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_getFixedPrefix___lambda__3___closed__2; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_mkEqns___spec__1___lambda__1___closed__2; static lean_object* l_Lean_Elab_WF_instInhabitedEqnInfo___closed__3; @@ -3686,7 +3686,7 @@ lean_inc(x_11); x_12 = lean_ctor_get(x_10, 1); lean_inc(x_12); lean_dec(x_10); -x_13 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_11, x_5, x_6, x_7, x_8, x_12); +x_13 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_11, x_5, x_6, x_7, x_8, x_12); x_14 = lean_ctor_get(x_13, 0); lean_inc(x_14); x_15 = lean_ctor_get(x_13, 1); @@ -5894,7 +5894,7 @@ lean_object* x_30; lean_object* x_31; x_30 = lean_ctor_get(x_29, 1); lean_inc(x_30); lean_dec(x_29); -x_31 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_11, x_4, x_5, x_6, x_7, x_30); +x_31 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_11, x_4, x_5, x_6, x_7, x_30); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); diff --git a/stage0/stdlib/Lean/Elab/PreDefinition/WF/Fix.c b/stage0/stdlib/Lean/Elab/PreDefinition/WF/Fix.c index 386fb2162d5c..98e5174137a0 100644 --- a/stage0/stdlib/Lean/Elab/PreDefinition/WF/Fix.c +++ b/stage0/stdlib/Lean/Elab/PreDefinition/WF/Fix.c @@ -150,7 +150,7 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_PreDefi static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_PreDefinition_WF_Fix_0__Lean_Elab_WF_replaceRecApps_loop___spec__5___lambda__2___closed__6; lean_object* l_Lean_Elab_pushInfoTree___at_Lean_Elab_Term_addDotCompletionInfo___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___at___private_Lean_Elab_PreDefinition_WF_Fix_0__Lean_Elab_WF_processSumCasesOn___spec__1___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* l_Lean_Elab_Term_reportUnsolvedGoals(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Elab_Term_reportUnsolvedGoals(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_WF_Fix_0__Lean_Elab_WF_replaceRecApps(lean_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_WF_mkFix___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*); static lean_object* l___private_Lean_Elab_PreDefinition_WF_Fix_0__Lean_Elab_WF_processPSigmaCasesOn___lambda__2___closed__2; @@ -301,7 +301,7 @@ lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_14 = l_Lean_Elab_Term_reportUnsolvedGoals(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_14 = l_Lean_Elab_Term_reportUnsolvedGoals(x_1, x_4, x_5, x_6, x_7, x_8); if (lean_obj_tag(x_14) == 0) { lean_object* x_15; @@ -356,7 +356,7 @@ lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_25 = l_Lean_Elab_Term_reportUnsolvedGoals(x_24, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_25 = l_Lean_Elab_Term_reportUnsolvedGoals(x_24, x_4, x_5, x_6, x_7, x_8); if (lean_obj_tag(x_25) == 0) { lean_object* x_26; diff --git a/stage0/stdlib/Lean/Elab/PreDefinition/WF/PackDomain.c b/stage0/stdlib/Lean/Elab/PreDefinition/WF/PackDomain.c index 12df15288c7a..0b4c9555abcd 100644 --- a/stage0/stdlib/Lean/Elab/PreDefinition/WF/PackDomain.c +++ b/stage0/stdlib/Lean/Elab/PreDefinition/WF/PackDomain.c @@ -106,7 +106,6 @@ lean_object* lean_st_ref_get(lean_object*, lean_object*); lean_object* lean_array_pop(lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_mkUnaryArg_go___closed__8; -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_WF_packDomain_packApplications___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_WF_packDomain___spec__4___closed__9; lean_object* l_Lean_Name_append(lean_object*, lean_object*); @@ -168,6 +167,7 @@ uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_packDomain___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*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Elab_WF_packDomain_packApplications_visit___spec__4___rarg(lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_packDomain___spec__4___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_packDomain___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* l_Lean_Name_mkStr2(lean_object*, lean_object*); @@ -1066,7 +1066,7 @@ lean_object* x_18; lean_object* x_19; x_18 = lean_ctor_get(x_17, 1); lean_inc(x_18); lean_dec(x_17); -x_19 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_12, x_5, x_6, x_7, x_8, x_18); +x_19 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_12, x_5, x_6, x_7, x_8, x_18); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); diff --git a/stage0/stdlib/Lean/Elab/PreDefinition/WF/PackMutual.c b/stage0/stdlib/Lean/Elab/PreDefinition/WF/PackMutual.c index dff65d1283a1..0aad9a9c3d13 100644 --- a/stage0/stdlib/Lean/Elab/PreDefinition/WF/PackMutual.c +++ b/stage0/stdlib/Lean/Elab/PreDefinition/WF/PackMutual.c @@ -89,7 +89,6 @@ static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_PreDefinit lean_object* lean_array_pop(lean_object*); lean_object* l_Lean_Meta_mkAppM(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_getCodomainLevel___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(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_forInUnsafe_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_mkNewDomain___spec__1___closed__2; extern lean_object* l_Lean_levelZero; @@ -144,6 +143,7 @@ extern lean_object* l_Lean_pp_sanitizeNames; static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_packValues___spec__1___closed__2; static lean_object* l___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_mkNewCoDomain_go___closed__5; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_getCodomainsLevel___spec__1___lambda__2___closed__3; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_withFixedPrefix___spec__1(size_t, size_t, lean_object*); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); @@ -2798,7 +2798,7 @@ lean_object* x_21; lean_object* x_22; x_21 = lean_ctor_get(x_20, 1); lean_inc(x_21); lean_dec(x_20); -x_22 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_15, x_4, x_5, x_6, x_7, x_21); +x_22 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_15, x_4, x_5, x_6, x_7, x_21); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); diff --git a/stage0/stdlib/Lean/Elab/Quotation.c b/stage0/stdlib/Lean/Elab/Quotation.c index 6eb934bdc4e9..34fc161b2e1a 100644 --- a/stage0/stdlib/Lean/Elab/Quotation.c +++ b/stage0/stdlib/Lean/Elab/Quotation.c @@ -20,7 +20,6 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotati 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__23___lambda__3___closed__28; LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__19(lean_object*, lean_object*); -uint8_t l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__27; 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__10___closed__20; @@ -55,7 +54,6 @@ LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Term_Quotation_ma LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__5(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_ArrayStxBuilder_mkNode___closed__45; 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*, 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___hyg_20075_(lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__1; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__2; static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__12___closed__2; @@ -86,7 +84,6 @@ lean_object* l_Lean_RBNode_find___at_Lean_sanitizeName___spec__1(lean_object*, l static lean_object* l_Lean_Elab_Term_Quotation_mkTuple___closed__2; 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 uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_1116____at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__3(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__9; static lean_object* l_Lean_Elab_Term_Quotation_mkTuple___closed__5; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__12; static lean_object* l_Lean_Elab_Term_Quotation_mkSyntaxQuotation___closed__32; @@ -117,7 +114,6 @@ static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_ static lean_object* l_Lean_Elab_Term_Quotation_mkSyntaxQuotation___closed__21; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__3; static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__11___closed__5; -static lean_object* l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__3; 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*); static lean_object* l_Lean_Elab_Term_Quotation_mkSyntaxQuotation___closed__11; static lean_object* l_Lean_Elab_Term_Quotation_mkSyntaxQuotation___closed__2; @@ -248,7 +244,6 @@ static lean_object* l_List_forIn_loop___at___private_Lean_Elab_Quotation_0__Lean static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_8015____closed__16; LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__27(lean_object*, lean_object*); 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_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__1; static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__33; 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*); 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*); @@ -261,7 +256,6 @@ static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_ static lean_object* l_Lean_Elab_Term_Quotation_instQuotePreresolvedMkStr1___closed__10; static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabNoErrorIfUnused_declRange___closed__4; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23(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_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__2; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__27___closed__5; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_Quotation_match__syntax_expand___spec__4(size_t, size_t, lean_object*); @@ -439,7 +433,6 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean 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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__37; uint8_t l_List_foldr___at_List_and___spec__1(uint8_t, lean_object*); -static lean_object* l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__7; static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__1___rarg___closed__1; 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_Lean_Elab_Term_Quotation_mkSyntaxQuotation___closed__40; @@ -526,6 +519,7 @@ static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation_____ 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__23___closed__20; 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_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__22___boxed(lean_object**); static lean_object* l_Lean_Elab_Term_Quotation_instReprMatchResult___closed__3; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__8; @@ -545,6 +539,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quota static lean_object* l_Lean_Elab_Term_Quotation_mkSyntaxQuotation___closed__22; 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_object* l_Lean_mkCIdentFrom(lean_object*, lean_object*, uint8_t); +static lean_object* l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__7; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteArray_go___at_Lean_Elab_Term_Quotation_ArrayStxBuilder_build___spec__3___boxed(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_object* l_Lean_Elab_getBetterRef(lean_object*, lean_object*); @@ -557,7 +552,6 @@ static lean_object* l_Lean_Elab_Term_Quotation_instQuotePreresolvedMkStr1___clos static lean_object* l_Lean_Elab_Term_Quotation_ArrayStxBuilder_mkNode___closed__82; static lean_object* l_Lean_Elab_Term_Quotation_instReprMatchResult___closed__8; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__14___closed__1; -static lean_object* l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__8; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___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_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_deduplicate___lambda__3___boxed(lean_object*, lean_object*); lean_object* l_ReaderT_instMonadLiftReaderT(lean_object*, lean_object*, lean_object*); @@ -593,6 +587,7 @@ static lean_object* l_Lean_Elab_Term_Quotation_mkTuple___closed__4; 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__23___closed__60; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepFromSplice___closed__2; +static lean_object* l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__9; static lean_object* l_Lean_Elab_Term_Quotation_ArrayStxBuilder_mkNode___closed__48; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__3; uint8_t l_List_isEmpty___rarg(lean_object*); @@ -692,7 +687,6 @@ static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_ lean_object* l_List_lengthTRAux___rarg(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__9; static lean_object* l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__6; -static lean_object* l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__6; static lean_object* l_Lean_Elab_Term_Quotation_ArrayStxBuilder_mkNode___closed__73; lean_object* l_Lean_addMacroScope(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__23___closed__69; @@ -731,6 +725,7 @@ static lean_object* l_Lean_Elab_Term_Quotation_ArrayStxBuilder_mkNode___closed__ static lean_object* l_List_forIn_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__6___closed__2; static lean_object* l_Lean_Elab_Term_Quotation_ArrayStxBuilder_mkNode___closed__29; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_checkUnusedAlts(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_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); static lean_object* l___private_Init_Meta_0__Lean_quoteArray___at_Lean_Elab_Term_Quotation_ArrayStxBuilder_build___spec__1___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot_declRange____x40_Lean_Elab_Quotation___hyg_8015____closed__1; static lean_object* l_Lean_Elab_Term_Quotation_instQuotePreresolvedMkStr1___closed__15; @@ -841,6 +836,7 @@ static lean_object* l_Lean_Elab_Term_Quotation_ArrayStxBuilder_mkNode___closed__ static lean_object* l_Lean_Elab_Term_Quotation_instReprMatchResult___closed__4; static lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__4___closed__2; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__13; +static lean_object* l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__5; static lean_object* l_List_forIn_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__6___closed__7; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___closed__12; LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___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_object*, lean_object*); @@ -872,6 +868,7 @@ static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_ 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*); lean_object* l_Lean_Elab_Term_addTermInfo_x27(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___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__3; +static lean_object* l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__3; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__17; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__2___closed__9; static lean_object* l___private_Init_Meta_0__Lean_quoteList___at_Lean_Elab_Term_Quotation_ArrayStxBuilder_build___spec__2___closed__6; @@ -908,6 +905,7 @@ static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_ static lean_object* l___private_Init_Meta_0__Lean_quoteList___at_Lean_Elab_Term_Quotation_ArrayStxBuilder_build___spec__2___closed__4; 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_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__27; +LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064_(lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__18; LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_findUsedAlts_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*); static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot_declRange____x40_Lean_Elab_Quotation___hyg_8107____closed__3; @@ -997,7 +995,6 @@ static lean_object* l_Lean_Elab_Term_Quotation_elabMatchSyntax___closed__1; lean_object* l_Lean_Syntax_node1(lean_object*, lean_object*, 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__27___closed__4; -static lean_object* l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__10; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_ArrayStxBuilder_empty; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___closed__12; static lean_object* l_Lean_Elab_Term_Quotation___aux__Lean__Elab__Quotation______macroRules__Lean__Elab__Term__Quotation__commandElab__stx__quot____1___closed__20; @@ -1075,6 +1072,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quota static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_findUsedAlts___closed__1; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__9; 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_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__6; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___closed__9; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___closed__10; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__58; @@ -1158,6 +1156,7 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Quotati static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__5; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___closed__15; static lean_object* l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__5; +static lean_object* l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__4; lean_object* lean_array_uget(lean_object*, size_t); static lean_object* l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__5; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabMatchSyntax_declRange(lean_object*); @@ -1227,7 +1226,6 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quota static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__1___closed__19; static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabMatchSyntax_declRange___closed__6; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__10___closed__5; -static lean_object* l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__4; static lean_object* l_Lean_Elab_Term_Quotation_mkSyntaxQuotation___closed__13; LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__13___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__14(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_mkSyntaxQuotation___closed__7; @@ -1248,6 +1246,7 @@ static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__8; lean_object* l_Array_mkArray2___rarg(lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__36; static lean_object* l_Lean_Elab_Term_Quotation_ArrayStxBuilder_mkNode___closed__31; +static lean_object* l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__2; extern lean_object* l_Lean_Elab_Term_Quotation_hygiene; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___lambda__2___closed__10; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__46; @@ -1259,7 +1258,6 @@ lean_object* lean_nat_add(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__13; lean_object* l_Lean_Syntax_antiquotKinds(lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_mkSyntaxQuotation___closed__31; -static lean_object* l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__5; static lean_object* l_List_foldlM___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__17___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_8061_(lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot_declRange____x40_Lean_Elab_Quotation___hyg_8084_(lean_object*); @@ -1280,6 +1278,7 @@ static lean_object* l_Lean_Elab_Term_Quotation_mkSyntaxQuotation___closed__23; 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* l_Lean_Meta_mkConstWithFreshMVarLevels(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___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__10; LEAN_EXPORT lean_object* l_List_foldl___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_markRhss___spec__6___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_markRhss___spec__7(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_Quotation_getQuotKind___closed__14; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_instReprMatchResult___boxed(lean_object*, lean_object*); @@ -1289,6 +1288,7 @@ static lean_object* l_Lean_Elab_Term_Quotation_ArrayStxBuilder_mkNode___closed__ static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__25; 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*); static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_elabMatchSyntax_declRange___closed__2; +static lean_object* l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__8; static lean_object* l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___closed__5; lean_object* l_String_toSubstring_x27(lean_object*); uint8_t l_Lean_Syntax_isAntiquots(lean_object*); @@ -14049,7 +14049,7 @@ lean_dec(x_79); x_80 = lean_ctor_get(x_6, 2); lean_inc(x_80); x_81 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__1; -x_82 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_80, x_81); +x_82 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_80, x_81); lean_dec(x_80); if (x_82 == 0) { @@ -14284,7 +14284,7 @@ lean_dec(x_1); x_181 = lean_ctor_get(x_6, 2); lean_inc(x_181); x_182 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___closed__1; -x_183 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_181, x_182); +x_183 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_181, x_182); lean_dec(x_181); if (x_183 == 0) { @@ -19731,7 +19731,7 @@ return x_18; } else { -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; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_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; size_t x_65; size_t x_66; +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; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_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; size_t x_64; size_t x_65; x_19 = lean_array_uget(x_6, x_8); x_20 = lean_ctor_get(x_14, 5); lean_inc(x_20); @@ -19761,83 +19761,80 @@ x_32 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_32, 0, x_22); lean_ctor_set(x_32, 1, x_4); lean_ctor_set(x_32, 2, x_31); -lean_inc(x_32); -lean_inc(x_19); -lean_inc(x_4); +x_33 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__13; lean_inc(x_22); -x_33 = l_Lean_Syntax_node2(x_22, x_4, x_19, x_32); -x_34 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__13; -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_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__10; -x_37 = l_Lean_addMacroScope(x_28, x_36, x_23); -x_38 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__8; +x_34 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_34, 0, x_22); +lean_ctor_set(x_34, 1, x_33); +x_35 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__10; +x_36 = l_Lean_addMacroScope(x_28, x_35, x_23); +x_37 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__8; lean_inc(x_3); lean_inc(x_22); -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_3); -x_40 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__11; +x_38 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_38, 0, x_22); +lean_ctor_set(x_38, 1, x_37); +lean_ctor_set(x_38, 2, x_36); +lean_ctor_set(x_38, 3, x_3); +x_39 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__11; lean_inc(x_22); -x_41 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_41, 0, x_22); -lean_ctor_set(x_41, 1, x_40); -x_42 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__12; +x_40 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_40, 0, x_22); +lean_ctor_set(x_40, 1, x_39); +x_41 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__12; lean_inc(x_22); -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_42 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_42, 0, x_22); +lean_ctor_set(x_42, 1, x_41); lean_inc(x_5); lean_inc(x_4); lean_inc(x_22); -x_44 = l_Lean_Syntax_node1(x_22, x_4, x_5); -x_45 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__16; +x_43 = l_Lean_Syntax_node1(x_22, x_4, x_5); +x_44 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__16; lean_inc(x_22); -x_46 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_46, 0, x_22); -lean_ctor_set(x_46, 1, x_45); -x_47 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__15; +x_45 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_45, 0, x_22); +lean_ctor_set(x_45, 1, x_44); +x_46 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__15; +lean_inc(x_19); lean_inc(x_32); lean_inc(x_22); -x_48 = l_Lean_Syntax_node4(x_22, x_47, x_44, x_32, x_46, x_19); -x_49 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__13; +x_47 = l_Lean_Syntax_node4(x_22, x_46, x_43, x_32, x_45, x_19); +x_48 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__13; lean_inc(x_22); -x_50 = l_Lean_Syntax_node2(x_22, x_49, x_43, x_48); -x_51 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__17; +x_49 = l_Lean_Syntax_node2(x_22, x_48, x_42, x_47); +x_50 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__17; 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_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__10; +x_51 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_51, 0, x_22); +lean_ctor_set(x_51, 1, x_50); +x_52 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__10; lean_inc(x_22); -x_54 = l_Lean_Syntax_node3(x_22, x_53, x_41, x_50, x_52); +x_53 = l_Lean_Syntax_node3(x_22, x_52, x_40, x_49, x_51); lean_inc(x_4); lean_inc(x_22); -x_55 = l_Lean_Syntax_node1(x_22, x_4, x_54); +x_54 = l_Lean_Syntax_node1(x_22, x_4, x_53); lean_inc(x_2); lean_inc(x_22); -x_56 = l_Lean_Syntax_node2(x_22, x_2, x_39, x_55); -x_57 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__6; +x_55 = l_Lean_Syntax_node2(x_22, x_2, x_38, x_54); +x_56 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__6; +lean_inc(x_32); lean_inc(x_22); -x_58 = l_Lean_Syntax_node4(x_22, x_57, x_33, x_32, x_35, x_56); -x_59 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__4; +x_57 = l_Lean_Syntax_node5(x_22, x_56, x_19, x_32, x_32, x_34, x_55); +x_58 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__4; lean_inc(x_22); -x_60 = l_Lean_Syntax_node1(x_22, x_59, x_58); -x_61 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_59 = l_Lean_Syntax_node1(x_22, x_58, x_57); +x_60 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; lean_inc(x_22); -x_62 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_62, 0, x_22); -lean_ctor_set(x_62, 1, x_61); -x_63 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__2; -x_64 = l_Lean_Syntax_node4(x_22, x_63, x_30, x_60, x_62, x_9); -x_65 = 1; -x_66 = lean_usize_add(x_8, x_65); -x_8 = x_66; -x_9 = x_64; +x_61 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_61, 0, x_22); +lean_ctor_set(x_61, 1, x_60); +x_62 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__2; +x_63 = l_Lean_Syntax_node4(x_22, x_62, x_30, x_59, x_61, x_9); +x_64 = 1; +x_65 = lean_usize_add(x_8, x_64); +x_8 = x_65; +x_9 = x_63; x_16 = x_26; goto _start; } @@ -19863,7 +19860,7 @@ return x_18; } else { -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; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_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; size_t x_65; size_t x_66; +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; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_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; size_t x_64; size_t x_65; x_19 = lean_array_uget(x_6, x_8); x_20 = lean_ctor_get(x_14, 5); lean_inc(x_20); @@ -19893,83 +19890,80 @@ x_32 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_32, 0, x_22); lean_ctor_set(x_32, 1, x_4); lean_ctor_set(x_32, 2, x_31); -lean_inc(x_32); -lean_inc(x_19); -lean_inc(x_4); -lean_inc(x_22); -x_33 = l_Lean_Syntax_node2(x_22, x_4, x_19, x_32); -x_34 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__13; +x_33 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__13; 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_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__10; -x_37 = l_Lean_addMacroScope(x_28, x_36, x_23); -x_38 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__8; +x_34 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_34, 0, x_22); +lean_ctor_set(x_34, 1, x_33); +x_35 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__10; +x_36 = l_Lean_addMacroScope(x_28, x_35, x_23); +x_37 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__8; lean_inc(x_3); lean_inc(x_22); -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_3); -x_40 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__11; +x_38 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_38, 0, x_22); +lean_ctor_set(x_38, 1, x_37); +lean_ctor_set(x_38, 2, x_36); +lean_ctor_set(x_38, 3, x_3); +x_39 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__11; lean_inc(x_22); -x_41 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_41, 0, x_22); -lean_ctor_set(x_41, 1, x_40); -x_42 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__12; +x_40 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_40, 0, x_22); +lean_ctor_set(x_40, 1, x_39); +x_41 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__12; lean_inc(x_22); -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_42 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_42, 0, x_22); +lean_ctor_set(x_42, 1, x_41); lean_inc(x_5); lean_inc(x_4); lean_inc(x_22); -x_44 = l_Lean_Syntax_node1(x_22, x_4, x_5); -x_45 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__16; +x_43 = l_Lean_Syntax_node1(x_22, x_4, x_5); +x_44 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__16; lean_inc(x_22); -x_46 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_46, 0, x_22); -lean_ctor_set(x_46, 1, x_45); -x_47 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__15; +x_45 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_45, 0, x_22); +lean_ctor_set(x_45, 1, x_44); +x_46 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__15; +lean_inc(x_19); lean_inc(x_32); lean_inc(x_22); -x_48 = l_Lean_Syntax_node4(x_22, x_47, x_44, x_32, x_46, x_19); -x_49 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__13; +x_47 = l_Lean_Syntax_node4(x_22, x_46, x_43, x_32, x_45, x_19); +x_48 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__13; lean_inc(x_22); -x_50 = l_Lean_Syntax_node2(x_22, x_49, x_43, x_48); -x_51 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__17; +x_49 = l_Lean_Syntax_node2(x_22, x_48, x_42, x_47); +x_50 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__17; 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_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__10; +x_51 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_51, 0, x_22); +lean_ctor_set(x_51, 1, x_50); +x_52 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__10; lean_inc(x_22); -x_54 = l_Lean_Syntax_node3(x_22, x_53, x_41, x_50, x_52); +x_53 = l_Lean_Syntax_node3(x_22, x_52, x_40, x_49, x_51); lean_inc(x_4); lean_inc(x_22); -x_55 = l_Lean_Syntax_node1(x_22, x_4, x_54); +x_54 = l_Lean_Syntax_node1(x_22, x_4, x_53); lean_inc(x_2); lean_inc(x_22); -x_56 = l_Lean_Syntax_node2(x_22, x_2, x_39, x_55); -x_57 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__6; +x_55 = l_Lean_Syntax_node2(x_22, x_2, x_38, x_54); +x_56 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__6; +lean_inc(x_32); lean_inc(x_22); -x_58 = l_Lean_Syntax_node4(x_22, x_57, x_33, x_32, x_35, x_56); -x_59 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__4; +x_57 = l_Lean_Syntax_node5(x_22, x_56, x_19, x_32, x_32, x_34, x_55); +x_58 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__4; lean_inc(x_22); -x_60 = l_Lean_Syntax_node1(x_22, x_59, x_58); -x_61 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_59 = l_Lean_Syntax_node1(x_22, x_58, x_57); +x_60 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; lean_inc(x_22); -x_62 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_62, 0, x_22); -lean_ctor_set(x_62, 1, x_61); -x_63 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__2; -x_64 = l_Lean_Syntax_node4(x_22, x_63, x_30, x_60, x_62, x_9); -x_65 = 1; -x_66 = lean_usize_add(x_8, x_65); -x_8 = x_66; -x_9 = x_64; +x_61 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_61, 0, x_22); +lean_ctor_set(x_61, 1, x_60); +x_62 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__2; +x_63 = l_Lean_Syntax_node4(x_22, x_62, x_30, x_59, x_61, x_9); +x_64 = 1; +x_65 = lean_usize_add(x_8, x_64); +x_8 = x_65; +x_9 = x_63; x_16 = x_26; goto _start; } @@ -19995,7 +19989,7 @@ return x_18; } else { -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; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_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; size_t x_65; size_t x_66; +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; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_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; size_t x_64; size_t x_65; x_19 = lean_array_uget(x_6, x_8); x_20 = lean_ctor_get(x_14, 5); lean_inc(x_20); @@ -20025,83 +20019,80 @@ x_32 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_32, 0, x_22); lean_ctor_set(x_32, 1, x_4); lean_ctor_set(x_32, 2, x_31); -lean_inc(x_32); -lean_inc(x_19); -lean_inc(x_4); -lean_inc(x_22); -x_33 = l_Lean_Syntax_node2(x_22, x_4, x_19, x_32); -x_34 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__13; +x_33 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__13; 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_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__10; -x_37 = l_Lean_addMacroScope(x_28, x_36, x_23); -x_38 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__8; +x_34 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_34, 0, x_22); +lean_ctor_set(x_34, 1, x_33); +x_35 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__10; +x_36 = l_Lean_addMacroScope(x_28, x_35, x_23); +x_37 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__8; lean_inc(x_3); lean_inc(x_22); -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_3); -x_40 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__11; +x_38 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_38, 0, x_22); +lean_ctor_set(x_38, 1, x_37); +lean_ctor_set(x_38, 2, x_36); +lean_ctor_set(x_38, 3, x_3); +x_39 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__11; lean_inc(x_22); -x_41 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_41, 0, x_22); -lean_ctor_set(x_41, 1, x_40); -x_42 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__12; +x_40 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_40, 0, x_22); +lean_ctor_set(x_40, 1, x_39); +x_41 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__12; lean_inc(x_22); -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_42 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_42, 0, x_22); +lean_ctor_set(x_42, 1, x_41); lean_inc(x_5); lean_inc(x_4); lean_inc(x_22); -x_44 = l_Lean_Syntax_node1(x_22, x_4, x_5); -x_45 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__16; +x_43 = l_Lean_Syntax_node1(x_22, x_4, x_5); +x_44 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__16; lean_inc(x_22); -x_46 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_46, 0, x_22); -lean_ctor_set(x_46, 1, x_45); -x_47 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__15; +x_45 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_45, 0, x_22); +lean_ctor_set(x_45, 1, x_44); +x_46 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__15; +lean_inc(x_19); lean_inc(x_32); lean_inc(x_22); -x_48 = l_Lean_Syntax_node4(x_22, x_47, x_44, x_32, x_46, x_19); -x_49 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__13; +x_47 = l_Lean_Syntax_node4(x_22, x_46, x_43, x_32, x_45, x_19); +x_48 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__13; lean_inc(x_22); -x_50 = l_Lean_Syntax_node2(x_22, x_49, x_43, x_48); -x_51 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__17; +x_49 = l_Lean_Syntax_node2(x_22, x_48, x_42, x_47); +x_50 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__17; 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_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__10; +x_51 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_51, 0, x_22); +lean_ctor_set(x_51, 1, x_50); +x_52 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__10; lean_inc(x_22); -x_54 = l_Lean_Syntax_node3(x_22, x_53, x_41, x_50, x_52); +x_53 = l_Lean_Syntax_node3(x_22, x_52, x_40, x_49, x_51); lean_inc(x_4); lean_inc(x_22); -x_55 = l_Lean_Syntax_node1(x_22, x_4, x_54); +x_54 = l_Lean_Syntax_node1(x_22, x_4, x_53); lean_inc(x_2); lean_inc(x_22); -x_56 = l_Lean_Syntax_node2(x_22, x_2, x_39, x_55); -x_57 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__6; +x_55 = l_Lean_Syntax_node2(x_22, x_2, x_38, x_54); +x_56 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__6; +lean_inc(x_32); lean_inc(x_22); -x_58 = l_Lean_Syntax_node4(x_22, x_57, x_33, x_32, x_35, x_56); -x_59 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__4; +x_57 = l_Lean_Syntax_node5(x_22, x_56, x_19, x_32, x_32, x_34, x_55); +x_58 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__4; lean_inc(x_22); -x_60 = l_Lean_Syntax_node1(x_22, x_59, x_58); -x_61 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_59 = l_Lean_Syntax_node1(x_22, x_58, x_57); +x_60 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; lean_inc(x_22); -x_62 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_62, 0, x_22); -lean_ctor_set(x_62, 1, x_61); -x_63 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__2; -x_64 = l_Lean_Syntax_node4(x_22, x_63, x_30, x_60, x_62, x_9); -x_65 = 1; -x_66 = lean_usize_add(x_8, x_65); -x_8 = x_66; -x_9 = x_64; +x_61 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_61, 0, x_22); +lean_ctor_set(x_61, 1, x_60); +x_62 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__2; +x_63 = l_Lean_Syntax_node4(x_22, x_62, x_30, x_59, x_61, x_9); +x_64 = 1; +x_65 = lean_usize_add(x_8, x_64); +x_8 = x_65; +x_9 = x_63; x_16 = x_26; goto _start; } @@ -20215,7 +20206,7 @@ return x_18; } else { -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; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_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; size_t x_65; size_t x_66; +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; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_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; size_t x_64; size_t x_65; x_19 = lean_array_uget(x_6, x_8); x_20 = lean_ctor_get(x_14, 5); lean_inc(x_20); @@ -20245,83 +20236,80 @@ x_32 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_32, 0, x_22); lean_ctor_set(x_32, 1, x_4); lean_ctor_set(x_32, 2, x_31); -lean_inc(x_32); -lean_inc(x_19); -lean_inc(x_4); +x_33 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__13; lean_inc(x_22); -x_33 = l_Lean_Syntax_node2(x_22, x_4, x_19, x_32); -x_34 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__13; -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_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__10; -x_37 = l_Lean_addMacroScope(x_28, x_36, x_23); -x_38 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__8; +x_34 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_34, 0, x_22); +lean_ctor_set(x_34, 1, x_33); +x_35 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__10; +x_36 = l_Lean_addMacroScope(x_28, x_35, x_23); +x_37 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__8; lean_inc(x_3); lean_inc(x_22); -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_3); -x_40 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__11; +x_38 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_38, 0, x_22); +lean_ctor_set(x_38, 1, x_37); +lean_ctor_set(x_38, 2, x_36); +lean_ctor_set(x_38, 3, x_3); +x_39 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__11; lean_inc(x_22); -x_41 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_41, 0, x_22); -lean_ctor_set(x_41, 1, x_40); -x_42 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__12; +x_40 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_40, 0, x_22); +lean_ctor_set(x_40, 1, x_39); +x_41 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__12; lean_inc(x_22); -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_42 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_42, 0, x_22); +lean_ctor_set(x_42, 1, x_41); lean_inc(x_5); lean_inc(x_4); lean_inc(x_22); -x_44 = l_Lean_Syntax_node1(x_22, x_4, x_5); -x_45 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__16; +x_43 = l_Lean_Syntax_node1(x_22, x_4, x_5); +x_44 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__16; lean_inc(x_22); -x_46 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_46, 0, x_22); -lean_ctor_set(x_46, 1, x_45); -x_47 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__15; +x_45 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_45, 0, x_22); +lean_ctor_set(x_45, 1, x_44); +x_46 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__15; +lean_inc(x_19); lean_inc(x_32); lean_inc(x_22); -x_48 = l_Lean_Syntax_node4(x_22, x_47, x_44, x_32, x_46, x_19); -x_49 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__13; +x_47 = l_Lean_Syntax_node4(x_22, x_46, x_43, x_32, x_45, x_19); +x_48 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__13; lean_inc(x_22); -x_50 = l_Lean_Syntax_node2(x_22, x_49, x_43, x_48); -x_51 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__17; +x_49 = l_Lean_Syntax_node2(x_22, x_48, x_42, x_47); +x_50 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__17; 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_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__10; +x_51 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_51, 0, x_22); +lean_ctor_set(x_51, 1, x_50); +x_52 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__10; lean_inc(x_22); -x_54 = l_Lean_Syntax_node3(x_22, x_53, x_41, x_50, x_52); +x_53 = l_Lean_Syntax_node3(x_22, x_52, x_40, x_49, x_51); lean_inc(x_4); lean_inc(x_22); -x_55 = l_Lean_Syntax_node1(x_22, x_4, x_54); +x_54 = l_Lean_Syntax_node1(x_22, x_4, x_53); lean_inc(x_2); lean_inc(x_22); -x_56 = l_Lean_Syntax_node2(x_22, x_2, x_39, x_55); -x_57 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__6; +x_55 = l_Lean_Syntax_node2(x_22, x_2, x_38, x_54); +x_56 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__6; +lean_inc(x_32); lean_inc(x_22); -x_58 = l_Lean_Syntax_node4(x_22, x_57, x_33, x_32, x_35, x_56); -x_59 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__4; +x_57 = l_Lean_Syntax_node5(x_22, x_56, x_19, x_32, x_32, x_34, x_55); +x_58 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__4; lean_inc(x_22); -x_60 = l_Lean_Syntax_node1(x_22, x_59, x_58); -x_61 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_59 = l_Lean_Syntax_node1(x_22, x_58, x_57); +x_60 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; lean_inc(x_22); -x_62 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_62, 0, x_22); -lean_ctor_set(x_62, 1, x_61); -x_63 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__2; -x_64 = l_Lean_Syntax_node4(x_22, x_63, x_30, x_60, x_62, x_9); -x_65 = 1; -x_66 = lean_usize_add(x_8, x_65); -x_8 = x_66; -x_9 = x_64; +x_61 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_61, 0, x_22); +lean_ctor_set(x_61, 1, x_60); +x_62 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__2; +x_63 = l_Lean_Syntax_node4(x_22, x_62, x_30, x_59, x_61, x_9); +x_64 = 1; +x_65 = lean_usize_add(x_8, x_64); +x_8 = x_65; +x_9 = x_63; x_16 = x_26; goto _start; } @@ -20347,7 +20335,7 @@ return x_18; } else { -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; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_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; size_t x_65; size_t x_66; +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; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_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; size_t x_64; size_t x_65; x_19 = lean_array_uget(x_6, x_8); x_20 = lean_ctor_get(x_14, 5); lean_inc(x_20); @@ -20377,83 +20365,80 @@ x_32 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_32, 0, x_22); lean_ctor_set(x_32, 1, x_4); lean_ctor_set(x_32, 2, x_31); -lean_inc(x_32); -lean_inc(x_19); -lean_inc(x_4); +x_33 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__13; lean_inc(x_22); -x_33 = l_Lean_Syntax_node2(x_22, x_4, x_19, x_32); -x_34 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__13; -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_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__10; -x_37 = l_Lean_addMacroScope(x_28, x_36, x_23); -x_38 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__8; +x_34 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_34, 0, x_22); +lean_ctor_set(x_34, 1, x_33); +x_35 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__10; +x_36 = l_Lean_addMacroScope(x_28, x_35, x_23); +x_37 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__8; lean_inc(x_3); lean_inc(x_22); -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_3); -x_40 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__11; +x_38 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_38, 0, x_22); +lean_ctor_set(x_38, 1, x_37); +lean_ctor_set(x_38, 2, x_36); +lean_ctor_set(x_38, 3, x_3); +x_39 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__11; lean_inc(x_22); -x_41 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_41, 0, x_22); -lean_ctor_set(x_41, 1, x_40); -x_42 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__12; +x_40 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_40, 0, x_22); +lean_ctor_set(x_40, 1, x_39); +x_41 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__12; lean_inc(x_22); -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_42 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_42, 0, x_22); +lean_ctor_set(x_42, 1, x_41); lean_inc(x_5); lean_inc(x_4); lean_inc(x_22); -x_44 = l_Lean_Syntax_node1(x_22, x_4, x_5); -x_45 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__16; +x_43 = l_Lean_Syntax_node1(x_22, x_4, x_5); +x_44 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__16; lean_inc(x_22); -x_46 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_46, 0, x_22); -lean_ctor_set(x_46, 1, x_45); -x_47 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__15; +x_45 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_45, 0, x_22); +lean_ctor_set(x_45, 1, x_44); +x_46 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__15; +lean_inc(x_19); lean_inc(x_32); lean_inc(x_22); -x_48 = l_Lean_Syntax_node4(x_22, x_47, x_44, x_32, x_46, x_19); -x_49 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__13; +x_47 = l_Lean_Syntax_node4(x_22, x_46, x_43, x_32, x_45, x_19); +x_48 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__13; lean_inc(x_22); -x_50 = l_Lean_Syntax_node2(x_22, x_49, x_43, x_48); -x_51 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__17; +x_49 = l_Lean_Syntax_node2(x_22, x_48, x_42, x_47); +x_50 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__17; 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_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__10; +x_51 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_51, 0, x_22); +lean_ctor_set(x_51, 1, x_50); +x_52 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__10; lean_inc(x_22); -x_54 = l_Lean_Syntax_node3(x_22, x_53, x_41, x_50, x_52); +x_53 = l_Lean_Syntax_node3(x_22, x_52, x_40, x_49, x_51); lean_inc(x_4); lean_inc(x_22); -x_55 = l_Lean_Syntax_node1(x_22, x_4, x_54); +x_54 = l_Lean_Syntax_node1(x_22, x_4, x_53); lean_inc(x_2); lean_inc(x_22); -x_56 = l_Lean_Syntax_node2(x_22, x_2, x_39, x_55); -x_57 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__6; +x_55 = l_Lean_Syntax_node2(x_22, x_2, x_38, x_54); +x_56 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__6; +lean_inc(x_32); lean_inc(x_22); -x_58 = l_Lean_Syntax_node4(x_22, x_57, x_33, x_32, x_35, x_56); -x_59 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__4; +x_57 = l_Lean_Syntax_node5(x_22, x_56, x_19, x_32, x_32, x_34, x_55); +x_58 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__4; lean_inc(x_22); -x_60 = l_Lean_Syntax_node1(x_22, x_59, x_58); -x_61 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_59 = l_Lean_Syntax_node1(x_22, x_58, x_57); +x_60 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; lean_inc(x_22); -x_62 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_62, 0, x_22); -lean_ctor_set(x_62, 1, x_61); -x_63 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__2; -x_64 = l_Lean_Syntax_node4(x_22, x_63, x_30, x_60, x_62, x_9); -x_65 = 1; -x_66 = lean_usize_add(x_8, x_65); -x_8 = x_66; -x_9 = x_64; +x_61 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_61, 0, x_22); +lean_ctor_set(x_61, 1, x_60); +x_62 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__2; +x_63 = l_Lean_Syntax_node4(x_22, x_62, x_30, x_59, x_61, x_9); +x_64 = 1; +x_65 = lean_usize_add(x_8, x_64); +x_8 = x_65; +x_9 = x_63; x_16 = x_26; goto _start; } @@ -20479,7 +20464,7 @@ return x_18; } else { -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; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_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; size_t x_65; size_t x_66; +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; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_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; size_t x_64; size_t x_65; x_19 = lean_array_uget(x_6, x_8); x_20 = lean_ctor_get(x_14, 5); lean_inc(x_20); @@ -20509,83 +20494,80 @@ x_32 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_32, 0, x_22); lean_ctor_set(x_32, 1, x_4); lean_ctor_set(x_32, 2, x_31); -lean_inc(x_32); -lean_inc(x_19); -lean_inc(x_4); -lean_inc(x_22); -x_33 = l_Lean_Syntax_node2(x_22, x_4, x_19, x_32); -x_34 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__13; +x_33 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__13; 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_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__10; -x_37 = l_Lean_addMacroScope(x_28, x_36, x_23); -x_38 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__8; +x_34 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_34, 0, x_22); +lean_ctor_set(x_34, 1, x_33); +x_35 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__10; +x_36 = l_Lean_addMacroScope(x_28, x_35, x_23); +x_37 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__8; lean_inc(x_3); lean_inc(x_22); -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_3); -x_40 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__11; +x_38 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_38, 0, x_22); +lean_ctor_set(x_38, 1, x_37); +lean_ctor_set(x_38, 2, x_36); +lean_ctor_set(x_38, 3, x_3); +x_39 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__11; lean_inc(x_22); -x_41 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_41, 0, x_22); -lean_ctor_set(x_41, 1, x_40); -x_42 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__12; +x_40 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_40, 0, x_22); +lean_ctor_set(x_40, 1, x_39); +x_41 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__12; lean_inc(x_22); -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_42 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_42, 0, x_22); +lean_ctor_set(x_42, 1, x_41); lean_inc(x_5); lean_inc(x_4); lean_inc(x_22); -x_44 = l_Lean_Syntax_node1(x_22, x_4, x_5); -x_45 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__16; +x_43 = l_Lean_Syntax_node1(x_22, x_4, x_5); +x_44 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__16; lean_inc(x_22); -x_46 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_46, 0, x_22); -lean_ctor_set(x_46, 1, x_45); -x_47 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__15; +x_45 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_45, 0, x_22); +lean_ctor_set(x_45, 1, x_44); +x_46 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__15; +lean_inc(x_19); lean_inc(x_32); lean_inc(x_22); -x_48 = l_Lean_Syntax_node4(x_22, x_47, x_44, x_32, x_46, x_19); -x_49 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__13; +x_47 = l_Lean_Syntax_node4(x_22, x_46, x_43, x_32, x_45, x_19); +x_48 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__13; lean_inc(x_22); -x_50 = l_Lean_Syntax_node2(x_22, x_49, x_43, x_48); -x_51 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__17; +x_49 = l_Lean_Syntax_node2(x_22, x_48, x_42, x_47); +x_50 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__17; 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_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__10; +x_51 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_51, 0, x_22); +lean_ctor_set(x_51, 1, x_50); +x_52 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__10; lean_inc(x_22); -x_54 = l_Lean_Syntax_node3(x_22, x_53, x_41, x_50, x_52); +x_53 = l_Lean_Syntax_node3(x_22, x_52, x_40, x_49, x_51); lean_inc(x_4); lean_inc(x_22); -x_55 = l_Lean_Syntax_node1(x_22, x_4, x_54); +x_54 = l_Lean_Syntax_node1(x_22, x_4, x_53); lean_inc(x_2); lean_inc(x_22); -x_56 = l_Lean_Syntax_node2(x_22, x_2, x_39, x_55); -x_57 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__6; +x_55 = l_Lean_Syntax_node2(x_22, x_2, x_38, x_54); +x_56 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__6; +lean_inc(x_32); lean_inc(x_22); -x_58 = l_Lean_Syntax_node4(x_22, x_57, x_33, x_32, x_35, x_56); -x_59 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__4; +x_57 = l_Lean_Syntax_node5(x_22, x_56, x_19, x_32, x_32, x_34, x_55); +x_58 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__4; lean_inc(x_22); -x_60 = l_Lean_Syntax_node1(x_22, x_59, x_58); -x_61 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_59 = l_Lean_Syntax_node1(x_22, x_58, x_57); +x_60 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; lean_inc(x_22); -x_62 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_62, 0, x_22); -lean_ctor_set(x_62, 1, x_61); -x_63 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__2; -x_64 = l_Lean_Syntax_node4(x_22, x_63, x_30, x_60, x_62, x_9); -x_65 = 1; -x_66 = lean_usize_add(x_8, x_65); -x_8 = x_66; -x_9 = x_64; +x_61 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_61, 0, x_22); +lean_ctor_set(x_61, 1, x_60); +x_62 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__2; +x_63 = l_Lean_Syntax_node4(x_22, x_62, x_30, x_59, x_61, x_9); +x_64 = 1; +x_65 = lean_usize_add(x_8, x_64); +x_8 = x_65; +x_9 = x_63; x_16 = x_26; goto _start; } @@ -21024,7 +21006,7 @@ x_15 = lean_st_ref_get(x_9, x_10); 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; 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_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_17 = lean_ctor_get(x_15, 0); x_18 = lean_ctor_get(x_17, 0); lean_inc(x_18); @@ -21042,97 +21024,93 @@ x_24 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_24, 0, x_13); lean_ctor_set(x_24, 1, x_22); lean_ctor_set(x_24, 2, x_23); -lean_inc(x_24); -lean_inc(x_13); -x_25 = l_Lean_Syntax_node2(x_13, x_22, x_1, x_24); -x_26 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__13; +x_25 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__13; lean_inc(x_13); -x_27 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_27, 0, x_13); -lean_ctor_set(x_27, 1, x_26); -x_28 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; -x_29 = l_Lean_addMacroScope(x_19, x_28, x_14); -x_30 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; +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_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; +x_28 = l_Lean_addMacroScope(x_19, x_27, x_14); +x_29 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; lean_inc(x_13); -x_31 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_31, 0, x_13); -lean_ctor_set(x_31, 1, x_30); -lean_ctor_set(x_31, 2, x_29); -lean_ctor_set(x_31, 3, x_2); -x_32 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__6; +x_30 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_30, 0, x_13); +lean_ctor_set(x_30, 1, x_29); +lean_ctor_set(x_30, 2, x_28); +lean_ctor_set(x_30, 3, x_1); +x_31 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__6; +lean_inc(x_24); lean_inc(x_13); -x_33 = l_Lean_Syntax_node4(x_13, x_32, x_25, x_24, x_27, x_31); -x_34 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__4; +x_32 = l_Lean_Syntax_node5(x_13, x_31, x_2, x_24, x_24, x_26, x_30); +x_33 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__4; lean_inc(x_13); -x_35 = l_Lean_Syntax_node1(x_13, x_34, x_33); -x_36 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_34 = l_Lean_Syntax_node1(x_13, x_33, x_32); +x_35 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; lean_inc(x_13); -x_37 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_37, 0, x_13); -lean_ctor_set(x_37, 1, x_36); -x_38 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__2; -x_39 = l_Lean_Syntax_node4(x_13, x_38, x_21, x_35, x_37, x_3); -lean_ctor_set(x_15, 0, x_39); +x_36 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_36, 0, x_13); +lean_ctor_set(x_36, 1, x_35); +x_37 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__2; +x_38 = l_Lean_Syntax_node4(x_13, x_37, x_21, x_34, x_36, x_3); +lean_ctor_set(x_15, 0, x_38); return x_15; } 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; -x_40 = lean_ctor_get(x_15, 0); -x_41 = lean_ctor_get(x_15, 1); -lean_inc(x_41); +lean_object* x_39; lean_object* x_40; lean_object* x_41; 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_39 = lean_ctor_get(x_15, 0); +x_40 = lean_ctor_get(x_15, 1); lean_inc(x_40); +lean_inc(x_39); lean_dec(x_15); -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_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; -lean_inc(x_13); -x_45 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_45, 0, x_13); -lean_ctor_set(x_45, 1, x_44); -x_46 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__11; -x_47 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__12; +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_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; lean_inc(x_13); -x_48 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_48, 0, x_13); -lean_ctor_set(x_48, 1, x_46); -lean_ctor_set(x_48, 2, x_47); -lean_inc(x_48); +x_44 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_44, 0, x_13); +lean_ctor_set(x_44, 1, x_43); +x_45 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__11; +x_46 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__12; lean_inc(x_13); -x_49 = l_Lean_Syntax_node2(x_13, x_46, x_1, x_48); -x_50 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__13; +x_47 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_47, 0, x_13); +lean_ctor_set(x_47, 1, x_45); +lean_ctor_set(x_47, 2, x_46); +x_48 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__13; 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_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; -x_53 = l_Lean_addMacroScope(x_43, x_52, x_14); -x_54 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; +x_49 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_49, 0, x_13); +lean_ctor_set(x_49, 1, x_48); +x_50 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; +x_51 = l_Lean_addMacroScope(x_42, x_50, x_14); +x_52 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; lean_inc(x_13); -x_55 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_55, 0, x_13); -lean_ctor_set(x_55, 1, x_54); -lean_ctor_set(x_55, 2, x_53); -lean_ctor_set(x_55, 3, x_2); -x_56 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__6; +x_53 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_53, 0, x_13); +lean_ctor_set(x_53, 1, x_52); +lean_ctor_set(x_53, 2, x_51); +lean_ctor_set(x_53, 3, x_1); +x_54 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__6; +lean_inc(x_47); lean_inc(x_13); -x_57 = l_Lean_Syntax_node4(x_13, x_56, x_49, x_48, x_51, x_55); -x_58 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__4; +x_55 = l_Lean_Syntax_node5(x_13, x_54, x_2, x_47, x_47, x_49, x_53); +x_56 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__4; lean_inc(x_13); -x_59 = l_Lean_Syntax_node1(x_13, x_58, x_57); -x_60 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_57 = l_Lean_Syntax_node1(x_13, x_56, x_55); +x_58 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; lean_inc(x_13); -x_61 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_61, 0, x_13); -lean_ctor_set(x_61, 1, x_60); -x_62 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__2; -x_63 = l_Lean_Syntax_node4(x_13, x_62, x_45, x_59, x_61, x_3); -x_64 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_64, 0, x_63); -lean_ctor_set(x_64, 1, x_41); -return x_64; +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_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__2; +x_61 = l_Lean_Syntax_node4(x_13, x_60, x_44, x_57, x_59, x_3); +x_62 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_62, 0, x_61); +lean_ctor_set(x_62, 1, x_40); +return x_62; } } } @@ -26496,7 +26474,7 @@ lean_dec(x_12); x_41 = !lean_is_exclusive(x_40); if (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; 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_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; x_42 = lean_ctor_get(x_40, 0); x_43 = lean_ctor_get(x_42, 0); lean_inc(x_43); @@ -26514,269 +26492,265 @@ x_49 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_49, 0, x_38); lean_ctor_set(x_49, 1, x_47); lean_ctor_set(x_49, 2, x_48); -lean_inc(x_49); -lean_inc(x_38); -x_50 = l_Lean_Syntax_node2(x_38, x_47, x_3, x_49); -x_51 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__13; +x_50 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__13; lean_inc(x_38); -x_52 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_52, 0, x_38); -lean_ctor_set(x_52, 1, x_51); -x_53 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__21; +x_51 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_51, 0, x_38); +lean_ctor_set(x_51, 1, x_50); +x_52 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__21; lean_inc(x_38); -x_54 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_54, 0, x_38); -lean_ctor_set(x_54, 1, x_53); -x_55 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___closed__3; +x_53 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_53, 0, x_38); +lean_ctor_set(x_53, 1, x_52); +x_54 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___closed__3; lean_inc(x_39); lean_inc(x_44); -x_56 = l_Lean_addMacroScope(x_44, x_55, x_39); -x_57 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___closed__4; -lean_inc(x_4); -x_58 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_58, 0, x_57); -lean_ctor_set(x_58, 1, x_4); -x_59 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___closed__5; -lean_inc(x_4); +x_55 = l_Lean_addMacroScope(x_44, x_54, x_39); +x_56 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___closed__4; +lean_inc(x_3); +x_57 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_57, 0, x_56); +lean_ctor_set(x_57, 1, x_3); +x_58 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___closed__5; +lean_inc(x_3); +x_59 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_59, 0, x_58); +lean_ctor_set(x_59, 1, x_3); x_60 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_60, 0, x_59); -lean_ctor_set(x_60, 1, x_4); -x_61 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_61, 0, x_58); -lean_ctor_set(x_61, 1, x_60); -x_62 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___closed__2; +lean_ctor_set(x_60, 0, x_57); +lean_ctor_set(x_60, 1, x_59); +x_61 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___closed__2; lean_inc(x_38); -x_63 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_63, 0, x_38); -lean_ctor_set(x_63, 1, x_62); -lean_ctor_set(x_63, 2, x_56); -lean_ctor_set(x_63, 3, x_61); -x_64 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__20; +x_62 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_62, 0, x_38); +lean_ctor_set(x_62, 1, x_61); +lean_ctor_set(x_62, 2, x_55); +lean_ctor_set(x_62, 3, x_60); +x_63 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__20; lean_inc(x_38); -x_65 = l_Lean_Syntax_node2(x_38, x_64, x_54, x_63); -x_66 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__22(x_5); -x_67 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepFromSplice(x_1); -x_68 = lean_box(2); -x_69 = l_Lean_Syntax_mkStrLit(x_67, x_68); -lean_dec(x_67); -x_70 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__11; +x_64 = l_Lean_Syntax_node2(x_38, x_63, x_53, x_62); +x_65 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__22(x_4); +x_66 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepFromSplice(x_1); +x_67 = lean_box(2); +x_68 = l_Lean_Syntax_mkStrLit(x_66, x_67); +lean_dec(x_66); +x_69 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__11; lean_inc(x_38); -x_71 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_71, 0, x_38); -lean_ctor_set(x_71, 1, x_70); -x_72 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__3; +x_70 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_70, 0, x_38); +lean_ctor_set(x_70, 1, x_69); +x_71 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__3; lean_inc(x_39); lean_inc(x_44); -x_73 = l_Lean_addMacroScope(x_44, x_72, x_39); -x_74 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__4; -lean_inc(x_4); -x_75 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_75, 0, x_74); -lean_ctor_set(x_75, 1, x_4); -x_76 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__5; -lean_inc(x_4); +x_72 = l_Lean_addMacroScope(x_44, x_71, x_39); +x_73 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__4; +lean_inc(x_3); +x_74 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_74, 0, x_73); +lean_ctor_set(x_74, 1, x_3); +x_75 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__5; +lean_inc(x_3); +x_76 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_76, 0, x_75); +lean_ctor_set(x_76, 1, x_3); x_77 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_77, 0, x_76); -lean_ctor_set(x_77, 1, x_4); -x_78 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_78, 0, x_75); -lean_ctor_set(x_78, 1, x_77); -x_79 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__2; +lean_ctor_set(x_77, 0, x_74); +lean_ctor_set(x_77, 1, x_76); +x_78 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__2; lean_inc(x_38); -x_80 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_80, 0, x_38); -lean_ctor_set(x_80, 1, x_79); -lean_ctor_set(x_80, 2, x_73); -lean_ctor_set(x_80, 3, x_78); -x_81 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; -x_82 = l_Lean_addMacroScope(x_44, x_81, x_39); -x_83 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; +x_79 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_79, 0, x_38); +lean_ctor_set(x_79, 1, x_78); +lean_ctor_set(x_79, 2, x_72); +lean_ctor_set(x_79, 3, x_77); +x_80 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; +x_81 = l_Lean_addMacroScope(x_44, x_80, x_39); +x_82 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; lean_inc(x_38); -x_84 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_84, 0, x_38); -lean_ctor_set(x_84, 1, x_83); -lean_ctor_set(x_84, 2, x_82); -lean_ctor_set(x_84, 3, x_4); +x_83 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_83, 0, x_38); +lean_ctor_set(x_83, 1, x_82); +lean_ctor_set(x_83, 2, x_81); +lean_ctor_set(x_83, 3, x_3); lean_inc(x_38); -x_85 = l_Lean_Syntax_node1(x_38, x_47, x_84); -x_86 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___closed__6; +x_84 = l_Lean_Syntax_node1(x_38, x_47, x_83); +x_85 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___closed__6; lean_inc(x_38); -x_87 = l_Lean_Syntax_node2(x_38, x_86, x_80, x_85); -x_88 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__17; +x_86 = l_Lean_Syntax_node2(x_38, x_85, x_79, x_84); +x_87 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__17; lean_inc(x_38); -x_89 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_89, 0, x_38); -lean_ctor_set(x_89, 1, x_88); -x_90 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__10; +x_88 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_88, 0, x_38); +lean_ctor_set(x_88, 1, x_87); +x_89 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__10; lean_inc(x_38); -x_91 = l_Lean_Syntax_node3(x_38, x_90, x_71, x_87, x_89); +x_90 = l_Lean_Syntax_node3(x_38, x_89, x_70, x_86, x_88); lean_inc(x_38); -x_92 = l_Lean_Syntax_node3(x_38, x_47, x_66, x_69, x_91); +x_91 = l_Lean_Syntax_node3(x_38, x_47, x_65, x_68, x_90); lean_inc(x_38); -x_93 = l_Lean_Syntax_node2(x_38, x_86, x_65, x_92); -x_94 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__6; +x_92 = l_Lean_Syntax_node2(x_38, x_85, x_64, x_91); +x_93 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__6; +lean_inc(x_49); lean_inc(x_38); -x_95 = l_Lean_Syntax_node4(x_38, x_94, x_50, x_49, x_52, x_93); -x_96 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__4; +x_94 = l_Lean_Syntax_node5(x_38, x_93, x_5, x_49, x_49, x_51, x_92); +x_95 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__4; lean_inc(x_38); -x_97 = l_Lean_Syntax_node1(x_38, x_96, x_95); -x_98 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_96 = l_Lean_Syntax_node1(x_38, x_95, x_94); +x_97 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; lean_inc(x_38); -x_99 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_99, 0, x_38); -lean_ctor_set(x_99, 1, x_98); -x_100 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__2; -x_101 = l_Lean_Syntax_node4(x_38, x_100, x_46, x_97, x_99, x_6); -lean_ctor_set(x_40, 0, x_101); +x_98 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_98, 0, x_38); +lean_ctor_set(x_98, 1, x_97); +x_99 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__2; +x_100 = l_Lean_Syntax_node4(x_38, x_99, x_46, x_96, x_98, x_6); +lean_ctor_set(x_40, 0, x_100); return x_40; } else { -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_102 = lean_ctor_get(x_40, 0); -x_103 = lean_ctor_get(x_40, 1); -lean_inc(x_103); +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; +x_101 = lean_ctor_get(x_40, 0); +x_102 = lean_ctor_get(x_40, 1); lean_inc(x_102); +lean_inc(x_101); lean_dec(x_40); -x_104 = lean_ctor_get(x_102, 0); -lean_inc(x_104); -lean_dec(x_102); -x_105 = lean_environment_main_module(x_104); -x_106 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; +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_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; lean_inc(x_38); -x_107 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_107, 0, x_38); -lean_ctor_set(x_107, 1, x_106); -x_108 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__11; -x_109 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__12; +x_106 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_106, 0, x_38); +lean_ctor_set(x_106, 1, x_105); +x_107 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__11; +x_108 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__12; lean_inc(x_38); -x_110 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_110, 0, x_38); -lean_ctor_set(x_110, 1, x_108); -lean_ctor_set(x_110, 2, x_109); -lean_inc(x_110); +x_109 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_109, 0, x_38); +lean_ctor_set(x_109, 1, x_107); +lean_ctor_set(x_109, 2, x_108); +x_110 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__13; lean_inc(x_38); -x_111 = l_Lean_Syntax_node2(x_38, x_108, x_3, x_110); -x_112 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__13; +x_111 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_111, 0, x_38); +lean_ctor_set(x_111, 1, x_110); +x_112 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__21; lean_inc(x_38); x_113 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_113, 0, x_38); lean_ctor_set(x_113, 1, x_112); -x_114 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__21; -lean_inc(x_38); -x_115 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_115, 0, x_38); -lean_ctor_set(x_115, 1, x_114); -x_116 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___closed__3; +x_114 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___closed__3; lean_inc(x_39); -lean_inc(x_105); -x_117 = l_Lean_addMacroScope(x_105, x_116, x_39); -x_118 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___closed__4; -lean_inc(x_4); +lean_inc(x_104); +x_115 = l_Lean_addMacroScope(x_104, x_114, x_39); +x_116 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___closed__4; +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__25___closed__5; +lean_inc(x_3); x_119 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_119, 0, x_118); -lean_ctor_set(x_119, 1, x_4); -x_120 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___closed__5; -lean_inc(x_4); -x_121 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_121, 0, x_120); -lean_ctor_set(x_121, 1, x_4); -x_122 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_122, 0, x_119); -lean_ctor_set(x_122, 1, x_121); -x_123 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___closed__2; +lean_ctor_set(x_119, 1, x_3); +x_120 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_120, 0, x_117); +lean_ctor_set(x_120, 1, x_119); +x_121 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___closed__2; lean_inc(x_38); -x_124 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_124, 0, x_38); -lean_ctor_set(x_124, 1, x_123); -lean_ctor_set(x_124, 2, x_117); -lean_ctor_set(x_124, 3, x_122); -x_125 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__20; +x_122 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_122, 0, x_38); +lean_ctor_set(x_122, 1, x_121); +lean_ctor_set(x_122, 2, x_115); +lean_ctor_set(x_122, 3, x_120); +x_123 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__20; lean_inc(x_38); -x_126 = l_Lean_Syntax_node2(x_38, x_125, x_115, x_124); -x_127 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__22(x_5); -x_128 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepFromSplice(x_1); -x_129 = lean_box(2); -x_130 = l_Lean_Syntax_mkStrLit(x_128, x_129); -lean_dec(x_128); -x_131 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__11; +x_124 = l_Lean_Syntax_node2(x_38, x_123, x_113, x_122); +x_125 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__22(x_4); +x_126 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepFromSplice(x_1); +x_127 = lean_box(2); +x_128 = l_Lean_Syntax_mkStrLit(x_126, x_127); +lean_dec(x_126); +x_129 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__11; lean_inc(x_38); -x_132 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_132, 0, x_38); -lean_ctor_set(x_132, 1, x_131); -x_133 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__3; +x_130 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_130, 0, x_38); +lean_ctor_set(x_130, 1, x_129); +x_131 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__3; lean_inc(x_39); -lean_inc(x_105); -x_134 = l_Lean_addMacroScope(x_105, x_133, x_39); -x_135 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__4; -lean_inc(x_4); +lean_inc(x_104); +x_132 = l_Lean_addMacroScope(x_104, x_131, x_39); +x_133 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__4; +lean_inc(x_3); +x_134 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_134, 0, x_133); +lean_ctor_set(x_134, 1, x_3); +x_135 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__5; +lean_inc(x_3); x_136 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_136, 0, x_135); -lean_ctor_set(x_136, 1, x_4); -x_137 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__5; -lean_inc(x_4); -x_138 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_138, 0, x_137); -lean_ctor_set(x_138, 1, x_4); -x_139 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_139, 0, x_136); -lean_ctor_set(x_139, 1, x_138); -x_140 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__2; +lean_ctor_set(x_136, 1, x_3); +x_137 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_137, 0, x_134); +lean_ctor_set(x_137, 1, x_136); +x_138 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__2; lean_inc(x_38); -x_141 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_141, 0, x_38); -lean_ctor_set(x_141, 1, x_140); -lean_ctor_set(x_141, 2, x_134); -lean_ctor_set(x_141, 3, x_139); -x_142 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; -x_143 = l_Lean_addMacroScope(x_105, x_142, x_39); -x_144 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; +x_139 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_139, 0, x_38); +lean_ctor_set(x_139, 1, x_138); +lean_ctor_set(x_139, 2, x_132); +lean_ctor_set(x_139, 3, x_137); +x_140 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; +x_141 = l_Lean_addMacroScope(x_104, x_140, x_39); +x_142 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; lean_inc(x_38); -x_145 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_145, 0, x_38); -lean_ctor_set(x_145, 1, x_144); -lean_ctor_set(x_145, 2, x_143); -lean_ctor_set(x_145, 3, x_4); +x_143 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_143, 0, x_38); +lean_ctor_set(x_143, 1, x_142); +lean_ctor_set(x_143, 2, x_141); +lean_ctor_set(x_143, 3, x_3); lean_inc(x_38); -x_146 = l_Lean_Syntax_node1(x_38, x_108, x_145); -x_147 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___closed__6; +x_144 = l_Lean_Syntax_node1(x_38, x_107, x_143); +x_145 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___closed__6; lean_inc(x_38); -x_148 = l_Lean_Syntax_node2(x_38, x_147, x_141, x_146); -x_149 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__17; +x_146 = l_Lean_Syntax_node2(x_38, x_145, x_139, x_144); +x_147 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__17; lean_inc(x_38); -x_150 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_150, 0, x_38); -lean_ctor_set(x_150, 1, x_149); -x_151 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__10; +x_148 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_148, 0, x_38); +lean_ctor_set(x_148, 1, x_147); +x_149 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__10; lean_inc(x_38); -x_152 = l_Lean_Syntax_node3(x_38, x_151, x_132, x_148, x_150); +x_150 = l_Lean_Syntax_node3(x_38, x_149, x_130, x_146, x_148); lean_inc(x_38); -x_153 = l_Lean_Syntax_node3(x_38, x_108, x_127, x_130, x_152); +x_151 = l_Lean_Syntax_node3(x_38, x_107, x_125, x_128, x_150); lean_inc(x_38); -x_154 = l_Lean_Syntax_node2(x_38, x_147, x_126, x_153); -x_155 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__6; +x_152 = l_Lean_Syntax_node2(x_38, x_145, x_124, x_151); +x_153 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__6; +lean_inc(x_109); lean_inc(x_38); -x_156 = l_Lean_Syntax_node4(x_38, x_155, x_111, x_110, x_113, x_154); -x_157 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__4; +x_154 = l_Lean_Syntax_node5(x_38, x_153, x_5, x_109, x_109, x_111, x_152); +x_155 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__4; lean_inc(x_38); -x_158 = l_Lean_Syntax_node1(x_38, x_157, x_156); -x_159 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_156 = l_Lean_Syntax_node1(x_38, x_155, x_154); +x_157 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; lean_inc(x_38); -x_160 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_160, 0, x_38); -lean_ctor_set(x_160, 1, x_159); -x_161 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__2; -x_162 = l_Lean_Syntax_node4(x_38, x_161, x_107, x_158, x_160, x_6); -x_163 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_163, 0, x_162); -lean_ctor_set(x_163, 1, x_103); -return x_163; +x_158 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_158, 0, x_38); +lean_ctor_set(x_158, 1, x_157); +x_159 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__2; +x_160 = l_Lean_Syntax_node4(x_38, x_159, x_106, x_156, x_158, x_6); +x_161 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_161, 0, x_160); +lean_ctor_set(x_161, 1, x_102); +return x_161; } } } else { -lean_object* x_164; uint8_t x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; uint8_t x_169; +lean_object* x_162; uint8_t x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; uint8_t x_167; lean_dec(x_19); lean_dec(x_17); lean_dec(x_10); @@ -26784,283 +26758,279 @@ lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_2); -x_164 = lean_ctor_get(x_11, 5); -lean_inc(x_164); -x_165 = 0; -x_166 = l_Lean_SourceInfo_fromRef(x_164, x_165); -x_167 = lean_ctor_get(x_11, 10); -lean_inc(x_167); +x_162 = lean_ctor_get(x_11, 5); +lean_inc(x_162); +x_163 = 0; +x_164 = l_Lean_SourceInfo_fromRef(x_162, x_163); +x_165 = lean_ctor_get(x_11, 10); +lean_inc(x_165); lean_dec(x_11); -x_168 = lean_st_ref_get(x_12, x_13); +x_166 = lean_st_ref_get(x_12, x_13); lean_dec(x_12); -x_169 = !lean_is_exclusive(x_168); -if (x_169 == 0) +x_167 = !lean_is_exclusive(x_166); +if (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; 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; -x_170 = lean_ctor_get(x_168, 0); -x_171 = lean_ctor_get(x_170, 0); -lean_inc(x_171); -lean_dec(x_170); -x_172 = lean_environment_main_module(x_171); -x_173 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; -lean_inc(x_166); -x_174 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_174, 0, x_166); -lean_ctor_set(x_174, 1, x_173); -x_175 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__11; -x_176 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__12; -lean_inc(x_166); -x_177 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_177, 0, x_166); -lean_ctor_set(x_177, 1, x_175); -lean_ctor_set(x_177, 2, x_176); -lean_inc(x_177); -lean_inc(x_166); -x_178 = l_Lean_Syntax_node2(x_166, x_175, x_3, x_177); -x_179 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__13; -lean_inc(x_166); -x_180 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_180, 0, x_166); -lean_ctor_set(x_180, 1, x_179); -x_181 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__21; -lean_inc(x_166); -x_182 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_182, 0, x_166); -lean_ctor_set(x_182, 1, x_181); -x_183 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___closed__8; -lean_inc(x_167); -lean_inc(x_172); -x_184 = l_Lean_addMacroScope(x_172, x_183, x_167); -x_185 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___closed__9; -lean_inc(x_4); -x_186 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_186, 0, x_185); -lean_ctor_set(x_186, 1, x_4); -lean_inc(x_4); -x_187 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_187, 0, x_186); -lean_ctor_set(x_187, 1, x_4); -x_188 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___closed__7; -lean_inc(x_166); -x_189 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_189, 0, x_166); -lean_ctor_set(x_189, 1, x_188); -lean_ctor_set(x_189, 2, x_184); -lean_ctor_set(x_189, 3, x_187); -x_190 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__20; -lean_inc(x_166); -x_191 = l_Lean_Syntax_node2(x_166, x_190, x_182, x_189); -x_192 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__22(x_5); -x_193 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__11; -lean_inc(x_166); -x_194 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_194, 0, x_166); -lean_ctor_set(x_194, 1, x_193); -x_195 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__3; -lean_inc(x_167); -lean_inc(x_172); -x_196 = l_Lean_addMacroScope(x_172, x_195, x_167); -x_197 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__4; -lean_inc(x_4); +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_168 = lean_ctor_get(x_166, 0); +x_169 = lean_ctor_get(x_168, 0); +lean_inc(x_169); +lean_dec(x_168); +x_170 = lean_environment_main_module(x_169); +x_171 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; +lean_inc(x_164); +x_172 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_172, 0, x_164); +lean_ctor_set(x_172, 1, x_171); +x_173 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__11; +x_174 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__12; +lean_inc(x_164); +x_175 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_175, 0, x_164); +lean_ctor_set(x_175, 1, x_173); +lean_ctor_set(x_175, 2, x_174); +x_176 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__13; +lean_inc(x_164); +x_177 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_177, 0, x_164); +lean_ctor_set(x_177, 1, x_176); +x_178 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__21; +lean_inc(x_164); +x_179 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_179, 0, x_164); +lean_ctor_set(x_179, 1, x_178); +x_180 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___closed__8; +lean_inc(x_165); +lean_inc(x_170); +x_181 = l_Lean_addMacroScope(x_170, x_180, x_165); +x_182 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___closed__9; +lean_inc(x_3); +x_183 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_183, 0, x_182); +lean_ctor_set(x_183, 1, x_3); +lean_inc(x_3); +x_184 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_184, 0, x_183); +lean_ctor_set(x_184, 1, x_3); +x_185 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___closed__7; +lean_inc(x_164); +x_186 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_186, 0, x_164); +lean_ctor_set(x_186, 1, x_185); +lean_ctor_set(x_186, 2, x_181); +lean_ctor_set(x_186, 3, x_184); +x_187 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__20; +lean_inc(x_164); +x_188 = l_Lean_Syntax_node2(x_164, x_187, x_179, x_186); +x_189 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__22(x_4); +x_190 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__11; +lean_inc(x_164); +x_191 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_191, 0, x_164); +lean_ctor_set(x_191, 1, x_190); +x_192 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__3; +lean_inc(x_165); +lean_inc(x_170); +x_193 = l_Lean_addMacroScope(x_170, x_192, x_165); +x_194 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__4; +lean_inc(x_3); +x_195 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_195, 0, x_194); +lean_ctor_set(x_195, 1, x_3); +x_196 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__5; +lean_inc(x_3); +x_197 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_197, 0, x_196); +lean_ctor_set(x_197, 1, x_3); x_198 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_198, 0, x_197); -lean_ctor_set(x_198, 1, x_4); -x_199 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__5; -lean_inc(x_4); -x_200 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_200, 0, x_199); -lean_ctor_set(x_200, 1, x_4); -x_201 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_201, 0, x_198); -lean_ctor_set(x_201, 1, x_200); -x_202 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__2; -lean_inc(x_166); -x_203 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_203, 0, x_166); -lean_ctor_set(x_203, 1, x_202); -lean_ctor_set(x_203, 2, x_196); -lean_ctor_set(x_203, 3, x_201); -x_204 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; -x_205 = l_Lean_addMacroScope(x_172, x_204, x_167); -x_206 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; -lean_inc(x_166); -x_207 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_207, 0, x_166); -lean_ctor_set(x_207, 1, x_206); -lean_ctor_set(x_207, 2, x_205); -lean_ctor_set(x_207, 3, x_4); -lean_inc(x_166); -x_208 = l_Lean_Syntax_node1(x_166, x_175, x_207); -x_209 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___closed__6; -lean_inc(x_166); -x_210 = l_Lean_Syntax_node2(x_166, x_209, x_203, x_208); -x_211 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__17; -lean_inc(x_166); -x_212 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_212, 0, x_166); -lean_ctor_set(x_212, 1, x_211); -x_213 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__10; -lean_inc(x_166); -x_214 = l_Lean_Syntax_node3(x_166, x_213, x_194, x_210, x_212); -lean_inc(x_166); -x_215 = l_Lean_Syntax_node2(x_166, x_175, x_192, x_214); -lean_inc(x_166); -x_216 = l_Lean_Syntax_node2(x_166, x_209, x_191, x_215); -x_217 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__6; -lean_inc(x_166); -x_218 = l_Lean_Syntax_node4(x_166, x_217, x_178, x_177, x_180, x_216); -x_219 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__4; -lean_inc(x_166); -x_220 = l_Lean_Syntax_node1(x_166, x_219, x_218); -x_221 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -lean_inc(x_166); -x_222 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_222, 0, x_166); -lean_ctor_set(x_222, 1, x_221); -x_223 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__2; -x_224 = l_Lean_Syntax_node4(x_166, x_223, x_174, x_220, x_222, x_6); -lean_ctor_set(x_168, 0, x_224); -return x_168; +lean_ctor_set(x_198, 0, x_195); +lean_ctor_set(x_198, 1, x_197); +x_199 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__2; +lean_inc(x_164); +x_200 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_200, 0, x_164); +lean_ctor_set(x_200, 1, x_199); +lean_ctor_set(x_200, 2, x_193); +lean_ctor_set(x_200, 3, x_198); +x_201 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; +x_202 = l_Lean_addMacroScope(x_170, x_201, x_165); +x_203 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; +lean_inc(x_164); +x_204 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_204, 0, x_164); +lean_ctor_set(x_204, 1, x_203); +lean_ctor_set(x_204, 2, x_202); +lean_ctor_set(x_204, 3, x_3); +lean_inc(x_164); +x_205 = l_Lean_Syntax_node1(x_164, x_173, x_204); +x_206 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___closed__6; +lean_inc(x_164); +x_207 = l_Lean_Syntax_node2(x_164, x_206, x_200, x_205); +x_208 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__17; +lean_inc(x_164); +x_209 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_209, 0, x_164); +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__23___lambda__3___closed__10; +lean_inc(x_164); +x_211 = l_Lean_Syntax_node3(x_164, x_210, x_191, x_207, x_209); +lean_inc(x_164); +x_212 = l_Lean_Syntax_node2(x_164, x_173, x_189, x_211); +lean_inc(x_164); +x_213 = l_Lean_Syntax_node2(x_164, x_206, x_188, x_212); +x_214 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__6; +lean_inc(x_175); +lean_inc(x_164); +x_215 = l_Lean_Syntax_node5(x_164, x_214, x_5, x_175, x_175, x_177, x_213); +x_216 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__4; +lean_inc(x_164); +x_217 = l_Lean_Syntax_node1(x_164, x_216, x_215); +x_218 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +lean_inc(x_164); +x_219 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_219, 0, x_164); +lean_ctor_set(x_219, 1, x_218); +x_220 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__2; +x_221 = l_Lean_Syntax_node4(x_164, x_220, x_172, x_217, x_219, x_6); +lean_ctor_set(x_166, 0, x_221); +return x_166; } else { -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; -x_225 = lean_ctor_get(x_168, 0); -x_226 = lean_ctor_get(x_168, 1); -lean_inc(x_226); +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; +x_222 = lean_ctor_get(x_166, 0); +x_223 = lean_ctor_get(x_166, 1); +lean_inc(x_223); +lean_inc(x_222); +lean_dec(x_166); +x_224 = lean_ctor_get(x_222, 0); +lean_inc(x_224); +lean_dec(x_222); +x_225 = lean_environment_main_module(x_224); +x_226 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; +lean_inc(x_164); +x_227 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_227, 0, x_164); +lean_ctor_set(x_227, 1, x_226); +x_228 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__11; +x_229 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__12; +lean_inc(x_164); +x_230 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_230, 0, x_164); +lean_ctor_set(x_230, 1, x_228); +lean_ctor_set(x_230, 2, x_229); +x_231 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__13; +lean_inc(x_164); +x_232 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_232, 0, x_164); +lean_ctor_set(x_232, 1, x_231); +x_233 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__21; +lean_inc(x_164); +x_234 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_234, 0, x_164); +lean_ctor_set(x_234, 1, x_233); +x_235 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___closed__8; +lean_inc(x_165); lean_inc(x_225); -lean_dec(x_168); -x_227 = lean_ctor_get(x_225, 0); -lean_inc(x_227); -lean_dec(x_225); -x_228 = lean_environment_main_module(x_227); -x_229 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; -lean_inc(x_166); -x_230 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_230, 0, x_166); -lean_ctor_set(x_230, 1, x_229); -x_231 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__11; -x_232 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__12; -lean_inc(x_166); -x_233 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_233, 0, x_166); -lean_ctor_set(x_233, 1, x_231); -lean_ctor_set(x_233, 2, x_232); -lean_inc(x_233); -lean_inc(x_166); -x_234 = l_Lean_Syntax_node2(x_166, x_231, x_3, x_233); -x_235 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__13; -lean_inc(x_166); -x_236 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_236, 0, x_166); -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__23___closed__21; -lean_inc(x_166); -x_238 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_238, 0, x_166); -lean_ctor_set(x_238, 1, x_237); -x_239 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___closed__8; -lean_inc(x_167); -lean_inc(x_228); -x_240 = l_Lean_addMacroScope(x_228, x_239, x_167); -x_241 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___closed__9; -lean_inc(x_4); -x_242 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_242, 0, x_241); -lean_ctor_set(x_242, 1, x_4); -lean_inc(x_4); -x_243 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_243, 0, x_242); -lean_ctor_set(x_243, 1, x_4); -x_244 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___closed__7; -lean_inc(x_166); -x_245 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_245, 0, x_166); -lean_ctor_set(x_245, 1, x_244); -lean_ctor_set(x_245, 2, x_240); -lean_ctor_set(x_245, 3, x_243); -x_246 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__20; -lean_inc(x_166); -x_247 = l_Lean_Syntax_node2(x_166, x_246, x_238, x_245); -x_248 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__22(x_5); -x_249 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__11; -lean_inc(x_166); -x_250 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_250, 0, x_166); -lean_ctor_set(x_250, 1, x_249); -x_251 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__3; -lean_inc(x_167); -lean_inc(x_228); -x_252 = l_Lean_addMacroScope(x_228, x_251, x_167); -x_253 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__4; -lean_inc(x_4); -x_254 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_254, 0, x_253); -lean_ctor_set(x_254, 1, x_4); -x_255 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__5; -lean_inc(x_4); -x_256 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_256, 0, x_255); -lean_ctor_set(x_256, 1, x_4); -x_257 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_257, 0, x_254); -lean_ctor_set(x_257, 1, x_256); -x_258 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__2; -lean_inc(x_166); +x_236 = l_Lean_addMacroScope(x_225, x_235, x_165); +x_237 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___closed__9; +lean_inc(x_3); +x_238 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_238, 0, x_237); +lean_ctor_set(x_238, 1, x_3); +lean_inc(x_3); +x_239 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_239, 0, x_238); +lean_ctor_set(x_239, 1, x_3); +x_240 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___closed__7; +lean_inc(x_164); +x_241 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_241, 0, x_164); +lean_ctor_set(x_241, 1, x_240); +lean_ctor_set(x_241, 2, x_236); +lean_ctor_set(x_241, 3, x_239); +x_242 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__20; +lean_inc(x_164); +x_243 = l_Lean_Syntax_node2(x_164, x_242, x_234, x_241); +x_244 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__22(x_4); +x_245 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__11; +lean_inc(x_164); +x_246 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_246, 0, x_164); +lean_ctor_set(x_246, 1, x_245); +x_247 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__3; +lean_inc(x_165); +lean_inc(x_225); +x_248 = l_Lean_addMacroScope(x_225, x_247, x_165); +x_249 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__4; +lean_inc(x_3); +x_250 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_250, 0, x_249); +lean_ctor_set(x_250, 1, x_3); +x_251 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__5; +lean_inc(x_3); +x_252 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_252, 0, x_251); +lean_ctor_set(x_252, 1, x_3); +x_253 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_253, 0, x_250); +lean_ctor_set(x_253, 1, x_252); +x_254 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__23___closed__2; +lean_inc(x_164); +x_255 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_255, 0, x_164); +lean_ctor_set(x_255, 1, x_254); +lean_ctor_set(x_255, 2, x_248); +lean_ctor_set(x_255, 3, x_253); +x_256 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; +x_257 = l_Lean_addMacroScope(x_225, x_256, x_165); +x_258 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; +lean_inc(x_164); x_259 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_259, 0, x_166); +lean_ctor_set(x_259, 0, x_164); lean_ctor_set(x_259, 1, x_258); -lean_ctor_set(x_259, 2, x_252); -lean_ctor_set(x_259, 3, x_257); -x_260 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; -x_261 = l_Lean_addMacroScope(x_228, x_260, x_167); -x_262 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; -lean_inc(x_166); -x_263 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_263, 0, x_166); -lean_ctor_set(x_263, 1, x_262); -lean_ctor_set(x_263, 2, x_261); -lean_ctor_set(x_263, 3, x_4); -lean_inc(x_166); -x_264 = l_Lean_Syntax_node1(x_166, x_231, x_263); -x_265 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___closed__6; -lean_inc(x_166); -x_266 = l_Lean_Syntax_node2(x_166, x_265, x_259, x_264); -x_267 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__17; -lean_inc(x_166); -x_268 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_268, 0, x_166); -lean_ctor_set(x_268, 1, x_267); -x_269 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__10; -lean_inc(x_166); -x_270 = l_Lean_Syntax_node3(x_166, x_269, x_250, x_266, x_268); -lean_inc(x_166); -x_271 = l_Lean_Syntax_node2(x_166, x_231, x_248, x_270); -lean_inc(x_166); -x_272 = l_Lean_Syntax_node2(x_166, x_265, x_247, x_271); -x_273 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__6; -lean_inc(x_166); -x_274 = l_Lean_Syntax_node4(x_166, x_273, x_234, x_233, x_236, x_272); -x_275 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__4; -lean_inc(x_166); -x_276 = l_Lean_Syntax_node1(x_166, x_275, x_274); -x_277 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -lean_inc(x_166); -x_278 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_278, 0, x_166); -lean_ctor_set(x_278, 1, x_277); -x_279 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__2; -x_280 = l_Lean_Syntax_node4(x_166, x_279, x_230, x_276, x_278, x_6); -x_281 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_281, 0, x_280); -lean_ctor_set(x_281, 1, x_226); -return x_281; +lean_ctor_set(x_259, 2, x_257); +lean_ctor_set(x_259, 3, x_3); +lean_inc(x_164); +x_260 = l_Lean_Syntax_node1(x_164, x_228, x_259); +x_261 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___closed__6; +lean_inc(x_164); +x_262 = l_Lean_Syntax_node2(x_164, x_261, x_255, x_260); +x_263 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__17; +lean_inc(x_164); +x_264 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_264, 0, x_164); +lean_ctor_set(x_264, 1, x_263); +x_265 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__10; +lean_inc(x_164); +x_266 = l_Lean_Syntax_node3(x_164, x_265, x_246, x_262, x_264); +lean_inc(x_164); +x_267 = l_Lean_Syntax_node2(x_164, x_228, x_244, x_266); +lean_inc(x_164); +x_268 = l_Lean_Syntax_node2(x_164, x_261, x_243, x_267); +x_269 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__6; +lean_inc(x_230); +lean_inc(x_164); +x_270 = l_Lean_Syntax_node5(x_164, x_269, x_5, x_230, x_230, x_232, x_268); +x_271 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__4; +lean_inc(x_164); +x_272 = l_Lean_Syntax_node1(x_164, x_271, x_270); +x_273 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +lean_inc(x_164); +x_274 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_274, 0, x_164); +lean_ctor_set(x_274, 1, x_273); +x_275 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__2; +x_276 = l_Lean_Syntax_node4(x_164, x_275, x_227, x_272, x_274, x_6); +x_277 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_277, 0, x_276); +lean_ctor_set(x_277, 1, x_223); +return x_277; } } } else { -lean_object* x_282; uint8_t x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; uint8_t x_287; +lean_object* x_278; uint8_t x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; uint8_t x_283; lean_dec(x_19); lean_dec(x_17); lean_dec(x_10); @@ -27068,405 +27038,401 @@ lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_2); -x_282 = lean_ctor_get(x_11, 5); -lean_inc(x_282); -x_283 = 0; -x_284 = l_Lean_SourceInfo_fromRef(x_282, x_283); -x_285 = lean_ctor_get(x_11, 10); -lean_inc(x_285); +x_278 = lean_ctor_get(x_11, 5); +lean_inc(x_278); +x_279 = 0; +x_280 = l_Lean_SourceInfo_fromRef(x_278, x_279); +x_281 = lean_ctor_get(x_11, 10); +lean_inc(x_281); lean_dec(x_11); -x_286 = lean_st_ref_get(x_12, x_13); +x_282 = lean_st_ref_get(x_12, x_13); lean_dec(x_12); -x_287 = !lean_is_exclusive(x_286); -if (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_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; -x_288 = lean_ctor_get(x_286, 0); -x_289 = lean_ctor_get(x_288, 0); -lean_inc(x_289); -lean_dec(x_288); -x_290 = lean_environment_main_module(x_289); -x_291 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; -lean_inc(x_284); -x_292 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_292, 0, x_284); -lean_ctor_set(x_292, 1, x_291); -x_293 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__11; -x_294 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__12; -lean_inc(x_284); -x_295 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_295, 0, x_284); -lean_ctor_set(x_295, 1, x_293); -lean_ctor_set(x_295, 2, x_294); -lean_inc(x_295); -lean_inc(x_284); -x_296 = l_Lean_Syntax_node2(x_284, x_293, x_3, x_295); -x_297 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__13; -lean_inc(x_284); -x_298 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_298, 0, x_284); -lean_ctor_set(x_298, 1, x_297); -x_299 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__40; -lean_inc(x_285); -lean_inc(x_290); -x_300 = l_Lean_addMacroScope(x_290, x_299, x_285); -lean_inc(x_4); -x_301 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_301, 0, x_299); -lean_ctor_set(x_301, 1, x_4); -x_302 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__42; -lean_inc(x_4); -x_303 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_303, 0, x_302); -lean_ctor_set(x_303, 1, x_4); -x_304 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_304, 0, x_301); -lean_ctor_set(x_304, 1, x_303); -x_305 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__39; -lean_inc(x_284); -x_306 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_306, 0, x_284); -lean_ctor_set(x_306, 1, x_305); -lean_ctor_set(x_306, 2, x_300); -lean_ctor_set(x_306, 3, x_304); -x_307 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__11; -lean_inc(x_284); -x_308 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_308, 0, x_284); -lean_ctor_set(x_308, 1, x_307); -x_309 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__21; -lean_inc(x_284); -x_310 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_310, 0, x_284); -lean_ctor_set(x_310, 1, x_309); -x_311 = l_Lean_Elab_Term_Quotation_mkSyntaxQuotation___closed__40; -lean_inc(x_285); -lean_inc(x_290); -x_312 = l_Lean_addMacroScope(x_290, x_311, x_285); -x_313 = l_Lean_Elab_Term_Quotation_mkSyntaxQuotation___closed__41; -lean_inc(x_4); -x_314 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_314, 0, x_313); -lean_ctor_set(x_314, 1, x_4); -x_315 = l_Lean_Elab_Term_Quotation_mkSyntaxQuotation___closed__43; -lean_inc(x_4); -x_316 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_316, 0, x_315); -lean_ctor_set(x_316, 1, x_4); -x_317 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_317, 0, x_314); -lean_ctor_set(x_317, 1, x_316); -x_318 = l_Lean_Elab_Term_Quotation_mkSyntaxQuotation___closed__39; -lean_inc(x_284); -x_319 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_319, 0, x_284); -lean_ctor_set(x_319, 1, x_318); -lean_ctor_set(x_319, 2, x_312); -lean_ctor_set(x_319, 3, x_317); -x_320 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__20; -lean_inc(x_284); -x_321 = l_Lean_Syntax_node2(x_284, x_320, x_310, x_319); -x_322 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__22(x_5); -lean_inc(x_284); -x_323 = l_Lean_Syntax_node1(x_284, x_293, x_322); -x_324 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___closed__6; -lean_inc(x_284); -x_325 = l_Lean_Syntax_node2(x_284, x_324, x_321, x_323); -x_326 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__17; -lean_inc(x_284); -x_327 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_327, 0, x_284); -lean_ctor_set(x_327, 1, x_326); -x_328 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__10; -lean_inc(x_327); -lean_inc(x_308); -lean_inc(x_284); -x_329 = l_Lean_Syntax_node3(x_284, x_328, x_308, x_325, x_327); -x_330 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___closed__13; +x_283 = !lean_is_exclusive(x_282); +if (x_283 == 0) +{ +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; +x_284 = lean_ctor_get(x_282, 0); +x_285 = lean_ctor_get(x_284, 0); lean_inc(x_285); -lean_inc(x_290); -x_331 = l_Lean_addMacroScope(x_290, x_330, x_285); -x_332 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___closed__14; -lean_inc(x_4); -x_333 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_333, 0, x_332); -lean_ctor_set(x_333, 1, x_4); -x_334 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___closed__15; -lean_inc(x_4); -x_335 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_335, 0, x_334); -lean_ctor_set(x_335, 1, x_4); -x_336 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_336, 0, x_333); -lean_ctor_set(x_336, 1, x_335); -x_337 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___closed__11; -lean_inc(x_284); -x_338 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_338, 0, x_284); -lean_ctor_set(x_338, 1, x_337); -lean_ctor_set(x_338, 2, x_331); -lean_ctor_set(x_338, 3, x_336); -x_339 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; -x_340 = l_Lean_addMacroScope(x_290, x_339, x_285); -x_341 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; -lean_inc(x_284); -x_342 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_342, 0, x_284); -lean_ctor_set(x_342, 1, x_341); -lean_ctor_set(x_342, 2, x_340); -lean_ctor_set(x_342, 3, x_4); -lean_inc(x_284); -x_343 = l_Lean_Syntax_node1(x_284, x_293, x_342); -lean_inc(x_284); -x_344 = l_Lean_Syntax_node2(x_284, x_324, x_338, x_343); -lean_inc(x_284); -x_345 = l_Lean_Syntax_node3(x_284, x_328, x_308, x_344, x_327); -lean_inc(x_284); -x_346 = l_Lean_Syntax_node2(x_284, x_293, x_329, x_345); -lean_inc(x_284); -x_347 = l_Lean_Syntax_node2(x_284, x_324, x_306, x_346); -x_348 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__6; -lean_inc(x_284); -x_349 = l_Lean_Syntax_node4(x_284, x_348, x_296, x_295, x_298, x_347); -x_350 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__4; -lean_inc(x_284); -x_351 = l_Lean_Syntax_node1(x_284, x_350, x_349); -x_352 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -lean_inc(x_284); -x_353 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_353, 0, x_284); -lean_ctor_set(x_353, 1, x_352); -x_354 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__2; -x_355 = l_Lean_Syntax_node4(x_284, x_354, x_292, x_351, x_353, x_6); -lean_ctor_set(x_286, 0, x_355); -return x_286; -} -else -{ -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; -x_356 = lean_ctor_get(x_286, 0); -x_357 = lean_ctor_get(x_286, 1); -lean_inc(x_357); -lean_inc(x_356); -lean_dec(x_286); -x_358 = lean_ctor_get(x_356, 0); -lean_inc(x_358); -lean_dec(x_356); -x_359 = lean_environment_main_module(x_358); -x_360 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; -lean_inc(x_284); +lean_dec(x_284); +x_286 = lean_environment_main_module(x_285); +x_287 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; +lean_inc(x_280); +x_288 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_288, 0, x_280); +lean_ctor_set(x_288, 1, x_287); +x_289 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__11; +x_290 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__12; +lean_inc(x_280); +x_291 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_291, 0, x_280); +lean_ctor_set(x_291, 1, x_289); +lean_ctor_set(x_291, 2, x_290); +x_292 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__13; +lean_inc(x_280); +x_293 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_293, 0, x_280); +lean_ctor_set(x_293, 1, x_292); +x_294 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__40; +lean_inc(x_281); +lean_inc(x_286); +x_295 = l_Lean_addMacroScope(x_286, x_294, x_281); +lean_inc(x_3); +x_296 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_296, 0, x_294); +lean_ctor_set(x_296, 1, x_3); +x_297 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__42; +lean_inc(x_3); +x_298 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_298, 0, x_297); +lean_ctor_set(x_298, 1, x_3); +x_299 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_299, 0, x_296); +lean_ctor_set(x_299, 1, x_298); +x_300 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__39; +lean_inc(x_280); +x_301 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_301, 0, x_280); +lean_ctor_set(x_301, 1, x_300); +lean_ctor_set(x_301, 2, x_295); +lean_ctor_set(x_301, 3, x_299); +x_302 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__11; +lean_inc(x_280); +x_303 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_303, 0, x_280); +lean_ctor_set(x_303, 1, x_302); +x_304 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__21; +lean_inc(x_280); +x_305 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_305, 0, x_280); +lean_ctor_set(x_305, 1, x_304); +x_306 = l_Lean_Elab_Term_Quotation_mkSyntaxQuotation___closed__40; +lean_inc(x_281); +lean_inc(x_286); +x_307 = l_Lean_addMacroScope(x_286, x_306, x_281); +x_308 = l_Lean_Elab_Term_Quotation_mkSyntaxQuotation___closed__41; +lean_inc(x_3); +x_309 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_309, 0, x_308); +lean_ctor_set(x_309, 1, x_3); +x_310 = l_Lean_Elab_Term_Quotation_mkSyntaxQuotation___closed__43; +lean_inc(x_3); +x_311 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_311, 0, x_310); +lean_ctor_set(x_311, 1, x_3); +x_312 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_312, 0, x_309); +lean_ctor_set(x_312, 1, x_311); +x_313 = l_Lean_Elab_Term_Quotation_mkSyntaxQuotation___closed__39; +lean_inc(x_280); +x_314 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_314, 0, x_280); +lean_ctor_set(x_314, 1, x_313); +lean_ctor_set(x_314, 2, x_307); +lean_ctor_set(x_314, 3, x_312); +x_315 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__20; +lean_inc(x_280); +x_316 = l_Lean_Syntax_node2(x_280, x_315, x_305, x_314); +x_317 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__22(x_4); +lean_inc(x_280); +x_318 = l_Lean_Syntax_node1(x_280, x_289, x_317); +x_319 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___closed__6; +lean_inc(x_280); +x_320 = l_Lean_Syntax_node2(x_280, x_319, x_316, x_318); +x_321 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__17; +lean_inc(x_280); +x_322 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_322, 0, x_280); +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__23___lambda__3___closed__10; +lean_inc(x_322); +lean_inc(x_303); +lean_inc(x_280); +x_324 = l_Lean_Syntax_node3(x_280, x_323, x_303, x_320, x_322); +x_325 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___closed__13; +lean_inc(x_281); +lean_inc(x_286); +x_326 = l_Lean_addMacroScope(x_286, x_325, x_281); +x_327 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___closed__14; +lean_inc(x_3); +x_328 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_328, 0, x_327); +lean_ctor_set(x_328, 1, x_3); +x_329 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___closed__15; +lean_inc(x_3); +x_330 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_330, 0, x_329); +lean_ctor_set(x_330, 1, x_3); +x_331 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_331, 0, x_328); +lean_ctor_set(x_331, 1, x_330); +x_332 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___closed__11; +lean_inc(x_280); +x_333 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_333, 0, x_280); +lean_ctor_set(x_333, 1, x_332); +lean_ctor_set(x_333, 2, x_326); +lean_ctor_set(x_333, 3, x_331); +x_334 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; +x_335 = l_Lean_addMacroScope(x_286, x_334, x_281); +x_336 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; +lean_inc(x_280); +x_337 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_337, 0, x_280); +lean_ctor_set(x_337, 1, x_336); +lean_ctor_set(x_337, 2, x_335); +lean_ctor_set(x_337, 3, x_3); +lean_inc(x_280); +x_338 = l_Lean_Syntax_node1(x_280, x_289, x_337); +lean_inc(x_280); +x_339 = l_Lean_Syntax_node2(x_280, x_319, x_333, x_338); +lean_inc(x_280); +x_340 = l_Lean_Syntax_node3(x_280, x_323, x_303, x_339, x_322); +lean_inc(x_280); +x_341 = l_Lean_Syntax_node2(x_280, x_289, x_324, x_340); +lean_inc(x_280); +x_342 = l_Lean_Syntax_node2(x_280, x_319, x_301, x_341); +x_343 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__6; +lean_inc(x_291); +lean_inc(x_280); +x_344 = l_Lean_Syntax_node5(x_280, x_343, x_5, x_291, x_291, x_293, x_342); +x_345 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__4; +lean_inc(x_280); +x_346 = l_Lean_Syntax_node1(x_280, x_345, x_344); +x_347 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +lean_inc(x_280); +x_348 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_348, 0, x_280); +lean_ctor_set(x_348, 1, x_347); +x_349 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__2; +x_350 = l_Lean_Syntax_node4(x_280, x_349, x_288, x_346, x_348, x_6); +lean_ctor_set(x_282, 0, x_350); +return x_282; +} +else +{ +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; +x_351 = lean_ctor_get(x_282, 0); +x_352 = lean_ctor_get(x_282, 1); +lean_inc(x_352); +lean_inc(x_351); +lean_dec(x_282); +x_353 = lean_ctor_get(x_351, 0); +lean_inc(x_353); +lean_dec(x_351); +x_354 = lean_environment_main_module(x_353); +x_355 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; +lean_inc(x_280); +x_356 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_356, 0, x_280); +lean_ctor_set(x_356, 1, x_355); +x_357 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__11; +x_358 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__12; +lean_inc(x_280); +x_359 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_359, 0, x_280); +lean_ctor_set(x_359, 1, x_357); +lean_ctor_set(x_359, 2, x_358); +x_360 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__13; +lean_inc(x_280); x_361 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_361, 0, x_284); +lean_ctor_set(x_361, 0, x_280); lean_ctor_set(x_361, 1, x_360); -x_362 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__11; -x_363 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__12; -lean_inc(x_284); -x_364 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_364, 0, x_284); -lean_ctor_set(x_364, 1, x_362); -lean_ctor_set(x_364, 2, x_363); -lean_inc(x_364); -lean_inc(x_284); -x_365 = l_Lean_Syntax_node2(x_284, x_362, x_3, x_364); -x_366 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__13; -lean_inc(x_284); -x_367 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_367, 0, x_284); +x_362 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__40; +lean_inc(x_281); +lean_inc(x_354); +x_363 = l_Lean_addMacroScope(x_354, x_362, x_281); +lean_inc(x_3); +x_364 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_364, 0, x_362); +lean_ctor_set(x_364, 1, x_3); +x_365 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__42; +lean_inc(x_3); +x_366 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_366, 0, x_365); +lean_ctor_set(x_366, 1, x_3); +x_367 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_367, 0, x_364); lean_ctor_set(x_367, 1, x_366); -x_368 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__40; -lean_inc(x_285); -lean_inc(x_359); -x_369 = l_Lean_addMacroScope(x_359, x_368, x_285); -lean_inc(x_4); -x_370 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_370, 0, x_368); -lean_ctor_set(x_370, 1, x_4); -x_371 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__42; -lean_inc(x_4); -x_372 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_372, 0, x_371); -lean_ctor_set(x_372, 1, x_4); -x_373 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_373, 0, x_370); +x_368 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__39; +lean_inc(x_280); +x_369 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_369, 0, x_280); +lean_ctor_set(x_369, 1, x_368); +lean_ctor_set(x_369, 2, x_363); +lean_ctor_set(x_369, 3, x_367); +x_370 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__11; +lean_inc(x_280); +x_371 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_371, 0, x_280); +lean_ctor_set(x_371, 1, x_370); +x_372 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__21; +lean_inc(x_280); +x_373 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_373, 0, x_280); lean_ctor_set(x_373, 1, x_372); -x_374 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__39; -lean_inc(x_284); -x_375 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_375, 0, x_284); -lean_ctor_set(x_375, 1, x_374); -lean_ctor_set(x_375, 2, x_369); -lean_ctor_set(x_375, 3, x_373); -x_376 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__11; -lean_inc(x_284); -x_377 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_377, 0, x_284); -lean_ctor_set(x_377, 1, x_376); -x_378 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__21; -lean_inc(x_284); -x_379 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_379, 0, x_284); -lean_ctor_set(x_379, 1, x_378); -x_380 = l_Lean_Elab_Term_Quotation_mkSyntaxQuotation___closed__40; -lean_inc(x_285); -lean_inc(x_359); -x_381 = l_Lean_addMacroScope(x_359, x_380, x_285); -x_382 = l_Lean_Elab_Term_Quotation_mkSyntaxQuotation___closed__41; -lean_inc(x_4); -x_383 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_383, 0, x_382); -lean_ctor_set(x_383, 1, x_4); -x_384 = l_Lean_Elab_Term_Quotation_mkSyntaxQuotation___closed__43; -lean_inc(x_4); -x_385 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_385, 0, x_384); -lean_ctor_set(x_385, 1, x_4); -x_386 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_386, 0, x_383); -lean_ctor_set(x_386, 1, x_385); -x_387 = l_Lean_Elab_Term_Quotation_mkSyntaxQuotation___closed__39; -lean_inc(x_284); -x_388 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_388, 0, x_284); -lean_ctor_set(x_388, 1, x_387); -lean_ctor_set(x_388, 2, x_381); -lean_ctor_set(x_388, 3, x_386); -x_389 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__20; -lean_inc(x_284); -x_390 = l_Lean_Syntax_node2(x_284, x_389, x_379, x_388); -x_391 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__22(x_5); -lean_inc(x_284); -x_392 = l_Lean_Syntax_node1(x_284, x_362, x_391); -x_393 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___closed__6; -lean_inc(x_284); -x_394 = l_Lean_Syntax_node2(x_284, x_393, x_390, x_392); -x_395 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__17; -lean_inc(x_284); -x_396 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_396, 0, x_284); -lean_ctor_set(x_396, 1, x_395); -x_397 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__10; -lean_inc(x_396); -lean_inc(x_377); -lean_inc(x_284); -x_398 = l_Lean_Syntax_node3(x_284, x_397, x_377, x_394, x_396); -x_399 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___closed__13; -lean_inc(x_285); -lean_inc(x_359); -x_400 = l_Lean_addMacroScope(x_359, x_399, x_285); -x_401 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___closed__14; -lean_inc(x_4); -x_402 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_402, 0, x_401); -lean_ctor_set(x_402, 1, x_4); -x_403 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___closed__15; -lean_inc(x_4); -x_404 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_404, 0, x_403); -lean_ctor_set(x_404, 1, x_4); -x_405 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_405, 0, x_402); +x_374 = l_Lean_Elab_Term_Quotation_mkSyntaxQuotation___closed__40; +lean_inc(x_281); +lean_inc(x_354); +x_375 = l_Lean_addMacroScope(x_354, x_374, x_281); +x_376 = l_Lean_Elab_Term_Quotation_mkSyntaxQuotation___closed__41; +lean_inc(x_3); +x_377 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_377, 0, x_376); +lean_ctor_set(x_377, 1, x_3); +x_378 = l_Lean_Elab_Term_Quotation_mkSyntaxQuotation___closed__43; +lean_inc(x_3); +x_379 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_379, 0, x_378); +lean_ctor_set(x_379, 1, x_3); +x_380 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_380, 0, x_377); +lean_ctor_set(x_380, 1, x_379); +x_381 = l_Lean_Elab_Term_Quotation_mkSyntaxQuotation___closed__39; +lean_inc(x_280); +x_382 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_382, 0, x_280); +lean_ctor_set(x_382, 1, x_381); +lean_ctor_set(x_382, 2, x_375); +lean_ctor_set(x_382, 3, x_380); +x_383 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__20; +lean_inc(x_280); +x_384 = l_Lean_Syntax_node2(x_280, x_383, x_373, x_382); +x_385 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__22(x_4); +lean_inc(x_280); +x_386 = l_Lean_Syntax_node1(x_280, x_357, x_385); +x_387 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___closed__6; +lean_inc(x_280); +x_388 = l_Lean_Syntax_node2(x_280, x_387, x_384, x_386); +x_389 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__17; +lean_inc(x_280); +x_390 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_390, 0, x_280); +lean_ctor_set(x_390, 1, x_389); +x_391 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___lambda__3___closed__10; +lean_inc(x_390); +lean_inc(x_371); +lean_inc(x_280); +x_392 = l_Lean_Syntax_node3(x_280, x_391, x_371, x_388, x_390); +x_393 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___closed__13; +lean_inc(x_281); +lean_inc(x_354); +x_394 = l_Lean_addMacroScope(x_354, x_393, x_281); +x_395 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___closed__14; +lean_inc(x_3); +x_396 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_396, 0, x_395); +lean_ctor_set(x_396, 1, x_3); +x_397 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___closed__15; +lean_inc(x_3); +x_398 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_398, 0, x_397); +lean_ctor_set(x_398, 1, x_3); +x_399 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_399, 0, x_396); +lean_ctor_set(x_399, 1, x_398); +x_400 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___closed__11; +lean_inc(x_280); +x_401 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_401, 0, x_280); +lean_ctor_set(x_401, 1, x_400); +lean_ctor_set(x_401, 2, x_394); +lean_ctor_set(x_401, 3, x_399); +x_402 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; +x_403 = l_Lean_addMacroScope(x_354, x_402, x_281); +x_404 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; +lean_inc(x_280); +x_405 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_405, 0, x_280); lean_ctor_set(x_405, 1, x_404); -x_406 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___closed__11; -lean_inc(x_284); -x_407 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_407, 0, x_284); -lean_ctor_set(x_407, 1, x_406); -lean_ctor_set(x_407, 2, x_400); -lean_ctor_set(x_407, 3, x_405); -x_408 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; -x_409 = l_Lean_addMacroScope(x_359, x_408, x_285); -x_410 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; -lean_inc(x_284); -x_411 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_411, 0, x_284); -lean_ctor_set(x_411, 1, x_410); -lean_ctor_set(x_411, 2, x_409); -lean_ctor_set(x_411, 3, x_4); -lean_inc(x_284); -x_412 = l_Lean_Syntax_node1(x_284, x_362, x_411); -lean_inc(x_284); -x_413 = l_Lean_Syntax_node2(x_284, x_393, x_407, x_412); -lean_inc(x_284); -x_414 = l_Lean_Syntax_node3(x_284, x_397, x_377, x_413, x_396); -lean_inc(x_284); -x_415 = l_Lean_Syntax_node2(x_284, x_362, x_398, x_414); -lean_inc(x_284); -x_416 = l_Lean_Syntax_node2(x_284, x_393, x_375, x_415); -x_417 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__6; -lean_inc(x_284); -x_418 = l_Lean_Syntax_node4(x_284, x_417, x_365, x_364, x_367, x_416); -x_419 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__4; -lean_inc(x_284); -x_420 = l_Lean_Syntax_node1(x_284, x_419, x_418); -x_421 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -lean_inc(x_284); -x_422 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_422, 0, x_284); -lean_ctor_set(x_422, 1, x_421); -x_423 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__2; -x_424 = l_Lean_Syntax_node4(x_284, x_423, x_361, x_420, x_422, x_6); -x_425 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_425, 0, x_424); -lean_ctor_set(x_425, 1, x_357); -return x_425; +lean_ctor_set(x_405, 2, x_403); +lean_ctor_set(x_405, 3, x_3); +lean_inc(x_280); +x_406 = l_Lean_Syntax_node1(x_280, x_357, x_405); +lean_inc(x_280); +x_407 = l_Lean_Syntax_node2(x_280, x_387, x_401, x_406); +lean_inc(x_280); +x_408 = l_Lean_Syntax_node3(x_280, x_391, x_371, x_407, x_390); +lean_inc(x_280); +x_409 = l_Lean_Syntax_node2(x_280, x_357, x_392, x_408); +lean_inc(x_280); +x_410 = l_Lean_Syntax_node2(x_280, x_387, x_369, x_409); +x_411 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__6; +lean_inc(x_359); +lean_inc(x_280); +x_412 = l_Lean_Syntax_node5(x_280, x_411, x_5, x_359, x_359, x_361, x_410); +x_413 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__4; +lean_inc(x_280); +x_414 = l_Lean_Syntax_node1(x_280, x_413, x_412); +x_415 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +lean_inc(x_280); +x_416 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_416, 0, x_280); +lean_ctor_set(x_416, 1, x_415); +x_417 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__2; +x_418 = l_Lean_Syntax_node4(x_280, x_417, x_356, x_414, x_416, x_6); +x_419 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_419, 0, x_418); +lean_ctor_set(x_419, 1, x_352); +return x_419; } } } 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_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_dec(x_18); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_426 = l_Lean_MessageData_ofName(x_17); -x_427 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__15; +x_420 = l_Lean_MessageData_ofName(x_17); +x_421 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__15; +x_422 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_422, 0, x_421); +lean_ctor_set(x_422, 1, x_420); +x_423 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__17; +x_424 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_424, 0, x_422); +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__23___closed__5; +x_426 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_426, 0, x_425); +lean_ctor_set(x_426, 1, x_424); +x_427 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__11; x_428 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_428, 0, x_427); -lean_ctor_set(x_428, 1, x_426); -x_429 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__17; -x_430 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_430, 0, x_428); -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__23___closed__5; -x_432 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_432, 0, x_431); -lean_ctor_set(x_432, 1, x_430); -x_433 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__11; -x_434 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_434, 0, x_432); -lean_ctor_set(x_434, 1, x_433); -x_435 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__20(x_2, x_434, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_435; +lean_ctor_set(x_428, 0, x_426); +lean_ctor_set(x_428, 1, x_427); +x_429 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__20(x_2, x_428, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_429; } } 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_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_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_436 = l_Lean_MessageData_ofName(x_17); -x_437 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__15; +x_430 = l_Lean_MessageData_ofName(x_17); +x_431 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__15; +x_432 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_432, 0, x_431); +lean_ctor_set(x_432, 1, x_430); +x_433 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__17; +x_434 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_434, 0, x_432); +lean_ctor_set(x_434, 1, x_433); +x_435 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__5; +x_436 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_436, 0, x_435); +lean_ctor_set(x_436, 1, x_434); +x_437 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__11; x_438 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_438, 0, x_437); -lean_ctor_set(x_438, 1, x_436); -x_439 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__17; -x_440 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_440, 0, x_438); -lean_ctor_set(x_440, 1, x_439); -x_441 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__5; -x_442 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_442, 0, x_441); -lean_ctor_set(x_442, 1, x_440); -x_443 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__11; -x_444 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_444, 0, x_442); -lean_ctor_set(x_444, 1, x_443); -x_445 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__20(x_2, x_444, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -return x_445; +lean_ctor_set(x_438, 0, x_436); +lean_ctor_set(x_438, 1, x_437); +x_439 = l_Lean_throwErrorAt___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__20(x_2, x_438, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_439; } } } @@ -27833,7 +27799,7 @@ x_16 = lean_st_ref_get(x_10, x_11); 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; lean_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_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; x_18 = lean_ctor_get(x_16, 0); x_19 = lean_ctor_get(x_18, 0); lean_inc(x_19); @@ -27851,173 +27817,169 @@ x_25 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_25, 0, x_14); lean_ctor_set(x_25, 1, x_23); lean_ctor_set(x_25, 2, x_24); -lean_inc(x_25); -lean_inc(x_14); -x_26 = l_Lean_Syntax_node2(x_14, x_23, x_1, x_25); -x_27 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__13; +x_26 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__13; lean_inc(x_14); -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_14); -lean_ctor_set(x_28, 1, x_27); -x_29 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__21; +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_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__21; 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_Lean_Elab_Term_Quotation_mkSyntaxQuotation___closed__40; +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_Elab_Term_Quotation_mkSyntaxQuotation___closed__40; lean_inc(x_15); lean_inc(x_20); -x_32 = l_Lean_addMacroScope(x_20, x_31, x_15); -x_33 = l_Lean_Elab_Term_Quotation_mkSyntaxQuotation___closed__41; -lean_inc(x_2); -x_34 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_34, 0, x_33); -lean_ctor_set(x_34, 1, x_2); -x_35 = l_Lean_Elab_Term_Quotation_mkSyntaxQuotation___closed__43; -lean_inc(x_2); +x_31 = l_Lean_addMacroScope(x_20, x_30, x_15); +x_32 = l_Lean_Elab_Term_Quotation_mkSyntaxQuotation___closed__41; +lean_inc(x_1); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_1); +x_34 = l_Lean_Elab_Term_Quotation_mkSyntaxQuotation___closed__43; +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 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_36, 0, x_35); -lean_ctor_set(x_36, 1, x_2); -x_37 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_37, 0, x_34); -lean_ctor_set(x_37, 1, x_36); -x_38 = l_Lean_Elab_Term_Quotation_mkSyntaxQuotation___closed__39; +lean_ctor_set(x_36, 0, x_33); +lean_ctor_set(x_36, 1, x_35); +x_37 = l_Lean_Elab_Term_Quotation_mkSyntaxQuotation___closed__39; 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_32); -lean_ctor_set(x_39, 3, x_37); -x_40 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__20; +x_38 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_38, 0, x_14); +lean_ctor_set(x_38, 1, x_37); +lean_ctor_set(x_38, 2, x_31); +lean_ctor_set(x_38, 3, x_36); +x_39 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__20; lean_inc(x_14); -x_41 = l_Lean_Syntax_node2(x_14, x_40, x_30, x_39); -x_42 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__22(x_3); -x_43 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; -x_44 = l_Lean_addMacroScope(x_20, x_43, x_15); -x_45 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; +x_40 = l_Lean_Syntax_node2(x_14, x_39, x_29, x_38); +x_41 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__22(x_2); +x_42 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; +x_43 = l_Lean_addMacroScope(x_20, x_42, x_15); +x_44 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; lean_inc(x_14); -x_46 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_46, 0, x_14); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_46, 2, x_44); -lean_ctor_set(x_46, 3, x_2); +x_45 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_45, 0, x_14); +lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_45, 2, x_43); +lean_ctor_set(x_45, 3, x_1); lean_inc(x_14); -x_47 = l_Lean_Syntax_node2(x_14, x_23, x_42, x_46); -x_48 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___closed__6; +x_46 = l_Lean_Syntax_node2(x_14, x_23, x_41, x_45); +x_47 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___closed__6; lean_inc(x_14); -x_49 = l_Lean_Syntax_node2(x_14, x_48, x_41, x_47); -x_50 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__6; +x_48 = l_Lean_Syntax_node2(x_14, x_47, x_40, x_46); +x_49 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__6; +lean_inc(x_25); lean_inc(x_14); -x_51 = l_Lean_Syntax_node4(x_14, x_50, x_26, x_25, x_28, x_49); -x_52 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__4; +x_50 = l_Lean_Syntax_node5(x_14, x_49, x_3, x_25, x_25, x_27, x_48); +x_51 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__4; lean_inc(x_14); -x_53 = l_Lean_Syntax_node1(x_14, x_52, x_51); -x_54 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_52 = l_Lean_Syntax_node1(x_14, x_51, x_50); +x_53 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; lean_inc(x_14); -x_55 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_55, 0, x_14); -lean_ctor_set(x_55, 1, x_54); -x_56 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__2; -x_57 = l_Lean_Syntax_node4(x_14, x_56, x_22, x_53, x_55, x_4); -lean_ctor_set(x_16, 0, x_57); +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_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__2; +x_56 = l_Lean_Syntax_node4(x_14, x_55, x_22, x_52, x_54, x_4); +lean_ctor_set(x_16, 0, x_56); return x_16; } 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; -x_58 = lean_ctor_get(x_16, 0); -x_59 = lean_ctor_get(x_16, 1); -lean_inc(x_59); +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; +x_57 = lean_ctor_get(x_16, 0); +x_58 = lean_ctor_get(x_16, 1); lean_inc(x_58); +lean_inc(x_57); lean_dec(x_16); -x_60 = lean_ctor_get(x_58, 0); -lean_inc(x_60); -lean_dec(x_58); -x_61 = lean_environment_main_module(x_60); -x_62 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; +x_59 = lean_ctor_get(x_57, 0); +lean_inc(x_59); +lean_dec(x_57); +x_60 = lean_environment_main_module(x_59); +x_61 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; lean_inc(x_14); -x_63 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_63, 0, x_14); -lean_ctor_set(x_63, 1, x_62); -x_64 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__11; -x_65 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__12; +x_62 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_62, 0, x_14); +lean_ctor_set(x_62, 1, x_61); +x_63 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__11; +x_64 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__12; lean_inc(x_14); -x_66 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_66, 0, x_14); -lean_ctor_set(x_66, 1, x_64); -lean_ctor_set(x_66, 2, x_65); -lean_inc(x_66); +x_65 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_65, 0, x_14); +lean_ctor_set(x_65, 1, x_63); +lean_ctor_set(x_65, 2, x_64); +x_66 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__13; lean_inc(x_14); -x_67 = l_Lean_Syntax_node2(x_14, x_64, x_1, x_66); -x_68 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__13; +x_67 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_67, 0, x_14); +lean_ctor_set(x_67, 1, x_66); +x_68 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__21; lean_inc(x_14); x_69 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_69, 0, x_14); lean_ctor_set(x_69, 1, x_68); -x_70 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__21; -lean_inc(x_14); -x_71 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_71, 0, x_14); -lean_ctor_set(x_71, 1, x_70); -x_72 = l_Lean_Elab_Term_Quotation_mkSyntaxQuotation___closed__40; +x_70 = l_Lean_Elab_Term_Quotation_mkSyntaxQuotation___closed__40; lean_inc(x_15); -lean_inc(x_61); -x_73 = l_Lean_addMacroScope(x_61, x_72, x_15); -x_74 = l_Lean_Elab_Term_Quotation_mkSyntaxQuotation___closed__41; -lean_inc(x_2); +lean_inc(x_60); +x_71 = l_Lean_addMacroScope(x_60, x_70, x_15); +x_72 = l_Lean_Elab_Term_Quotation_mkSyntaxQuotation___closed__41; +lean_inc(x_1); +x_73 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_73, 0, x_72); +lean_ctor_set(x_73, 1, x_1); +x_74 = l_Lean_Elab_Term_Quotation_mkSyntaxQuotation___closed__43; +lean_inc(x_1); x_75 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_75, 0, x_74); -lean_ctor_set(x_75, 1, x_2); -x_76 = l_Lean_Elab_Term_Quotation_mkSyntaxQuotation___closed__43; -lean_inc(x_2); -x_77 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_77, 0, x_76); -lean_ctor_set(x_77, 1, x_2); -x_78 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_78, 0, x_75); -lean_ctor_set(x_78, 1, x_77); -x_79 = l_Lean_Elab_Term_Quotation_mkSyntaxQuotation___closed__39; +lean_ctor_set(x_75, 1, x_1); +x_76 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_76, 0, x_73); +lean_ctor_set(x_76, 1, x_75); +x_77 = l_Lean_Elab_Term_Quotation_mkSyntaxQuotation___closed__39; lean_inc(x_14); -x_80 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_80, 0, x_14); -lean_ctor_set(x_80, 1, x_79); -lean_ctor_set(x_80, 2, x_73); -lean_ctor_set(x_80, 3, x_78); -x_81 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__20; +x_78 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_78, 0, x_14); +lean_ctor_set(x_78, 1, x_77); +lean_ctor_set(x_78, 2, x_71); +lean_ctor_set(x_78, 3, x_76); +x_79 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__23___closed__20; lean_inc(x_14); -x_82 = l_Lean_Syntax_node2(x_14, x_81, x_71, x_80); -x_83 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__22(x_3); -x_84 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; -x_85 = l_Lean_addMacroScope(x_61, x_84, x_15); -x_86 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; +x_80 = l_Lean_Syntax_node2(x_14, x_79, x_69, x_78); +x_81 = l___private_Init_Meta_0__Lean_quoteList___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_quoteSyntax___spec__22(x_2); +x_82 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; +x_83 = l_Lean_addMacroScope(x_60, x_82, x_15); +x_84 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; lean_inc(x_14); -x_87 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_87, 0, x_14); -lean_ctor_set(x_87, 1, x_86); -lean_ctor_set(x_87, 2, x_85); -lean_ctor_set(x_87, 3, x_2); +x_85 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_85, 0, x_14); +lean_ctor_set(x_85, 1, x_84); +lean_ctor_set(x_85, 2, x_83); +lean_ctor_set(x_85, 3, x_1); lean_inc(x_14); -x_88 = l_Lean_Syntax_node2(x_14, x_64, x_83, x_87); -x_89 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___closed__6; +x_86 = l_Lean_Syntax_node2(x_14, x_63, x_81, x_85); +x_87 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getSepStxFromSplice___closed__6; lean_inc(x_14); -x_90 = l_Lean_Syntax_node2(x_14, x_89, x_82, x_88); -x_91 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__6; +x_88 = l_Lean_Syntax_node2(x_14, x_87, x_80, x_86); +x_89 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__6; +lean_inc(x_65); lean_inc(x_14); -x_92 = l_Lean_Syntax_node4(x_14, x_91, x_67, x_66, x_69, x_90); -x_93 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__4; +x_90 = l_Lean_Syntax_node5(x_14, x_89, x_3, x_65, x_65, x_67, x_88); +x_91 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__4; lean_inc(x_14); -x_94 = l_Lean_Syntax_node1(x_14, x_93, x_92); -x_95 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_92 = l_Lean_Syntax_node1(x_14, x_91, x_90); +x_93 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; 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_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__2; -x_98 = l_Lean_Syntax_node4(x_14, x_97, x_63, x_94, x_96, x_4); -x_99 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_99, 0, x_98); -lean_ctor_set(x_99, 1, x_59); -return x_99; +x_94 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_94, 0, x_14); +lean_ctor_set(x_94, 1, x_93); +x_95 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__2; +x_96 = l_Lean_Syntax_node4(x_14, x_95, x_62, x_92, x_94, x_4); +x_97 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_97, 0, x_96); +lean_ctor_set(x_97, 1, x_58); +return x_97; } } } @@ -28036,7 +27998,7 @@ x_15 = lean_st_ref_get(x_9, x_10); 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; 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_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; x_17 = lean_ctor_get(x_15, 0); x_18 = lean_ctor_get(x_17, 0); lean_inc(x_18); @@ -28055,98 +28017,94 @@ x_25 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_25, 0, x_13); lean_ctor_set(x_25, 1, x_23); lean_ctor_set(x_25, 2, x_24); -lean_inc(x_25); -lean_inc(x_13); -x_26 = l_Lean_Syntax_node2(x_13, x_23, x_22, x_25); -x_27 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__13; +x_26 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__13; 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_List_mapM_loop___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 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; +x_27 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_27, 0, x_13); +lean_ctor_set(x_27, 1, x_26); +x_28 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; +x_29 = l_Lean_addMacroScope(x_19, x_28, x_14); +x_30 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; lean_inc(x_13); -x_32 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_32, 0, x_13); -lean_ctor_set(x_32, 1, x_31); -lean_ctor_set(x_32, 2, x_30); -lean_ctor_set(x_32, 3, x_2); -x_33 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__6; +x_31 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_31, 0, x_13); +lean_ctor_set(x_31, 1, x_30); +lean_ctor_set(x_31, 2, x_29); +lean_ctor_set(x_31, 3, x_2); +x_32 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__6; +lean_inc(x_25); lean_inc(x_13); -x_34 = l_Lean_Syntax_node4(x_13, x_33, x_26, x_25, x_28, x_32); -x_35 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__4; +x_33 = l_Lean_Syntax_node5(x_13, x_32, x_22, x_25, x_25, x_27, x_31); +x_34 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__4; lean_inc(x_13); -x_36 = l_Lean_Syntax_node1(x_13, x_35, x_34); -x_37 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_35 = l_Lean_Syntax_node1(x_13, x_34, x_33); +x_36 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; lean_inc(x_13); -x_38 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_38, 0, x_13); -lean_ctor_set(x_38, 1, x_37); -x_39 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__2; -x_40 = l_Lean_Syntax_node4(x_13, x_39, x_21, x_36, x_38, x_3); -lean_ctor_set(x_15, 0, x_40); +x_37 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_37, 0, x_13); +lean_ctor_set(x_37, 1, x_36); +x_38 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__2; +x_39 = l_Lean_Syntax_node4(x_13, x_38, x_21, x_35, x_37, x_3); +lean_ctor_set(x_15, 0, x_39); return x_15; } 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; 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_41 = lean_ctor_get(x_15, 0); -x_42 = lean_ctor_get(x_15, 1); -lean_inc(x_42); +lean_object* x_40; lean_object* x_41; 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_40 = lean_ctor_get(x_15, 0); +x_41 = lean_ctor_get(x_15, 1); lean_inc(x_41); +lean_inc(x_40); lean_dec(x_15); -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_getHeadInfo___spec__9___closed__1; -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_Lean_Syntax_getAntiquotTerm(x_1); -x_48 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__11; -x_49 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__12; +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_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; lean_inc(x_13); -x_50 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_50, 0, x_13); -lean_ctor_set(x_50, 1, x_48); -lean_ctor_set(x_50, 2, x_49); -lean_inc(x_50); +x_45 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_45, 0, x_13); +lean_ctor_set(x_45, 1, x_44); +x_46 = l_Lean_Syntax_getAntiquotTerm(x_1); +x_47 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__11; +x_48 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__12; lean_inc(x_13); -x_51 = l_Lean_Syntax_node2(x_13, x_48, x_47, x_50); -x_52 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__13; +x_49 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_49, 0, x_13); +lean_ctor_set(x_49, 1, x_47); +lean_ctor_set(x_49, 2, x_48); +x_50 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__13; 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 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; -x_55 = l_Lean_addMacroScope(x_44, x_54, x_14); -x_56 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; +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_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; +x_53 = l_Lean_addMacroScope(x_43, x_52, x_14); +x_54 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; lean_inc(x_13); -x_57 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_57, 0, x_13); -lean_ctor_set(x_57, 1, x_56); -lean_ctor_set(x_57, 2, x_55); -lean_ctor_set(x_57, 3, x_2); -x_58 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__6; +x_55 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_55, 0, x_13); +lean_ctor_set(x_55, 1, x_54); +lean_ctor_set(x_55, 2, x_53); +lean_ctor_set(x_55, 3, x_2); +x_56 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__6; +lean_inc(x_49); lean_inc(x_13); -x_59 = l_Lean_Syntax_node4(x_13, x_58, x_51, x_50, x_53, x_57); -x_60 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__4; +x_57 = l_Lean_Syntax_node5(x_13, x_56, x_46, x_49, x_49, x_51, x_55); +x_58 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__4; lean_inc(x_13); -x_61 = l_Lean_Syntax_node1(x_13, x_60, x_59); -x_62 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_59 = l_Lean_Syntax_node1(x_13, x_58, x_57); +x_60 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; lean_inc(x_13); -x_63 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_63, 0, x_13); -lean_ctor_set(x_63, 1, x_62); -x_64 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__2; -x_65 = l_Lean_Syntax_node4(x_13, x_64, x_46, x_61, x_63, x_3); -x_66 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_66, 0, x_65); -lean_ctor_set(x_66, 1, x_42); -return x_66; +x_61 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_61, 0, x_13); +lean_ctor_set(x_61, 1, x_60); +x_62 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__2; +x_63 = l_Lean_Syntax_node4(x_13, x_62, x_45, x_59, x_61, x_3); +x_64 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_64, 0, x_63); +lean_ctor_set(x_64, 1, x_41); +return x_64; } } } @@ -28311,9 +28269,9 @@ lean_object* x_146; lean_inc(x_115); lean_inc(x_3); x_146 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__28___boxed), 11, 3); -lean_closure_set(x_146, 0, x_117); -lean_closure_set(x_146, 1, x_3); -lean_closure_set(x_146, 2, x_115); +lean_closure_set(x_146, 0, x_3); +lean_closure_set(x_146, 1, x_115); +lean_closure_set(x_146, 2, x_117); x_123 = x_146; goto block_142; } @@ -28641,9 +28599,9 @@ lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean x_89 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__25___boxed), 13, 5); lean_closure_set(x_89, 0, x_65); lean_closure_set(x_89, 1, x_1); -lean_closure_set(x_89, 2, x_78); -lean_closure_set(x_89, 3, x_3); -lean_closure_set(x_89, 4, x_76); +lean_closure_set(x_89, 2, x_3); +lean_closure_set(x_89, 3, x_76); +lean_closure_set(x_89, 4, x_78); x_90 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__4), 2, 1); lean_closure_set(x_90, 0, x_89); x_91 = lean_box(0); @@ -29026,8 +28984,8 @@ lean_inc(x_54); lean_inc(x_52); x_55 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__3), 4, 3); lean_closure_set(x_55, 0, x_52); -lean_closure_set(x_55, 1, x_29); -lean_closure_set(x_55, 2, x_9); +lean_closure_set(x_55, 1, x_9); +lean_closure_set(x_55, 2, x_29); x_56 = !lean_is_exclusive(x_52); if (x_56 == 0) { @@ -29068,8 +29026,8 @@ lean_inc(x_64); lean_inc(x_61); x_65 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__3), 4, 3); lean_closure_set(x_65, 0, x_61); -lean_closure_set(x_65, 1, x_29); -lean_closure_set(x_65, 2, x_9); +lean_closure_set(x_65, 1, x_9); +lean_closure_set(x_65, 2, x_29); if (lean_is_exclusive(x_61)) { lean_ctor_release(x_61, 0); lean_ctor_release(x_61, 1); @@ -29133,8 +29091,8 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); x_73 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__1___boxed), 10, 2); -lean_closure_set(x_73, 0, x_14); -lean_closure_set(x_73, 1, x_9); +lean_closure_set(x_73, 0, x_9); +lean_closure_set(x_73, 1, x_14); x_74 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__4), 2, 1); lean_closure_set(x_74, 0, x_73); x_75 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__5___boxed), 10, 1); @@ -29406,8 +29364,8 @@ lean_inc(x_150); lean_inc(x_146); x_151 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__3), 4, 3); lean_closure_set(x_151, 0, x_146); -lean_closure_set(x_151, 1, x_123); -lean_closure_set(x_151, 2, x_9); +lean_closure_set(x_151, 1, x_9); +lean_closure_set(x_151, 2, x_123); if (lean_is_exclusive(x_146)) { lean_ctor_release(x_146, 0); lean_ctor_release(x_146, 1); @@ -29475,8 +29433,8 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); x_159 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__1___boxed), 10, 2); -lean_closure_set(x_159, 0, x_108); -lean_closure_set(x_159, 1, x_9); +lean_closure_set(x_159, 0, x_9); +lean_closure_set(x_159, 1, x_108); x_160 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__4), 2, 1); lean_closure_set(x_160, 0, x_159); x_161 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___lambda__5___boxed), 10, 1); @@ -33355,7 +33313,7 @@ lean_dec(x_11); x_83 = !lean_is_exclusive(x_82); if (x_83 == 0) { -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_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_84 = lean_ctor_get(x_82, 0); x_85 = lean_ctor_get(x_84, 0); lean_inc(x_85); @@ -33381,92 +33339,88 @@ x_94 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_94, 0, x_80); lean_ctor_set(x_94, 1, x_93); lean_ctor_set(x_94, 2, x_65); -lean_inc(x_94); -lean_inc(x_80); -x_95 = l_Lean_Syntax_node2(x_80, x_93, x_92, x_94); -x_96 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__13; +x_95 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__13; lean_inc(x_80); -x_97 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_97, 0, x_80); -lean_ctor_set(x_97, 1, x_96); -x_98 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__6; +x_96 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_96, 0, x_80); +lean_ctor_set(x_96, 1, x_95); +x_97 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__6; +lean_inc(x_94); lean_inc(x_80); -x_99 = l_Lean_Syntax_node4(x_80, x_98, x_95, x_94, x_97, x_39); -x_100 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__4; +x_98 = l_Lean_Syntax_node5(x_80, x_97, x_92, x_94, x_94, x_96, x_39); +x_99 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__4; lean_inc(x_80); -x_101 = l_Lean_Syntax_node1(x_80, x_100, x_99); -x_102 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_100 = l_Lean_Syntax_node1(x_80, x_99, x_98); +x_101 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; lean_inc(x_80); -x_103 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_103, 0, x_80); -lean_ctor_set(x_103, 1, x_102); -x_104 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__2; -x_105 = l_Lean_Syntax_node4(x_80, x_104, x_88, x_101, x_103, x_76); -lean_ctor_set(x_82, 0, x_105); +x_102 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_102, 0, x_80); +lean_ctor_set(x_102, 1, x_101); +x_103 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__2; +x_104 = l_Lean_Syntax_node4(x_80, x_103, x_88, x_100, x_102, x_76); +lean_ctor_set(x_82, 0, x_104); return x_82; } else { -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; -x_106 = lean_ctor_get(x_82, 0); -x_107 = lean_ctor_get(x_82, 1); -lean_inc(x_107); +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; +x_105 = lean_ctor_get(x_82, 0); +x_106 = lean_ctor_get(x_82, 1); lean_inc(x_106); +lean_inc(x_105); lean_dec(x_82); -x_108 = lean_ctor_get(x_106, 0); -lean_inc(x_108); -lean_dec(x_106); -x_109 = lean_environment_main_module(x_108); -x_110 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; -lean_inc(x_80); -x_111 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_111, 0, x_80); -lean_ctor_set(x_111, 1, x_110); -x_112 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; -x_113 = l_Lean_addMacroScope(x_109, x_112, x_81); -x_114 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; +x_107 = lean_ctor_get(x_105, 0); +lean_inc(x_107); +lean_dec(x_105); +x_108 = lean_environment_main_module(x_107); +x_109 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; lean_inc(x_80); -x_115 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_115, 0, x_80); -lean_ctor_set(x_115, 1, x_114); -lean_ctor_set(x_115, 2, x_113); -lean_ctor_set(x_115, 3, x_46); -x_116 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__11; +x_110 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_110, 0, x_80); +lean_ctor_set(x_110, 1, x_109); +x_111 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; +x_112 = l_Lean_addMacroScope(x_108, x_111, x_81); +x_113 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; lean_inc(x_80); -x_117 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_117, 0, x_80); -lean_ctor_set(x_117, 1, x_116); -lean_ctor_set(x_117, 2, x_65); -lean_inc(x_117); +x_114 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_114, 0, x_80); +lean_ctor_set(x_114, 1, x_113); +lean_ctor_set(x_114, 2, x_112); +lean_ctor_set(x_114, 3, x_46); +x_115 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__11; lean_inc(x_80); -x_118 = l_Lean_Syntax_node2(x_80, x_116, x_115, x_117); -x_119 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__13; +x_116 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_116, 0, x_80); +lean_ctor_set(x_116, 1, x_115); +lean_ctor_set(x_116, 2, x_65); +x_117 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__13; lean_inc(x_80); -x_120 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_120, 0, x_80); -lean_ctor_set(x_120, 1, x_119); -x_121 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__6; +x_118 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_118, 0, x_80); +lean_ctor_set(x_118, 1, x_117); +x_119 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__6; +lean_inc(x_116); lean_inc(x_80); -x_122 = l_Lean_Syntax_node4(x_80, x_121, x_118, x_117, x_120, x_39); -x_123 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__4; +x_120 = l_Lean_Syntax_node5(x_80, x_119, x_114, x_116, x_116, x_118, x_39); +x_121 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__4; lean_inc(x_80); -x_124 = l_Lean_Syntax_node1(x_80, x_123, x_122); -x_125 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +x_122 = l_Lean_Syntax_node1(x_80, x_121, x_120); +x_123 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; lean_inc(x_80); -x_126 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_126, 0, x_80); -lean_ctor_set(x_126, 1, x_125); -x_127 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__2; -x_128 = l_Lean_Syntax_node4(x_80, x_127, x_111, x_124, x_126, x_76); -x_129 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_129, 0, x_128); -lean_ctor_set(x_129, 1, x_107); -return x_129; +x_124 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_124, 0, x_80); +lean_ctor_set(x_124, 1, x_123); +x_125 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__2; +x_126 = l_Lean_Syntax_node4(x_80, x_125, x_110, x_122, x_124, x_76); +x_127 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_127, 0, x_126); +lean_ctor_set(x_127, 1, x_106); +return x_127; } } else { -uint8_t x_130; +uint8_t x_128; lean_dec(x_60); lean_dec(x_39); lean_dec(x_11); @@ -33475,29 +33429,29 @@ lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -x_130 = !lean_is_exclusive(x_69); -if (x_130 == 0) +x_128 = !lean_is_exclusive(x_69); +if (x_128 == 0) { return x_69; } else { -lean_object* x_131; lean_object* x_132; lean_object* x_133; -x_131 = lean_ctor_get(x_69, 0); -x_132 = lean_ctor_get(x_69, 1); -lean_inc(x_132); -lean_inc(x_131); +lean_object* x_129; lean_object* x_130; lean_object* x_131; +x_129 = lean_ctor_get(x_69, 0); +x_130 = lean_ctor_get(x_69, 1); +lean_inc(x_130); +lean_inc(x_129); lean_dec(x_69); -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; +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; } } } else { -uint8_t x_134; +uint8_t x_132; lean_dec(x_44); lean_free_object(x_1); lean_dec(x_40); @@ -33508,29 +33462,29 @@ lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -x_134 = !lean_is_exclusive(x_55); -if (x_134 == 0) +x_132 = !lean_is_exclusive(x_55); +if (x_132 == 0) { return x_55; } else { -lean_object* x_135; lean_object* x_136; lean_object* x_137; -x_135 = lean_ctor_get(x_55, 0); -x_136 = lean_ctor_get(x_55, 1); -lean_inc(x_136); -lean_inc(x_135); +lean_object* x_133; lean_object* x_134; lean_object* x_135; +x_133 = lean_ctor_get(x_55, 0); +x_134 = lean_ctor_get(x_55, 1); +lean_inc(x_134); +lean_inc(x_133); lean_dec(x_55); -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; +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_138; +uint8_t x_136; lean_dec(x_44); lean_free_object(x_2); lean_dec(x_41); @@ -33545,29 +33499,29 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); -x_138 = !lean_is_exclusive(x_47); -if (x_138 == 0) +x_136 = !lean_is_exclusive(x_47); +if (x_136 == 0) { return x_47; } else { -lean_object* x_139; lean_object* x_140; lean_object* x_141; -x_139 = lean_ctor_get(x_47, 0); -x_140 = lean_ctor_get(x_47, 1); -lean_inc(x_140); -lean_inc(x_139); +lean_object* x_137; lean_object* x_138; lean_object* x_139; +x_137 = lean_ctor_get(x_47, 0); +x_138 = lean_ctor_get(x_47, 1); +lean_inc(x_138); +lean_inc(x_137); lean_dec(x_47); -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; +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; } } } else { -uint8_t x_142; +uint8_t x_140; lean_free_object(x_2); lean_dec(x_42); lean_dec(x_41); @@ -33582,35 +33536,35 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); -x_142 = !lean_is_exclusive(x_43); -if (x_142 == 0) +x_140 = !lean_is_exclusive(x_43); +if (x_140 == 0) { return x_43; } else { -lean_object* x_143; lean_object* x_144; lean_object* x_145; -x_143 = lean_ctor_get(x_43, 0); -x_144 = lean_ctor_get(x_43, 1); -lean_inc(x_144); -lean_inc(x_143); +lean_object* x_141; lean_object* x_142; lean_object* x_143; +x_141 = lean_ctor_get(x_43, 0); +x_142 = lean_ctor_get(x_43, 1); +lean_inc(x_142); +lean_inc(x_141); lean_dec(x_43); -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; +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; } } } else { -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_1, 0); -x_147 = lean_ctor_get(x_1, 1); -x_148 = lean_ctor_get(x_2, 0); -x_149 = lean_ctor_get(x_2, 1); -lean_inc(x_149); -lean_inc(x_148); +lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; +x_144 = lean_ctor_get(x_1, 0); +x_145 = lean_ctor_get(x_1, 1); +x_146 = lean_ctor_get(x_2, 0); +x_147 = lean_ctor_get(x_2, 1); +lean_inc(x_147); +lean_inc(x_146); lean_dec(x_2); lean_inc(x_11); lean_inc(x_10); @@ -33618,275 +33572,273 @@ lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -lean_inc(x_148); -x_150 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo(x_148, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_150) == 0) +lean_inc(x_146); +x_148 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo(x_146, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_148) == 0) { -lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; -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_153 = lean_box(0); +lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; +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_box(0); 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_151); -x_154 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__2(x_151, x_149, x_153, x_6, x_7, x_8, x_9, x_10, x_11, x_152); -if (lean_obj_tag(x_154) == 0) +lean_inc(x_149); +x_152 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__2(x_149, x_147, x_151, x_6, x_7, x_8, x_9, x_10, x_11, x_150); +if (lean_obj_tag(x_152) == 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; lean_object* x_161; lean_object* x_162; lean_object* x_163; -x_155 = lean_ctor_get(x_154, 0); +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_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_155 = lean_ctor_get(x_149, 1); lean_inc(x_155); -x_156 = lean_ctor_get(x_154, 1); +x_156 = lean_ctor_get(x_149, 0); lean_inc(x_156); -lean_dec(x_154); -x_157 = lean_ctor_get(x_151, 1); -lean_inc(x_157); -x_158 = lean_ctor_get(x_151, 0); -lean_inc(x_158); -x_159 = lean_apply_1(x_157, 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_148); -x_161 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_161, 0, x_160); -lean_ctor_set(x_161, 1, x_155); -x_162 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__8; +x_157 = lean_apply_1(x_155, x_156); +x_158 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_158, 0, x_157); +lean_ctor_set(x_158, 1, x_146); +x_159 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_159, 0, x_158); +lean_ctor_set(x_159, 1, x_153); +x_160 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__8; 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_163 = l_List_forIn_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__6(x_3, x_4, x_3, x_4, x_161, x_162, x_6, x_7, x_8, x_9, x_10, x_11, x_156); +x_161 = l_List_forIn_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__6(x_3, x_4, x_3, x_4, x_159, x_160, x_6, x_7, x_8, x_9, x_10, x_11, x_154); lean_dec(x_4); lean_dec(x_3); -if (lean_obj_tag(x_163) == 0) +if (lean_obj_tag(x_161) == 0) { -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; -x_164 = lean_ctor_get(x_163, 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; lean_object* x_174; lean_object* x_175; +x_162 = lean_ctor_get(x_161, 0); +lean_inc(x_162); +x_163 = lean_ctor_get(x_162, 1); +lean_inc(x_163); +x_164 = lean_ctor_get(x_163, 1); lean_inc(x_164); -x_165 = lean_ctor_get(x_164, 1); +x_165 = lean_ctor_get(x_161, 1); lean_inc(x_165); -x_166 = lean_ctor_get(x_165, 1); +lean_dec(x_161); +x_166 = lean_ctor_get(x_162, 0); lean_inc(x_166); -x_167 = lean_ctor_get(x_163, 1); +lean_dec(x_162); +x_167 = lean_ctor_get(x_163, 0); lean_inc(x_167); lean_dec(x_163); x_168 = lean_ctor_get(x_164, 0); lean_inc(x_168); -lean_dec(x_164); -x_169 = lean_ctor_get(x_165, 0); +x_169 = lean_ctor_get(x_164, 1); lean_inc(x_169); -lean_dec(x_165); -x_170 = lean_ctor_get(x_166, 0); +lean_dec(x_164); +x_170 = lean_ctor_get(x_149, 2); lean_inc(x_170); -x_171 = lean_ctor_get(x_166, 1); -lean_inc(x_171); -lean_dec(x_166); -x_172 = lean_ctor_get(x_151, 2); -lean_inc(x_172); -lean_dec(x_151); -x_173 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__12; -lean_inc(x_147); -x_174 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__3), 13, 5); -lean_closure_set(x_174, 0, x_147); -lean_closure_set(x_174, 1, x_170); -lean_closure_set(x_174, 2, x_153); -lean_closure_set(x_174, 3, x_173); -lean_closure_set(x_174, 4, x_171); -lean_inc(x_146); -x_175 = lean_array_to_list(lean_box(0), x_169); -x_176 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4), 9, 2); -lean_closure_set(x_176, 0, x_1); -lean_closure_set(x_176, 1, x_175); +lean_dec(x_149); +x_171 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__12; +lean_inc(x_145); +x_172 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__3), 13, 5); +lean_closure_set(x_172, 0, x_145); +lean_closure_set(x_172, 1, x_168); +lean_closure_set(x_172, 2, x_151); +lean_closure_set(x_172, 3, x_171); +lean_closure_set(x_172, 4, x_169); +lean_inc(x_144); +x_173 = lean_array_to_list(lean_box(0), x_167); +x_174 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4), 9, 2); +lean_closure_set(x_174, 0, x_1); +lean_closure_set(x_174, 1, x_173); 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_177 = lean_apply_9(x_172, x_174, x_176, x_6, x_7, x_8, x_9, x_10, x_11, x_167); -if (lean_obj_tag(x_177) == 0) +x_175 = lean_apply_9(x_170, x_172, x_174, x_6, x_7, x_8, x_9, x_10, x_11, x_165); +if (lean_obj_tag(x_175) == 0) { -lean_object* x_178; lean_object* x_179; lean_object* x_180; size_t x_181; size_t x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; uint8_t 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_178 = lean_ctor_get(x_177, 0); -lean_inc(x_178); -x_179 = lean_ctor_get(x_177, 1); -lean_inc(x_179); -lean_dec(x_177); -x_180 = lean_array_get_size(x_168); -x_181 = lean_usize_of_nat(x_180); -lean_dec(x_180); -x_182 = 0; +lean_object* x_176; lean_object* x_177; lean_object* x_178; size_t x_179; size_t x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; uint8_t 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_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_array_get_size(x_166); +x_179 = lean_usize_of_nat(x_178); +lean_dec(x_178); +x_180 = 0; lean_inc(x_10); -x_183 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__11(x_168, x_181, x_182, x_178, x_6, x_7, x_8, x_9, x_10, x_11, x_179); +x_181 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__11(x_166, x_179, x_180, x_176, x_6, x_7, x_8, x_9, x_10, x_11, x_177); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_168); -x_184 = lean_ctor_get(x_183, 0); +lean_dec(x_166); +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_ctor_get(x_10, 5); 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_10, 5); -lean_inc(x_186); -x_187 = 0; -x_188 = l_Lean_SourceInfo_fromRef(x_186, x_187); -x_189 = lean_ctor_get(x_10, 10); -lean_inc(x_189); +x_185 = 0; +x_186 = l_Lean_SourceInfo_fromRef(x_184, x_185); +x_187 = lean_ctor_get(x_10, 10); +lean_inc(x_187); lean_dec(x_10); -x_190 = lean_st_ref_get(x_11, x_185); +x_188 = lean_st_ref_get(x_11, x_183); lean_dec(x_11); -x_191 = lean_ctor_get(x_190, 0); -lean_inc(x_191); -x_192 = lean_ctor_get(x_190, 1); -lean_inc(x_192); -if (lean_is_exclusive(x_190)) { - lean_ctor_release(x_190, 0); - lean_ctor_release(x_190, 1); - x_193 = x_190; +x_189 = lean_ctor_get(x_188, 0); +lean_inc(x_189); +x_190 = lean_ctor_get(x_188, 1); +lean_inc(x_190); +if (lean_is_exclusive(x_188)) { + lean_ctor_release(x_188, 0); + lean_ctor_release(x_188, 1); + x_191 = x_188; } else { - lean_dec_ref(x_190); - x_193 = lean_box(0); + lean_dec_ref(x_188); + x_191 = lean_box(0); } -x_194 = lean_ctor_get(x_191, 0); -lean_inc(x_194); -lean_dec(x_191); -x_195 = lean_environment_main_module(x_194); -x_196 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; -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_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; -x_199 = l_Lean_addMacroScope(x_195, x_198, x_189); -x_200 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; -lean_inc(x_188); -x_201 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_201, 0, x_188); +x_192 = lean_ctor_get(x_189, 0); +lean_inc(x_192); +lean_dec(x_189); +x_193 = lean_environment_main_module(x_192); +x_194 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; +lean_inc(x_186); +x_195 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_195, 0, x_186); +lean_ctor_set(x_195, 1, x_194); +x_196 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; +x_197 = l_Lean_addMacroScope(x_193, x_196, x_187); +x_198 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; +lean_inc(x_186); +x_199 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_199, 0, x_186); +lean_ctor_set(x_199, 1, x_198); +lean_ctor_set(x_199, 2, x_197); +lean_ctor_set(x_199, 3, x_151); +x_200 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__11; +lean_inc(x_186); +x_201 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_201, 0, x_186); lean_ctor_set(x_201, 1, x_200); -lean_ctor_set(x_201, 2, x_199); -lean_ctor_set(x_201, 3, x_153); -x_202 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__11; -lean_inc(x_188); -x_203 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_203, 0, x_188); +lean_ctor_set(x_201, 2, x_171); +x_202 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__13; +lean_inc(x_186); +x_203 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_203, 0, x_186); lean_ctor_set(x_203, 1, x_202); -lean_ctor_set(x_203, 2, x_173); -lean_inc(x_203); -lean_inc(x_188); -x_204 = l_Lean_Syntax_node2(x_188, x_202, x_201, x_203); -x_205 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__13; -lean_inc(x_188); -x_206 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_206, 0, x_188); -lean_ctor_set(x_206, 1, x_205); -x_207 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__6; -lean_inc(x_188); -x_208 = l_Lean_Syntax_node4(x_188, x_207, x_204, x_203, x_206, x_146); -x_209 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__4; -lean_inc(x_188); -x_210 = l_Lean_Syntax_node1(x_188, x_209, x_208); -x_211 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -lean_inc(x_188); -x_212 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_212, 0, x_188); -lean_ctor_set(x_212, 1, x_211); -x_213 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__2; -x_214 = l_Lean_Syntax_node4(x_188, x_213, x_197, x_210, x_212, x_184); -if (lean_is_scalar(x_193)) { - x_215 = lean_alloc_ctor(0, 2, 0); +x_204 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__6; +lean_inc(x_201); +lean_inc(x_186); +x_205 = l_Lean_Syntax_node5(x_186, x_204, x_199, x_201, x_201, x_203, x_144); +x_206 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__4; +lean_inc(x_186); +x_207 = l_Lean_Syntax_node1(x_186, x_206, x_205); +x_208 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +lean_inc(x_186); +x_209 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_209, 0, x_186); +lean_ctor_set(x_209, 1, x_208); +x_210 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__2; +x_211 = l_Lean_Syntax_node4(x_186, x_210, x_195, x_207, x_209, x_182); +if (lean_is_scalar(x_191)) { + x_212 = lean_alloc_ctor(0, 2, 0); } else { - x_215 = x_193; + x_212 = x_191; } -lean_ctor_set(x_215, 0, x_214); -lean_ctor_set(x_215, 1, x_192); -return x_215; +lean_ctor_set(x_212, 0, x_211); +lean_ctor_set(x_212, 1, x_190); +return x_212; } else { -lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; -lean_dec(x_168); -lean_dec(x_146); +lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; +lean_dec(x_166); +lean_dec(x_144); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -x_216 = lean_ctor_get(x_177, 0); -lean_inc(x_216); -x_217 = lean_ctor_get(x_177, 1); -lean_inc(x_217); -if (lean_is_exclusive(x_177)) { - lean_ctor_release(x_177, 0); - lean_ctor_release(x_177, 1); - x_218 = x_177; +x_213 = lean_ctor_get(x_175, 0); +lean_inc(x_213); +x_214 = lean_ctor_get(x_175, 1); +lean_inc(x_214); +if (lean_is_exclusive(x_175)) { + lean_ctor_release(x_175, 0); + lean_ctor_release(x_175, 1); + x_215 = x_175; } else { - lean_dec_ref(x_177); - x_218 = lean_box(0); + lean_dec_ref(x_175); + x_215 = lean_box(0); } -if (lean_is_scalar(x_218)) { - x_219 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_215)) { + x_216 = lean_alloc_ctor(1, 2, 0); } else { - x_219 = x_218; + x_216 = x_215; } -lean_ctor_set(x_219, 0, x_216); -lean_ctor_set(x_219, 1, x_217); -return x_219; +lean_ctor_set(x_216, 0, x_213); +lean_ctor_set(x_216, 1, x_214); +return x_216; } } else { -lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; -lean_dec(x_151); +lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; +lean_dec(x_149); lean_free_object(x_1); -lean_dec(x_147); -lean_dec(x_146); +lean_dec(x_145); +lean_dec(x_144); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -x_220 = lean_ctor_get(x_163, 0); -lean_inc(x_220); -x_221 = lean_ctor_get(x_163, 1); -lean_inc(x_221); -if (lean_is_exclusive(x_163)) { - lean_ctor_release(x_163, 0); - lean_ctor_release(x_163, 1); - x_222 = x_163; +x_217 = lean_ctor_get(x_161, 0); +lean_inc(x_217); +x_218 = lean_ctor_get(x_161, 1); +lean_inc(x_218); +if (lean_is_exclusive(x_161)) { + lean_ctor_release(x_161, 0); + lean_ctor_release(x_161, 1); + x_219 = x_161; } else { - lean_dec_ref(x_163); - x_222 = lean_box(0); + lean_dec_ref(x_161); + x_219 = lean_box(0); } -if (lean_is_scalar(x_222)) { - x_223 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_219)) { + x_220 = lean_alloc_ctor(1, 2, 0); } else { - x_223 = x_222; + x_220 = x_219; } -lean_ctor_set(x_223, 0, x_220); -lean_ctor_set(x_223, 1, x_221); -return x_223; +lean_ctor_set(x_220, 0, x_217); +lean_ctor_set(x_220, 1, x_218); +return x_220; } } else { -lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; -lean_dec(x_151); -lean_dec(x_148); -lean_free_object(x_1); -lean_dec(x_147); +lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; +lean_dec(x_149); lean_dec(x_146); +lean_free_object(x_1); +lean_dec(x_145); +lean_dec(x_144); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -33895,36 +33847,36 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); -x_224 = lean_ctor_get(x_154, 0); -lean_inc(x_224); -x_225 = lean_ctor_get(x_154, 1); -lean_inc(x_225); -if (lean_is_exclusive(x_154)) { - lean_ctor_release(x_154, 0); - lean_ctor_release(x_154, 1); - x_226 = x_154; +x_221 = lean_ctor_get(x_152, 0); +lean_inc(x_221); +x_222 = lean_ctor_get(x_152, 1); +lean_inc(x_222); +if (lean_is_exclusive(x_152)) { + lean_ctor_release(x_152, 0); + lean_ctor_release(x_152, 1); + x_223 = x_152; } else { - lean_dec_ref(x_154); - x_226 = lean_box(0); + lean_dec_ref(x_152); + x_223 = lean_box(0); } -if (lean_is_scalar(x_226)) { - x_227 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_223)) { + x_224 = lean_alloc_ctor(1, 2, 0); } else { - x_227 = x_226; + x_224 = x_223; } -lean_ctor_set(x_227, 0, x_224); -lean_ctor_set(x_227, 1, x_225); -return x_227; +lean_ctor_set(x_224, 0, x_221); +lean_ctor_set(x_224, 1, x_222); +return x_224; } } else { -lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; -lean_dec(x_149); -lean_dec(x_148); -lean_free_object(x_1); +lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_dec(x_147); lean_dec(x_146); +lean_free_object(x_1); +lean_dec(x_145); +lean_dec(x_144); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -33933,48 +33885,48 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); -x_228 = lean_ctor_get(x_150, 0); -lean_inc(x_228); -x_229 = lean_ctor_get(x_150, 1); -lean_inc(x_229); -if (lean_is_exclusive(x_150)) { - lean_ctor_release(x_150, 0); - lean_ctor_release(x_150, 1); - x_230 = x_150; +x_225 = lean_ctor_get(x_148, 0); +lean_inc(x_225); +x_226 = lean_ctor_get(x_148, 1); +lean_inc(x_226); +if (lean_is_exclusive(x_148)) { + lean_ctor_release(x_148, 0); + lean_ctor_release(x_148, 1); + x_227 = x_148; } else { - lean_dec_ref(x_150); - x_230 = lean_box(0); + lean_dec_ref(x_148); + x_227 = lean_box(0); } -if (lean_is_scalar(x_230)) { - x_231 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_227)) { + x_228 = lean_alloc_ctor(1, 2, 0); } else { - x_231 = x_230; + x_228 = x_227; } -lean_ctor_set(x_231, 0, x_228); -lean_ctor_set(x_231, 1, x_229); -return x_231; +lean_ctor_set(x_228, 0, x_225); +lean_ctor_set(x_228, 1, x_226); +return x_228; } } } 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; -x_232 = lean_ctor_get(x_1, 0); -x_233 = lean_ctor_get(x_1, 1); -lean_inc(x_233); -lean_inc(x_232); +lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; +x_229 = lean_ctor_get(x_1, 0); +x_230 = lean_ctor_get(x_1, 1); +lean_inc(x_230); +lean_inc(x_229); lean_dec(x_1); -x_234 = lean_ctor_get(x_2, 0); -lean_inc(x_234); -x_235 = lean_ctor_get(x_2, 1); -lean_inc(x_235); +x_231 = lean_ctor_get(x_2, 0); +lean_inc(x_231); +x_232 = lean_ctor_get(x_2, 1); +lean_inc(x_232); if (lean_is_exclusive(x_2)) { lean_ctor_release(x_2, 0); lean_ctor_release(x_2, 1); - x_236 = x_2; + x_233 = x_2; } else { lean_dec_ref(x_2); - x_236 = lean_box(0); + x_233 = lean_box(0); } lean_inc(x_11); lean_inc(x_10); @@ -33982,227 +33934,259 @@ lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -lean_inc(x_234); -x_237 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo(x_234, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_237) == 0) +lean_inc(x_231); +x_234 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo(x_231, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_234) == 0) { -lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; -x_238 = lean_ctor_get(x_237, 0); -lean_inc(x_238); -x_239 = lean_ctor_get(x_237, 1); -lean_inc(x_239); -lean_dec(x_237); -x_240 = lean_box(0); +lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; +x_235 = lean_ctor_get(x_234, 0); +lean_inc(x_235); +x_236 = lean_ctor_get(x_234, 1); +lean_inc(x_236); +lean_dec(x_234); +x_237 = lean_box(0); 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_238); -x_241 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__2(x_238, x_235, x_240, x_6, x_7, x_8, x_9, x_10, x_11, x_239); -if (lean_obj_tag(x_241) == 0) +lean_inc(x_235); +x_238 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__2(x_235, x_232, x_237, x_6, x_7, x_8, x_9, x_10, x_11, x_236); +if (lean_obj_tag(x_238) == 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_object* x_248; lean_object* x_249; lean_object* x_250; -x_242 = lean_ctor_get(x_241, 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; +x_239 = lean_ctor_get(x_238, 0); +lean_inc(x_239); +x_240 = lean_ctor_get(x_238, 1); +lean_inc(x_240); +lean_dec(x_238); +x_241 = lean_ctor_get(x_235, 1); +lean_inc(x_241); +x_242 = lean_ctor_get(x_235, 0); lean_inc(x_242); -x_243 = lean_ctor_get(x_241, 1); -lean_inc(x_243); -lean_dec(x_241); -x_244 = lean_ctor_get(x_238, 1); -lean_inc(x_244); -x_245 = lean_ctor_get(x_238, 0); -lean_inc(x_245); -x_246 = lean_apply_1(x_244, x_245); -x_247 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_247, 0, x_246); -lean_ctor_set(x_247, 1, x_234); -if (lean_is_scalar(x_236)) { - x_248 = lean_alloc_ctor(1, 2, 0); +x_243 = lean_apply_1(x_241, x_242); +x_244 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_244, 0, x_243); +lean_ctor_set(x_244, 1, x_231); +if (lean_is_scalar(x_233)) { + x_245 = lean_alloc_ctor(1, 2, 0); } else { - x_248 = x_236; + x_245 = x_233; } -lean_ctor_set(x_248, 0, x_247); -lean_ctor_set(x_248, 1, x_242); -x_249 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__8; +lean_ctor_set(x_245, 0, x_244); +lean_ctor_set(x_245, 1, x_239); +x_246 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__5___closed__8; 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_250 = l_List_forIn_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__6(x_3, x_4, x_3, x_4, x_248, x_249, x_6, x_7, x_8, x_9, x_10, x_11, x_243); +x_247 = l_List_forIn_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__6(x_3, x_4, x_3, x_4, x_245, x_246, x_6, x_7, x_8, x_9, x_10, x_11, x_240); lean_dec(x_4); lean_dec(x_3); -if (lean_obj_tag(x_250) == 0) +if (lean_obj_tag(x_247) == 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; -x_251 = lean_ctor_get(x_250, 0); +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_248 = lean_ctor_get(x_247, 0); +lean_inc(x_248); +x_249 = lean_ctor_get(x_248, 1); +lean_inc(x_249); +x_250 = lean_ctor_get(x_249, 1); +lean_inc(x_250); +x_251 = lean_ctor_get(x_247, 1); lean_inc(x_251); -x_252 = lean_ctor_get(x_251, 1); +lean_dec(x_247); +x_252 = lean_ctor_get(x_248, 0); lean_inc(x_252); -x_253 = lean_ctor_get(x_252, 1); +lean_dec(x_248); +x_253 = lean_ctor_get(x_249, 0); lean_inc(x_253); -x_254 = lean_ctor_get(x_250, 1); +lean_dec(x_249); +x_254 = lean_ctor_get(x_250, 0); lean_inc(x_254); -lean_dec(x_250); -x_255 = lean_ctor_get(x_251, 0); +x_255 = lean_ctor_get(x_250, 1); lean_inc(x_255); -lean_dec(x_251); -x_256 = lean_ctor_get(x_252, 0); +lean_dec(x_250); +x_256 = lean_ctor_get(x_235, 2); lean_inc(x_256); -lean_dec(x_252); -x_257 = lean_ctor_get(x_253, 0); -lean_inc(x_257); -x_258 = lean_ctor_get(x_253, 1); -lean_inc(x_258); -lean_dec(x_253); -x_259 = lean_ctor_get(x_238, 2); -lean_inc(x_259); -lean_dec(x_238); -x_260 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__12; -lean_inc(x_233); -x_261 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__3), 13, 5); -lean_closure_set(x_261, 0, x_233); -lean_closure_set(x_261, 1, x_257); -lean_closure_set(x_261, 2, x_240); -lean_closure_set(x_261, 3, x_260); -lean_closure_set(x_261, 4, x_258); -lean_inc(x_232); -x_262 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_262, 0, x_232); -lean_ctor_set(x_262, 1, x_233); -x_263 = lean_array_to_list(lean_box(0), x_256); -x_264 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4), 9, 2); -lean_closure_set(x_264, 0, x_262); -lean_closure_set(x_264, 1, x_263); +lean_dec(x_235); +x_257 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__12; +lean_inc(x_230); +x_258 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__3), 13, 5); +lean_closure_set(x_258, 0, x_230); +lean_closure_set(x_258, 1, x_254); +lean_closure_set(x_258, 2, x_237); +lean_closure_set(x_258, 3, x_257); +lean_closure_set(x_258, 4, x_255); +lean_inc(x_229); +x_259 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_259, 0, x_229); +lean_ctor_set(x_259, 1, x_230); +x_260 = lean_array_to_list(lean_box(0), x_253); +x_261 = lean_alloc_closure((void*)(l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___lambda__4), 9, 2); +lean_closure_set(x_261, 0, x_259); +lean_closure_set(x_261, 1, x_260); 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_265 = lean_apply_9(x_259, x_261, x_264, x_6, x_7, x_8, x_9, x_10, x_11, x_254); -if (lean_obj_tag(x_265) == 0) +x_262 = lean_apply_9(x_256, x_258, x_261, x_6, x_7, x_8, x_9, x_10, x_11, x_251); +if (lean_obj_tag(x_262) == 0) { -lean_object* x_266; lean_object* x_267; lean_object* x_268; size_t x_269; size_t x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; uint8_t 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; -x_266 = lean_ctor_get(x_265, 0); -lean_inc(x_266); -x_267 = lean_ctor_get(x_265, 1); -lean_inc(x_267); +lean_object* x_263; lean_object* x_264; lean_object* x_265; size_t x_266; size_t x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; uint8_t 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; +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); +x_265 = lean_array_get_size(x_252); +x_266 = lean_usize_of_nat(x_265); lean_dec(x_265); -x_268 = lean_array_get_size(x_255); -x_269 = lean_usize_of_nat(x_268); -lean_dec(x_268); -x_270 = 0; +x_267 = 0; lean_inc(x_10); -x_271 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__11(x_255, x_269, x_270, x_266, x_6, x_7, x_8, x_9, x_10, x_11, x_267); +x_268 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__11(x_252, x_266, x_267, x_263, x_6, x_7, x_8, x_9, x_10, x_11, x_264); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_255); -x_272 = lean_ctor_get(x_271, 0); -lean_inc(x_272); -x_273 = lean_ctor_get(x_271, 1); -lean_inc(x_273); -lean_dec(x_271); -x_274 = lean_ctor_get(x_10, 5); +lean_dec(x_252); +x_269 = lean_ctor_get(x_268, 0); +lean_inc(x_269); +x_270 = lean_ctor_get(x_268, 1); +lean_inc(x_270); +lean_dec(x_268); +x_271 = lean_ctor_get(x_10, 5); +lean_inc(x_271); +x_272 = 0; +x_273 = l_Lean_SourceInfo_fromRef(x_271, x_272); +x_274 = lean_ctor_get(x_10, 10); lean_inc(x_274); -x_275 = 0; -x_276 = l_Lean_SourceInfo_fromRef(x_274, x_275); -x_277 = lean_ctor_get(x_10, 10); -lean_inc(x_277); lean_dec(x_10); -x_278 = lean_st_ref_get(x_11, x_273); +x_275 = lean_st_ref_get(x_11, x_270); lean_dec(x_11); -x_279 = lean_ctor_get(x_278, 0); +x_276 = lean_ctor_get(x_275, 0); +lean_inc(x_276); +x_277 = lean_ctor_get(x_275, 1); +lean_inc(x_277); +if (lean_is_exclusive(x_275)) { + lean_ctor_release(x_275, 0); + lean_ctor_release(x_275, 1); + x_278 = x_275; +} else { + lean_dec_ref(x_275); + x_278 = lean_box(0); +} +x_279 = lean_ctor_get(x_276, 0); lean_inc(x_279); -x_280 = lean_ctor_get(x_278, 1); -lean_inc(x_280); -if (lean_is_exclusive(x_278)) { - lean_ctor_release(x_278, 0); - lean_ctor_release(x_278, 1); - x_281 = x_278; +lean_dec(x_276); +x_280 = lean_environment_main_module(x_279); +x_281 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; +lean_inc(x_273); +x_282 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_282, 0, x_273); +lean_ctor_set(x_282, 1, x_281); +x_283 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; +x_284 = l_Lean_addMacroScope(x_280, x_283, x_274); +x_285 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; +lean_inc(x_273); +x_286 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_286, 0, x_273); +lean_ctor_set(x_286, 1, x_285); +lean_ctor_set(x_286, 2, x_284); +lean_ctor_set(x_286, 3, x_237); +x_287 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__11; +lean_inc(x_273); +x_288 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_288, 0, x_273); +lean_ctor_set(x_288, 1, x_287); +lean_ctor_set(x_288, 2, x_257); +x_289 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__13; +lean_inc(x_273); +x_290 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_290, 0, x_273); +lean_ctor_set(x_290, 1, x_289); +x_291 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__6; +lean_inc(x_288); +lean_inc(x_273); +x_292 = l_Lean_Syntax_node5(x_273, x_291, x_286, x_288, x_288, x_290, x_229); +x_293 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__4; +lean_inc(x_273); +x_294 = l_Lean_Syntax_node1(x_273, x_293, x_292); +x_295 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; +lean_inc(x_273); +x_296 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_296, 0, x_273); +lean_ctor_set(x_296, 1, x_295); +x_297 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__2; +x_298 = l_Lean_Syntax_node4(x_273, x_297, x_282, x_294, x_296, x_269); +if (lean_is_scalar(x_278)) { + x_299 = lean_alloc_ctor(0, 2, 0); } else { - lean_dec_ref(x_278); - x_281 = lean_box(0); -} -x_282 = lean_ctor_get(x_279, 0); -lean_inc(x_282); -lean_dec(x_279); -x_283 = lean_environment_main_module(x_282); -x_284 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__1; -lean_inc(x_276); -x_285 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_285, 0, x_276); -lean_ctor_set(x_285, 1, x_284); -x_286 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__12; -x_287 = l_Lean_addMacroScope(x_283, x_286, x_277); -x_288 = l_List_mapM_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__4___closed__11; -lean_inc(x_276); -x_289 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_289, 0, x_276); -lean_ctor_set(x_289, 1, x_288); -lean_ctor_set(x_289, 2, x_287); -lean_ctor_set(x_289, 3, x_240); -x_290 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__11; -lean_inc(x_276); -x_291 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_291, 0, x_276); -lean_ctor_set(x_291, 1, x_290); -lean_ctor_set(x_291, 2, x_260); -lean_inc(x_291); -lean_inc(x_276); -x_292 = l_Lean_Syntax_node2(x_276, x_290, x_289, x_291); -x_293 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__13; -lean_inc(x_276); -x_294 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_294, 0, x_276); -lean_ctor_set(x_294, 1, x_293); -x_295 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__6; -lean_inc(x_276); -x_296 = l_Lean_Syntax_node4(x_276, x_295, x_292, x_291, x_294, x_232); -x_297 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__4; -lean_inc(x_276); -x_298 = l_Lean_Syntax_node1(x_276, x_297, x_296); -x_299 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__14; -lean_inc(x_276); -x_300 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_300, 0, x_276); -lean_ctor_set(x_300, 1, x_299); -x_301 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_getHeadInfo___spec__9___closed__2; -x_302 = l_Lean_Syntax_node4(x_276, x_301, x_285, x_298, x_300, x_272); -if (lean_is_scalar(x_281)) { - x_303 = lean_alloc_ctor(0, 2, 0); + x_299 = x_278; +} +lean_ctor_set(x_299, 0, x_298); +lean_ctor_set(x_299, 1, x_277); +return x_299; +} +else +{ +lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; +lean_dec(x_252); +lean_dec(x_229); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +x_300 = lean_ctor_get(x_262, 0); +lean_inc(x_300); +x_301 = lean_ctor_get(x_262, 1); +lean_inc(x_301); +if (lean_is_exclusive(x_262)) { + lean_ctor_release(x_262, 0); + lean_ctor_release(x_262, 1); + x_302 = x_262; } else { - x_303 = x_281; + lean_dec_ref(x_262); + x_302 = lean_box(0); } -lean_ctor_set(x_303, 0, x_302); -lean_ctor_set(x_303, 1, x_280); +if (lean_is_scalar(x_302)) { + x_303 = lean_alloc_ctor(1, 2, 0); +} else { + x_303 = x_302; +} +lean_ctor_set(x_303, 0, x_300); +lean_ctor_set(x_303, 1, x_301); return x_303; } +} else { lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; -lean_dec(x_255); -lean_dec(x_232); +lean_dec(x_235); +lean_dec(x_230); +lean_dec(x_229); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -x_304 = lean_ctor_get(x_265, 0); +x_304 = lean_ctor_get(x_247, 0); lean_inc(x_304); -x_305 = lean_ctor_get(x_265, 1); +x_305 = lean_ctor_get(x_247, 1); lean_inc(x_305); -if (lean_is_exclusive(x_265)) { - lean_ctor_release(x_265, 0); - lean_ctor_release(x_265, 1); - x_306 = x_265; +if (lean_is_exclusive(x_247)) { + lean_ctor_release(x_247, 0); + lean_ctor_release(x_247, 1); + x_306 = x_247; } else { - lean_dec_ref(x_265); + lean_dec_ref(x_247); x_306 = lean_box(0); } if (lean_is_scalar(x_306)) { @@ -34218,25 +34202,29 @@ return x_307; else { lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; -lean_dec(x_238); +lean_dec(x_235); lean_dec(x_233); -lean_dec(x_232); +lean_dec(x_231); +lean_dec(x_230); +lean_dec(x_229); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -x_308 = lean_ctor_get(x_250, 0); +lean_dec(x_4); +lean_dec(x_3); +x_308 = lean_ctor_get(x_238, 0); lean_inc(x_308); -x_309 = lean_ctor_get(x_250, 1); +x_309 = lean_ctor_get(x_238, 1); lean_inc(x_309); -if (lean_is_exclusive(x_250)) { - lean_ctor_release(x_250, 0); - lean_ctor_release(x_250, 1); - x_310 = x_250; +if (lean_is_exclusive(x_238)) { + lean_ctor_release(x_238, 0); + lean_ctor_release(x_238, 1); + x_310 = x_238; } else { - lean_dec_ref(x_250); + lean_dec_ref(x_238); x_310 = lean_box(0); } if (lean_is_scalar(x_310)) { @@ -34252,11 +34240,11 @@ return x_311; else { lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; -lean_dec(x_238); -lean_dec(x_236); -lean_dec(x_234); lean_dec(x_233); lean_dec(x_232); +lean_dec(x_231); +lean_dec(x_230); +lean_dec(x_229); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -34265,16 +34253,16 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); -x_312 = lean_ctor_get(x_241, 0); +x_312 = lean_ctor_get(x_234, 0); lean_inc(x_312); -x_313 = lean_ctor_get(x_241, 1); +x_313 = lean_ctor_get(x_234, 1); lean_inc(x_313); -if (lean_is_exclusive(x_241)) { - lean_ctor_release(x_241, 0); - lean_ctor_release(x_241, 1); - x_314 = x_241; +if (lean_is_exclusive(x_234)) { + lean_ctor_release(x_234, 0); + lean_ctor_release(x_234, 1); + x_314 = x_234; } else { - lean_dec_ref(x_241); + lean_dec_ref(x_234); x_314 = lean_box(0); } if (lean_is_scalar(x_314)) { @@ -34287,44 +34275,6 @@ lean_ctor_set(x_315, 1, x_313); return x_315; } } -else -{ -lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; -lean_dec(x_236); -lean_dec(x_235); -lean_dec(x_234); -lean_dec(x_233); -lean_dec(x_232); -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_3); -x_316 = lean_ctor_get(x_237, 0); -lean_inc(x_316); -x_317 = lean_ctor_get(x_237, 1); -lean_inc(x_317); -if (lean_is_exclusive(x_237)) { - lean_ctor_release(x_237, 0); - lean_ctor_release(x_237, 1); - x_318 = x_237; -} else { - lean_dec_ref(x_237); - x_318 = lean_box(0); -} -if (lean_is_scalar(x_318)) { - x_319 = lean_alloc_ctor(1, 2, 0); -} else { - x_319 = x_318; -} -lean_ctor_set(x_319, 0, x_316); -lean_ctor_set(x_319, 1, x_317); -return x_319; -} -} } } } @@ -38094,7 +38044,7 @@ x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__1() { +static lean_object* _init_l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__1() { _start: { lean_object* x_1; @@ -38102,77 +38052,77 @@ x_1 = lean_mk_string_from_bytes("initFn", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__2() { +static lean_object* _init_l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____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_8015____closed__5; -x_2 = l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__1; +x_2 = l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__3() { +static lean_object* _init_l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__2; +x_1 = l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__2; x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_8015____closed__7; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__4() { +static lean_object* _init_l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__3; +x_1 = l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__3; x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_floatOutAntiquotTerms___lambda__1___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__5() { +static lean_object* _init_l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__4; +x_1 = l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__4; x_2 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__6() { +static lean_object* _init_l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__5; +x_1 = l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__5; x_2 = l_Lean_Elab_Term_Quotation_commandElab__stx__quot_____closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__7() { +static lean_object* _init_l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__6; +x_1 = l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__6; x_2 = l___regBuiltin_Lean_Elab_Term_Quotation_elabQuot____x40_Lean_Elab_Quotation___hyg_8015____closed__12; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__8() { +static lean_object* _init_l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__7; -x_2 = lean_unsigned_to_nat(20075u); +x_1 = l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__7; +x_2 = lean_unsigned_to_nat(20064u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__9() { +static lean_object* _init_l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__9() { _start: { lean_object* x_1; @@ -38180,24 +38130,24 @@ x_1 = lean_mk_string_from_bytes("alt", 3); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__10() { +static lean_object* _init_l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__10() { _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__1; x_2 = l_List_forIn_loop___at___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___spec__6___closed__1; -x_3 = l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__9; +x_3 = l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__9; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; x_2 = l___private_Lean_Elab_Quotation_0__Lean_Elab_Term_Quotation_compileStxMatch___closed__1; x_3 = 0; -x_4 = l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__8; +x_4 = l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__8; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); if (lean_obj_tag(x_5) == 0) { @@ -38205,7 +38155,7 @@ lean_object* x_6; lean_object* x_7; uint8_t 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_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__10; +x_7 = l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__10; x_8 = 1; x_9 = l_Lean_registerTraceClass(x_7, x_8, x_4, x_6); if (lean_obj_tag(x_9) == 0) @@ -40002,27 +39952,27 @@ lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Quotation_elabNoErrorIfUnused res = l___regBuiltin_Lean_Elab_Term_Quotation_elabNoErrorIfUnused_declRange(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__1 = _init_l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__1(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__1); -l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__2 = _init_l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__2(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__2); -l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__3 = _init_l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__3(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__3); -l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__4 = _init_l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__4(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__4); -l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__5 = _init_l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__5(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__5); -l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__6 = _init_l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__6(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__6); -l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__7 = _init_l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__7(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__7); -l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__8 = _init_l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__8(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__8); -l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__9 = _init_l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__9(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__9); -l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__10 = _init_l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__10(); -lean_mark_persistent(l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075____closed__10); -res = l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20075_(lean_io_mk_world()); +l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__1 = _init_l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__1(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__1); +l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__2 = _init_l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__2(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__2); +l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__3 = _init_l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__3(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__3); +l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__4 = _init_l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__4(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__4); +l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__5 = _init_l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__5(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__5); +l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__6 = _init_l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__6(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__6); +l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__7 = _init_l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__7(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__7); +l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__8 = _init_l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__8(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__8); +l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__9 = _init_l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__9(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__9); +l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__10 = _init_l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__10(); +lean_mark_persistent(l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064____closed__10); +res = l_Lean_Elab_Term_Quotation_initFn____x40_Lean_Elab_Quotation___hyg_20064_(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 ad3f3ccb5eba..06f1e4ec01bb 100644 --- a/stage0/stdlib/Lean/Elab/Quotation/Precheck.c +++ b/stage0/stdlib/Lean/Elab/Quotation/Precheck.c @@ -15,7 +15,6 @@ extern "C" { #endif static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_precheckIdent___closed__5; lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); -uint8_t l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_Quotation_precheckIdent(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_mkPrecheckAttribute___closed__2; static lean_object* l_Lean_Elab_Term_Quotation_precheckTypeAscription___closed__1; @@ -183,6 +182,7 @@ lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); static lean_object* l_List_forIn_loop___at_Lean_Elab_Term_Quotation_precheckIdent___spec__11___closed__1; LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_Quotation_precheckIdent___spec__6___boxed(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_Elab_Term_Quotation_precheckIdent___spec__6___closed__4; +uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_Quotation_precheck___spec__8___rarg___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_Quotation_precheckApp___closed__1; static lean_object* l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_Quotation_precheckIdent___spec__8___closed__2; @@ -2803,7 +2803,7 @@ lean_object* x_9; lean_object* x_10; uint8_t x_11; x_9 = lean_ctor_get(x_6, 2); lean_inc(x_9); x_10 = l_Lean_Elab_Term_Quotation_runPrecheck___closed__1; -x_11 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_9, x_10); +x_11 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_9, x_10); if (x_11 == 0) { lean_object* x_12; lean_object* x_13; @@ -2825,7 +2825,7 @@ else { lean_object* x_14; uint8_t x_15; x_14 = l_Lean_Elab_Term_Quotation_runPrecheck___closed__2; -x_15 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_9, x_14); +x_15 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_9, x_14); lean_dec(x_9); if (x_15 == 0) { @@ -3836,7 +3836,7 @@ x_26 = lean_ctor_get(x_24, 1); lean_inc(x_26); lean_dec(x_24); x_27 = l_List_forIn_loop___at_Lean_Elab_Term_Quotation_precheckIdent___spec__11___closed__1; -x_28 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_23, x_27); +x_28 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_23, x_27); if (x_28 == 0) { lean_object* x_29; diff --git a/stage0/stdlib/Lean/Elab/Structure.c b/stage0/stdlib/Lean/Elab/Structure.c index cff17352a75b..cbbc66e8d936 100644 --- a/stage0/stdlib/Lean/Elab/Structure.c +++ b/stage0/stdlib/Lean/Elab/Structure.c @@ -16,7 +16,6 @@ extern "C" { static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkAuxConstructions___closed__5; 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*); lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); -uint8_t l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(lean_object*, lean_object*); lean_object* lean_string_utf8_extract(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabStructureView___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_collectLevelParamsInFVars___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -463,7 +462,6 @@ static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabFie LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabStructureView___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*); static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_withFields_go___rarg___closed__8; -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCtor___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___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___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_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_withFields_go___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -553,6 +551,7 @@ lean_object* l_Lean_Syntax_node2(lean_object*, lean_object*, lean_object*, lean_ static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_16319____closed__9; static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___closed__4; lean_object* l_Lean_mkRecOn___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_mkAuxConstructions___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabStructure___closed__8; static lean_object* l_Lean_getDocStringText___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__14___closed__1; static lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__20___closed__19; @@ -737,6 +736,7 @@ uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_580____closed__15; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabFieldTypeValue___lambda__1___boxed(lean_object*, lean_object*); lean_object* lean_environment_main_module(lean_object*); +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_instantiateLevelMVars___at_Lean_Elab_Command_accLevelAtCtor___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_Array_forInUnsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__11(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_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___lambda__6___closed__5; @@ -13907,7 +13907,7 @@ x_52 = lean_ctor_get(x_29, 0); lean_inc(x_52); lean_dec(x_29); x_53 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_withParents_go___rarg___closed__1; -x_54 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_19, x_53); +x_54 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_19, x_53); lean_dec(x_19); if (x_54 == 0) { @@ -14171,7 +14171,7 @@ x_124 = lean_ctor_get(x_101, 0); lean_inc(x_124); lean_dec(x_101); x_125 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_withParents_go___rarg___closed__1; -x_126 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_83, x_125); +x_126 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_83, x_125); lean_dec(x_83); if (x_126 == 0) { @@ -24519,7 +24519,7 @@ lean_inc(x_31); x_32 = lean_ctor_get(x_30, 1); lean_inc(x_32); lean_dec(x_30); -x_33 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_31, x_14, x_15, x_16, x_17, x_32); +x_33 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_31, x_14, x_15, x_16, x_17, x_32); x_34 = lean_ctor_get(x_33, 0); lean_inc(x_34); x_35 = lean_ctor_get(x_33, 1); @@ -24861,7 +24861,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_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_26, x_12, x_13, x_14, x_15, x_27); +x_28 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_26, x_12, x_13, x_14, x_15, x_27); if (x_10 == 0) { lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; @@ -29866,7 +29866,6 @@ lean_closure_set(x_15, 4, x_5); lean_closure_set(x_15, 5, x_9); lean_closure_set(x_15, 6, x_14); lean_closure_set(x_15, 7, x_7); -lean_inc(x_10); x_16 = l_Lean_Elab_Command_runTermElabM___rarg(x_15, x_10, x_11, x_12); if (lean_obj_tag(x_16) == 0) { @@ -29888,6 +29887,7 @@ lean_closure_set(x_22, 0, x_2); lean_closure_set(x_22, 1, x_17); x_23 = l_Lean_Elab_Command_runTermElabM___rarg(x_22, x_10, x_11, x_18); lean_dec(x_11); +lean_dec(x_10); return x_23; } else @@ -29903,6 +29903,7 @@ lean_closure_set(x_25, 0, x_2); lean_closure_set(x_25, 1, x_17); x_26 = l_Lean_Elab_Command_runTermElabM___rarg(x_25, x_10, x_11, x_18); lean_dec(x_11); +lean_dec(x_10); return x_26; } else @@ -29927,6 +29928,7 @@ lean_closure_set(x_32, 0, x_2); lean_closure_set(x_32, 1, x_17); x_33 = l_Lean_Elab_Command_runTermElabM___rarg(x_32, x_10, x_11, x_31); lean_dec(x_11); +lean_dec(x_10); return x_33; } else diff --git a/stage0/stdlib/Lean/Elab/Syntax.c b/stage0/stdlib/Lean/Elab/Syntax.c index 426373348e4a..cd7c8cace1d1 100644 --- a/stage0/stdlib/Lean/Elab/Syntax.c +++ b/stage0/stdlib/Lean/Elab/Syntax.c @@ -293,7 +293,6 @@ uint8_t l_Char_isWhitespace(uint32_t); static lean_object* l_Lean_Elab_Term_toParserDescr_process___closed__2; static lean_object* l_Lean_Elab_Term_toParserDescr_processAtom___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processAlias___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_addTrace___at_Lean_Elab_Command_elabCommand___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabDeclareSyntaxCat___closed__9; 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*); static lean_object* l_Lean_Elab_Term_toParserDescr_processAlias___lambda__2___closed__2; @@ -367,6 +366,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processNonReserved(lean_ static lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__22; lean_object* lean_st_ref_get(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*); +uint8_t l___private_Lean_Parser_Basic_0__Lean_Parser_beqLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8753_(uint8_t, uint8_t); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Command_elabDeclareSyntaxCat_declRange(lean_object*); static lean_object* l_Lean_Elab_Term_checkLeftRec___closed__2; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_toParserDescr_processAlias___spec__6___lambda__1___closed__10; @@ -384,6 +384,7 @@ LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabSyntax__ static lean_object* l_Lean_Elab_Command_isLocalAttrKind___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Command_elabDeclareSyntaxCat_declRange___closed__5; static lean_object* l_Lean_Elab_Term_ensureUnaryOutput___closed__20; +lean_object* l_Lean_addTrace___at_Lean_Elab_Command_elabCommand___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabSyntax___spec__12___boxed(lean_object*, lean_object*); lean_object* l_Lean_Syntax_isStrLit_x3f(lean_object*); static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__70; @@ -483,7 +484,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processSepBy___lambda__1 static lean_object* l_Lean_Elab_Command_elabSyntax___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Command_elabSyntaxAbbrev___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processAlias___lambda__2(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* l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_elabCommand___spec__7(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_append___rarg(lean_object*, lean_object*); lean_object* l_Lean_throwError___at_Lean_Elab_Command_expandDeclId___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Environment_contains(lean_object*, lean_object*); @@ -543,6 +543,7 @@ uint8_t l_Lean_Name_hasMacroScopes(lean_object*); static lean_object* l_Lean_Elab_Term_toParserDescr_processSepBy1___lambda__1___closed__3; static lean_object* l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_declareSyntaxCatQuotParser___closed__47; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_inferMacroRulesAltKind___spec__1___rarg(lean_object*); +lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__13(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabSyntaxAbbrev_declRange___closed__7; static lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_0__Lean_Elab_Term_mkParserSeq___spec__1___closed__20; static lean_object* l_Lean_Elab_Command_elabSyntax___lambda__4___closed__11; @@ -582,6 +583,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Comma LEAN_EXPORT uint8_t l_Lean_Elab_Command_checkRuleKind(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_Lean_Elab_Term_toParserDescr_processNonReserved___closed__16; +lean_object* l_List_forM___at_Lean_Elab_Command_elabCommand___spec__8(lean_object*, lean_object*, lean_object*, lean_object*); 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__34; lean_object* l_Lean_Syntax_node7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -670,11 +672,9 @@ static lean_object* l_Lean_Elab_Command_elabSyntax___lambda__4___closed__16; lean_object* lean_array_uget(lean_object*, size_t); static lean_object* l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__24; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_toParserDescr_processAlias___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*); -uint8_t l___private_Lean_Parser_Basic_0__Lean_Parser_beqLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8688_(uint8_t, uint8_t); static lean_object* l_Lean_Elab_Term_toParserDescr_processAtom___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_toParserDescr_processAlias___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_object* l_Lean_Elab_mkUnusedBaseName(lean_object*, lean_object*, lean_object*); -lean_object* l_List_forM___at_Lean_Elab_Command_elabCommand___spec__9(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_object* lean_io_error_to_string(lean_object*); static lean_object* l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__1; @@ -4998,7 +4998,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_toParserDescr_processAtom___lambda__1( lean_object* x_12; uint8_t x_54; uint8_t x_55; uint8_t x_56; x_54 = lean_ctor_get_uint8(x_3, sizeof(void*)*1 + 2); x_55 = 0; -x_56 = l___private_Lean_Parser_Basic_0__Lean_Parser_beqLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8688_(x_54, x_55); +x_56 = l___private_Lean_Parser_Basic_0__Lean_Parser_beqLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8753_(x_54, x_55); if (x_56 == 0) { uint8_t x_57; @@ -11054,7 +11054,7 @@ static lean_object* _init_l___private_Lean_Elab_Syntax_0__Lean_Elab_Command_decl _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("|", 1); +x_1 = lean_mk_string_from_bytes("| ", 2); return x_1; } } @@ -14120,7 +14120,7 @@ x_53 = lean_ctor_get(x_44, 1); lean_inc(x_53); lean_dec(x_44); x_54 = l_List_reverse___rarg(x_53); -x_55 = l_List_forM___at_Lean_Elab_Command_elabCommand___spec__9(x_54, x_2, x_3, x_52); +x_55 = l_List_forM___at_Lean_Elab_Command_elabCommand___spec__8(x_54, x_2, x_3, x_52); lean_dec(x_3); lean_dec(x_2); x_56 = !lean_is_exclusive(x_55); @@ -14182,7 +14182,7 @@ x_71 = lean_ctor_get(x_44, 1); lean_inc(x_71); lean_dec(x_44); x_72 = l_List_reverse___rarg(x_71); -x_73 = l_List_forM___at_Lean_Elab_Command_elabCommand___spec__9(x_72, x_2, x_3, x_70); +x_73 = l_List_forM___at_Lean_Elab_Command_elabCommand___spec__8(x_72, x_2, x_3, x_70); lean_dec(x_3); lean_dec(x_2); x_74 = lean_ctor_get(x_73, 1); @@ -14473,7 +14473,7 @@ x_53 = lean_ctor_get(x_44, 1); lean_inc(x_53); lean_dec(x_44); x_54 = l_List_reverse___rarg(x_53); -x_55 = l_List_forM___at_Lean_Elab_Command_elabCommand___spec__9(x_54, x_2, x_3, x_52); +x_55 = l_List_forM___at_Lean_Elab_Command_elabCommand___spec__8(x_54, x_2, x_3, x_52); lean_dec(x_3); lean_dec(x_2); x_56 = !lean_is_exclusive(x_55); @@ -14535,7 +14535,7 @@ x_71 = lean_ctor_get(x_44, 1); lean_inc(x_71); lean_dec(x_44); x_72 = l_List_reverse___rarg(x_71); -x_73 = l_List_forM___at_Lean_Elab_Command_elabCommand___spec__9(x_72, x_2, x_3, x_70); +x_73 = l_List_forM___at_Lean_Elab_Command_elabCommand___spec__8(x_72, x_2, x_3, x_70); lean_dec(x_3); lean_dec(x_2); x_74 = lean_ctor_get(x_73, 1); @@ -14755,7 +14755,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabSyntax___lambda__3(lean_object* { lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; x_6 = l_Lean_Elab_Command_elabSyntax___lambda__3___closed__1; -x_7 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_elabCommand___spec__7(x_6, x_3, x_4, x_5); +x_7 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__13(x_6, x_3, x_4, x_5); x_8 = lean_ctor_get(x_7, 0); lean_inc(x_8); x_9 = lean_unbox(x_8); @@ -14780,7 +14780,7 @@ lean_inc(x_13); lean_dec(x_7); lean_inc(x_2); x_14 = l_Lean_MessageData_ofSyntax(x_2); -x_15 = l_Lean_addTrace___at_Lean_Elab_Command_elabCommand___spec__8(x_6, x_14, x_3, x_4, x_13); +x_15 = l_Lean_addTrace___at_Lean_Elab_Command_elabCommand___spec__7(x_6, x_14, x_3, x_4, x_13); x_16 = lean_ctor_get(x_15, 1); lean_inc(x_16); lean_dec(x_15); @@ -15056,7 +15056,6 @@ x_30 = l_Lean_mkIdentFrom(x_22, x_28, x_29); x_31 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabSyntax___lambda__2___boxed), 10, 2); lean_closure_set(x_31, 0, x_4); lean_closure_set(x_31, 1, x_3); -lean_inc(x_14); x_32 = l_Lean_Elab_Command_runTermElabM___rarg(x_31, x_14, x_15, x_24); if (lean_obj_tag(x_32) == 0) { @@ -15669,7 +15668,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabSyntax___lambda__5(lean_object* { lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; x_16 = l_Lean_Elab_Command_elabSyntax___lambda__5___closed__3; -x_17 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_elabCommand___spec__7(x_16, x_13, x_14, x_15); +x_17 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__13(x_16, x_13, x_14, x_15); x_18 = lean_ctor_get(x_17, 0); lean_inc(x_18); x_19 = lean_unbox(x_18); @@ -15700,7 +15699,7 @@ x_27 = l_Lean_Elab_Term_elabParserName_x3f___closed__4; x_28 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_28, 0, x_26); lean_ctor_set(x_28, 1, x_27); -x_29 = l_Lean_addTrace___at_Lean_Elab_Command_elabCommand___spec__8(x_16, x_28, x_13, x_14, x_23); +x_29 = l_Lean_addTrace___at_Lean_Elab_Command_elabCommand___spec__7(x_16, x_28, x_13, x_14, x_23); x_30 = lean_ctor_get(x_29, 0); lean_inc(x_30); x_31 = lean_ctor_get(x_29, 1); @@ -15800,7 +15799,6 @@ lean_inc(x_2); x_16 = lean_alloc_closure((void*)(l_Lean_Elab_Term_addCategoryInfo), 9, 2); lean_closure_set(x_16, 0, x_1); lean_closure_set(x_16, 1, x_2); -lean_inc(x_13); x_17 = l_Lean_Elab_Command_liftTermElabM___rarg(x_16, x_13, x_14, x_15); if (lean_obj_tag(x_17) == 0) { @@ -17003,7 +17001,6 @@ lean_inc(x_18); lean_dec(x_16); x_19 = lean_alloc_closure((void*)(l_Lean_Elab_Command_elabSyntaxAbbrev___lambda__1___boxed), 9, 1); lean_closure_set(x_19, 0, x_18); -lean_inc(x_4); x_20 = l_Lean_Elab_Command_runTermElabM___rarg(x_19, x_4, x_5, x_6); if (lean_obj_tag(x_20) == 0) { diff --git a/stage0/stdlib/Lean/Elab/SyntheticMVars.c b/stage0/stdlib/Lean/Elab/SyntheticMVars.c index 3441a758e094..7e5ade0070be 100644 --- a/stage0/stdlib/Lean/Elab/SyntheticMVars.c +++ b/stage0/stdlib/Lean/Elab/SyntheticMVars.c @@ -251,7 +251,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_ static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_synthesizeSyntheticMVars_loop___spec__1___closed__2; LEAN_EXPORT lean_object* l_Lean_occursCheck_visitMVar___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_resumePostponed___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___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefaultPrio_synthesizePendingInstMVar_x27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Elab_Term_reportUnsolvedGoals(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Elab_Term_reportUnsolvedGoals(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_LocalContext_foldlM___at_Lean_Elab_Term_runTactic___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_Elab_Term_withSynthesize(lean_object*, lean_object*); static lean_object* l_List_filterAuxM___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep___spec__1___lambda__2___closed__3; @@ -17453,7 +17453,7 @@ lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_44 = l_Lean_Elab_Term_reportUnsolvedGoals(x_41, x_4, x_5, x_6, x_7, x_8, x_9, x_42); +x_44 = l_Lean_Elab_Term_reportUnsolvedGoals(x_41, x_6, x_7, x_8, x_9, x_42); if (lean_obj_tag(x_44) == 0) { lean_dec(x_13); @@ -17512,7 +17512,7 @@ lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_51 = l_Lean_Elab_Term_reportUnsolvedGoals(x_48, x_4, x_5, x_6, x_7, x_8, x_9, x_49); +x_51 = l_Lean_Elab_Term_reportUnsolvedGoals(x_48, x_6, x_7, x_8, x_9, x_49); if (lean_obj_tag(x_51) == 0) { lean_dec(x_13); diff --git a/stage0/stdlib/Lean/Elab/Tactic/Basic.c b/stage0/stdlib/Lean/Elab/Tactic/Basic.c index d3c33be3db42..267d3747d010 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Basic.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Basic.c @@ -13,24 +13,23 @@ #ifdef __cplusplus extern "C" { #endif +lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___spec__1(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_tagUntaggedGoals(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_evalTactic_handleEx___closed__1; -uint8_t l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_getMainModule(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkSorry(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_instAlternativeTacticM___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_Tactic_Basic_0__Lean_Elab_Tactic_TacticM_runCore_x27___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_Elab_Term_logUnassignedUsingErrorInfos(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_logException___at_Lean_Elab_Tactic_closeUsingOrAdmit___spec__1___closed__1; -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10108____closed__1; -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__11; -static lean_object* l_Lean_Elab_Tactic_evalTactic___lambda__3___closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_instMonadExceptExceptionTacticM; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_getCurrMacroScope___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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*); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__6(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_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_evalTactic___spec__10___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_Tactic_mkTacticAttribute___closed__9; static lean_object* l_Lean_Elab_Tactic_instAlternativeTacticM___lambda__1___closed__1; +extern lean_object* l_Lean_profiler; uint8_t l_Lean_Elab_isAbortExceptionId(lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_instMonadExceptExceptionTacticM___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -38,26 +37,29 @@ LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Elab_Tactic_evalTactic_handle LEAN_EXPORT lean_object* l_Lean_Elab_admitGoal___lambda__1(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_Tactic_evalTactic_expandEval___spec__1___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_instMonadTacticM___closed__4; +LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_evalTactic___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___private_Lean_Elab_Tactic_Basic_0__Lean_Elab_Tactic_TacticM_runCore_x27(lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__4(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_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalTactic_expandEval___spec__6(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_Tactic_withTacticInfoContext___spec__2___rarg___closed__1; -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__16; LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at_Lean_Elab_Tactic_pruneSolvedGoals___spec__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_Tactic_ensureHasNoMVars(lean_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_mkInitialTacticInfo___elambda__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_Elab_Tactic_liftMetaMAtMain(lean_object*); +lean_object* l_Lean_PersistentArray_toArray___rarg(lean_object*); lean_object* l_Lean_profileitIOUnsafe___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Elab_Tactic_Context_recover___default; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_instOrElseTacticM(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_SavedState_restore___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_evalTactic___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_Tactic_evalTactic___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_Lean_Elab_Tactic_evalTactic___lambda__5___closed__4; extern lean_object* l_Lean_maxRecDepthErrorMessage; static lean_object* l_Lean_Elab_Tactic_mkTacticAttribute___closed__11; lean_object* l_Lean_indentD(lean_object*); +double lean_float_div(double, double); static lean_object* l_Lean_Elab_Tactic_instAlternativeTacticM___closed__1; +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__5; static lean_object* l_Lean_Elab_Tactic_instAlternativeTacticM___lambda__1___closed__2; -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__12; static lean_object* l_Lean_Elab_logException___at_Lean_Elab_Tactic_closeUsingOrAdmit___spec__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalTacticAt(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_Tactic_evalTactic_throwExs___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -69,11 +71,12 @@ static lean_object* l_Lean_Elab_throwAbortTactic___at_Lean_Elab_Tactic_done___sp LEAN_EXPORT lean_object* l_Lean_logAt___at_Lean_Elab_Tactic_closeUsingOrAdmit___spec__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_EXPORT lean_object* l_Lean_Elab_withMacroExpansionInfo___at_Lean_Elab_Tactic_withMacroExpansion___spec__1___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_Elab_Tactic_evalTactic_handleEx(lean_object*, 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_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__10; static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalTactic_expandEval___spec__6___rarg___closed__1; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___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_object*); +double l_Lean_trace_profiler_threshold_getSecs(lean_object*); lean_object* l_Lean_Syntax_getId(lean_object*); -static lean_object* l_Lean_Elab_Tactic_evalTactic___lambda__3___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_tacticElabAttribute; -LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_evalTactic___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_Elab_Tactic_instMonadBacktrackSavedStateTacticM___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_liftMacroM___at_Lean_Elab_Tactic_evalTactic_expandEval___spec__1___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentArray_push___rarg(lean_object*, lean_object*); @@ -94,7 +97,6 @@ lean_object* l_Lean_Syntax_getArgs(lean_object*); LEAN_EXPORT 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_EXPORT lean_object* l_Lean_Elab_Tactic_withMacroExpansion___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* l_Lean_replaceRef(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_evalTactic___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_EXPORT lean_object* l_Lean_Elab_Tactic_getMainGoal_loop(lean_object*, lean_object*, 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_Elab_Tactic_saveTacticInfoForToken___spec__1___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_Syntax_getPos_x3f(lean_object*, uint8_t); @@ -102,17 +104,18 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_withoutRecover___rarg(lean_object*, lean_object* lean_array_fget(lean_object*, lean_object*); lean_object* l_Lean_Syntax_getTailPos_x3f(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Elab_withInfoTreeContext___at_Lean_Elab_Tactic_evalTactic_eval___spec__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_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__14; static lean_object* l_Lean_Elab_Tactic_instMonadExceptExceptionTacticM___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_getMainGoal(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_instMonadBacktrackSavedStateTacticM___closed__3; lean_object* l_ReaderT_instMonadReaderT___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_focus(lean_object*); +uint8_t lean_float_decLt(double, double); static lean_object* l_List_forIn_loop___at_Lean_Elab_Tactic_tagUntaggedGoals___spec__2___closed__1; static lean_object* l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Tactic_withTacticInfoContext___spec__2___rarg___closed__2; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_throwNoGoalsToBeSolved___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_Tactic_evalTactic___lambda__5___closed__2; lean_object* l_Lean_MessageData_hasSyntheticSorry(lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__9; LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at_Lean_Elab_Tactic_getMainTarget___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_Elab_Tactic_getMainTag(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_Tactic_evalTactic_handleEx___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -121,7 +124,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_tryTactic_x3f(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_orElse___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_Lean_logAt___at_Lean_Elab_Tactic_closeUsingOrAdmit___spec__2___closed__1; static lean_object* l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Tactic_withTacticInfoContext___spec__2___rarg___closed__3; -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__15; LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Elab_goalsToMessageData___spec__1(lean_object*, lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); lean_object* l_Lean_Expr_mvar___override(lean_object*); @@ -130,15 +132,17 @@ LEAN_EXPORT lean_object* l_Lean_Elab_withInfoTreeContext___at_Lean_Elab_Tactic_w LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Tactic_evalTactic_handleEx___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*); uint8_t l___private_Lean_Util_Trace_0__Lean_checkTraceOption(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MetavarContext_setMVarUserName(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__3___closed__1; static lean_object* l_Lean_Elab_Tactic_mkTacticAttribute___closed__6; +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__13; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_SavedState_restore(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_throwAbortTactic___at_Lean_Elab_Tactic_done___spec__1(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_Elab_Tactic_saveTacticInfoForToken___spec__1(lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__13; uint8_t lean_string_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_tryTactic(lean_object*); lean_object* l_Lean_Exception_toMessageData(lean_object*); static lean_object* l_Lean_Elab_Tactic_withCaseRef___rarg___closed__1; +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__14; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_withCaseRef(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_mkTacticAttribute___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_run(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -146,7 +150,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_throwNoGoalsToBeSolved___rarg(lean_o LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_getMainModule___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_liftMetaTacticAux___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_Lean_Elab_Tactic_mkInitialTacticInfo___elambda__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_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_instMonadTacticM___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_Tactic_withCaseRef___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_instMonadBacktrackSavedStateTacticM___closed__2; @@ -171,17 +174,13 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_getNameOfIdent_x27(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_instMonadExceptExceptionTacticM___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_Lean_Elab_Tactic_getGoals___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_mkTacticAttribute(lean_object*); -LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Elab_Tactic_evalTactic___spec__5___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_EXPORT lean_object* l_Lean_Elab_Tactic_getUnsolvedGoals(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_focusAndDone___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_Tactic_instMonadExceptExceptionTacticM___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_instMonadTacticM___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_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__4; -lean_object* l_Lean_log___at_Lean_Elab_Term_exceptionToSorry___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_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__3; +static lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__4___closed__1; lean_object* lean_st_ref_take(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_getCurrMacroScope(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__10; LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Tactic_tagUntaggedGoals___spec__2___lambda__1___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_liftExcept___at_Lean_Elab_liftMacroM___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_withTacticInfoContext___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*); @@ -190,21 +189,26 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalTactic_expandEval___lambda__1(le LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_evalTactic_expandEval___spec__4(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_expandMacroImpl_x3f(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_evalTactic_throwExs___closed__2; -LEAN_EXPORT lean_object* l_List_forM___at_Lean_Elab_Term_reportUnsolvedGoals___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__3___closed__2; lean_object* l_Lean_ResolveName_resolveNamespace(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_liftMetaMAtMain___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_Lean_Elab_Tactic_evalTactic___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_Tactic_evalTactic___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_liftMacroM___at_Lean_Elab_Tactic_evalTactic_expandEval___spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_reportUnsolvedGoals___closed__2; static lean_object* l_Lean_Elab_Tactic_mkTacticAttribute___closed__5; lean_object* l_Lean_MessageData_ofSyntax(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_tryTactic_x3f___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_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___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_Lean_Elab_Tactic_instMonadBacktrackSavedStateTacticM___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_getGoals(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_getGoals___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_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985_(lean_object*); +static double l_Lean_withSeconds___at_Lean_Elab_Tactic_evalTactic___spec__5___closed__1; lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean_object*); lean_object* l_Lean_MessageLog_toList(lean_object*); LEAN_EXPORT lean_object* l_Lean_logAt___at_Lean_Elab_Tactic_closeUsingOrAdmit___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_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__4; +static lean_object* l_Lean_Elab_Tactic_evalTactic___lambda__4___closed__1; static lean_object* l_Lean_Elab_Tactic_mkTacticAttribute___closed__8; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_tagUntaggedGoals___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_withInfoTreeContext___at_Lean_Elab_Tactic_withMacroExpansion___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -214,29 +218,32 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_saveTacticInfoForToken(lean_object*, LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Elab_Tactic_evalTactic_handleEx___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* l_Lean_Syntax_getKind(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_focusAndDone(lean_object*); +lean_object* l_Lean_PersistentArray_append___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_mkTacticAttribute___closed__7; lean_object* l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_liftMetaTacticAux___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_Tactic_evalTactic_expandEval(lean_object*, lean_object*, 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_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__3___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_adaptExpander(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_instMonadTacticM; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_withMainContext___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_Tactic_evalTactic___lambda__5___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_instMonadBacktrackSavedStateTacticM; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalTactic___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_Elab_Tactic_evalTactic___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_EXPORT lean_object* l_Lean_Elab_withMacroExpansionInfo___at_Lean_Elab_Tactic_adaptExpander___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_withMacroExpansionInfo___at_Lean_Elab_Tactic_withMacroExpansion___spec__1(lean_object*); static lean_object* l_Lean_Elab_Tactic_instAlternativeTacticM___closed__3; LEAN_EXPORT lean_object* l_Lean_log___at_Lean_Elab_Tactic_closeUsingOrAdmit___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_st_ref_get(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__16; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_getMainTarget(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_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10108_(lean_object*); LEAN_EXPORT lean_object* l_Lean_log___at_Lean_Elab_Tactic_closeUsingOrAdmit___spec__3(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_List_forIn_loop___at_Lean_Elab_Tactic_tagUntaggedGoals___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_evalTactic___lambda__3___closed__2; uint8_t l_List_isEmpty___rarg(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Tactic_evalTactic___spec__4___rarg(lean_object*, lean_object*); static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_evalTactic_expandEval___spec__5___closed__1; lean_object* lean_st_mk_ref(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, double, lean_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_instAlternativeTacticM___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_liftMacroM___at_Lean_Elab_Tactic_evalTactic_expandEval___spec__1___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withMVarContextImp___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -246,38 +253,47 @@ lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_getMainDecl(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*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_orElse(lean_object*); +LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Elab_Tactic_evalTactic___spec__11(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalTactic_eval(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_io_mono_nanos_now(lean_object*); static lean_object* l_List_forIn_loop___at_Lean_Elab_Tactic_evalTactic_handleEx___spec__4___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_closeMainGoal___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_Tactic_instAlternativeTacticM___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_throwMaxRecDepthAt___at_Lean_Elab_Tactic_evalTactic_expandEval___spec__5___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalTactic_expandEval___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at_Lean_Elab_Tactic_getMainTarget___spec__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_List_forM___at_Lean_Elab_Term_reportUnsolvedGoals___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__2; +LEAN_EXPORT lean_object* l_List_forM___at_Lean_Elab_Term_reportUnsolvedGoals___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Core_withFreshMacroScope___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_evalTactic___spec__3(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_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__6; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___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*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalTactic___lambda__1___boxed(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_log___at_Lean_Meta_computeSynthOrder___spec__6(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Tactic_withTacticInfoContext___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_pruneSolvedGoals(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_evalTactic___lambda__5___closed__1; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__15; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalTactic_eval___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*); extern lean_object* l_Lean_Elab_macroAttribute; extern lean_object* l_Lean_warningAsError; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_getUnsolvedGoals___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_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__8; static lean_object* l_Lean_Elab_Tactic_mkTacticAttribute___closed__1; +uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Elab_Tactic_evalTactic_handleEx___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_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__12; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___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*); uint8_t l_Lean_MetavarContext_isAnonymousMVar(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_ensureHasNoMVars___closed__1; -static lean_object* l_Lean_Elab_Tactic_evalTactic___lambda__3___closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Tactic_evalTactic___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_Lean_Elab_Tactic_liftMetaTacticAux(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_evalTactic_expandEval___spec__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_isIdent(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Basic_0__Lean_Elab_Tactic_TacticM_runCore___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_Tactic_run___lambda__1___closed__1; -LEAN_EXPORT lean_object* l_Lean_Elab_Term_reportUnsolvedGoals(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_reportUnsolvedGoals(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at_Lean_Elab_Tactic_closeMainGoal___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_Lean_Elab_Tactic_ensureHasNoMVars___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_evalTactic_expandEval___spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -289,37 +305,46 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_tryCatch(lean_object*); lean_object* l_Lean_Elab_mkElabAttribute___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_PersistentHashMap_contains___at_Lean_MVarId_isAssigned___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_evalTactic___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Tactic_evalTactic___spec__4___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10028_(lean_object*); uint8_t l_Lean_Environment_contains(lean_object*, lean_object*); extern lean_object* l_Lean_inheritedTraceOptions; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalTactic___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_Tactic_throwNoGoalsToBeSolved(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_focusAndDone___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalTactic_expandEval___spec__6___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_pruneSolvedGoals___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_name_append_index_after(lean_object*, lean_object*); LEAN_EXPORT 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_Lean_Elab_Tactic_getNameOfIdent_x27___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Elab_Tactic_evalTactic___spec__5___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*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_closeUsingOrAdmit(lean_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_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Tactic_evalTactic___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___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_EXPORT lean_object* l_Lean_MVarId_withContext___at_Lean_Elab_Tactic_withMainContext___spec__1(lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_evalTactic___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_Lean_Elab_Tactic_throwNoGoalsToBeSolved___rarg___closed__1; LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Tactic_tagUntaggedGoals___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_withTacticInfoContext(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_closeMainGoal(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_liftMacroM___at_Lean_Elab_Tactic_evalTactic_expandEval___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__8; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Tactic_evalTactic_throwExs___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_EXPORT lean_object* l_Lean_Elab_Tactic_mkTacticInfo___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_throwMaxRecDepthAt___at_Lean_Elab_Tactic_evalTactic_expandEval___spec__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_Tactic_evalTactic___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_Lean_Elab_Tactic_evalTactic___lambda__3___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_getCurrMacroScope___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Tactic_evalTactic_expandEval___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_Tactic_evalTactic___lambda__4___closed__3; static lean_object* l_Lean_Elab_Term_reportUnsolvedGoals___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_evalTactic_expandEval___spec__1___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__3; lean_object* l_Lean_MVarId_getDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_evalTactic___spec__9(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_throwNoGoalsToBeSolved___rarg___closed__2; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); lean_object* lean_environment_main_module(lean_object*); +lean_object* lean_float_to_string(double); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Tactic_evalTactic___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_instMonadTacticM___closed__2; lean_object* l_Lean_PersistentHashMap_insert___at_Lean_MVarId_assign___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalTactic_expandEval___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*); @@ -329,8 +354,6 @@ LEAN_EXPORT lean_object* l_List_forM___at_Lean_Elab_Tactic_evalTactic_expandEval lean_object* l_Lean_indentExpr(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_evalTactic_throwExs___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_Lean_addTrace___at_Lean_Elab_Tactic_evalTactic_handleEx___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_profileitM___at_Lean_Elab_Tactic_evalTactic___spec__5(lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_evalTactic___spec__4(lean_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_logException___at_Lean_Elab_Tactic_closeUsingOrAdmit___spec__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_Tactic_withMacroExpansion(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Tactic_evalTactic___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -339,30 +362,37 @@ LEAN_EXPORT lean_object* l_Lean_Elab_withMacroExpansionInfo___at_Lean_Elab_Tacti static lean_object* l_Lean_Elab_Term_reportUnsolvedGoals___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_closeMainGoal___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_focus___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_MonadExcept_ofExcept___at_Lean_Elab_Tactic_evalTactic___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* 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_Tactic_withMacroExpansion___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_Elab_Tactic_appendGoals___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_isTracingEnabledFor___at_Lean_Elab_Tactic_evalTactic_handleEx___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_Tactic_evalTactic___lambda__4___closed__2; 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_Tactic_evalTactic___lambda__5___closed__3; +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__11; +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Tactic_evalTactic___spec__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_ResolveName_resolveGlobalName(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_nat_sub(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Elab_Tactic_evalTactic___spec__7(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*); +static lean_object* l_Lean_Elab_Tactic_evalTactic___lambda__4___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalTactic_handleEx___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_Tactic_tagUntaggedGoals___spec__2___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_withMainContext(lean_object*); +LEAN_EXPORT lean_object* l_Lean_withSeconds___at_Lean_Elab_Tactic_evalTactic___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_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__7; LEAN_EXPORT lean_object* l_Lean_Elab_throwAbortTactic___at_Lean_Elab_Tactic_done___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_throwError___at_Lean_Elab_Tactic_instAlternativeTacticM___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_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__7; static lean_object* l_Lean_Elab_Tactic_instOrElseTacticM___closed__1; static lean_object* l_Lean_Elab_Tactic_evalTactic_handleEx___closed__2; +static lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__3___closed__5; lean_object* l_List_reverse___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_tryCatch___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_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__6; 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_Tactic_mkInitialTacticInfo___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_throwError___at_Lean_Elab_Tactic_instAlternativeTacticM___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_liftMacroM___at_Lean_Elab_Tactic_evalTactic_expandEval___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_instantiateMVarsCore(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_reportUnsolvedGoals___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_liftMetaMAtMain___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_Tactic_evalTactic___lambda__1(lean_object*, lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); @@ -372,10 +402,12 @@ static lean_object* l_Lean_Elab_Tactic_instAlternativeTacticM___closed__2; static lean_object* l_Lean_Elab_Tactic_instMonadTacticM___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_instAlternativeTacticM; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_closeMainGoal___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_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10028____closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_done___boxed(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_Lean_throwError___at_Lean_Elab_Tactic_evalTactic___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_withInfoTreeContext___at_Lean_Elab_Tactic_withMacroExpansion___spec__2(lean_object*); +extern lean_object* l_Lean_trace_profiler; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_getMainModule___rarg(lean_object*, lean_object*); static lean_object* l_Lean_addTrace___at_Lean_Elab_Tactic_evalTactic_handleEx___spec__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_mkInitialTacticInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -387,6 +419,7 @@ LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at_Lean_Elab_Tactic_pruneSol LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_withCaseRef___rarg___lambda__1___boxed(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_Elab_Tactic_evalTactic___lambda__5___closed__7; +LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Elab_Tactic_evalTactic___spec__11___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*); lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_throwNoGoalsToBeSolved___spec__1(lean_object*); LEAN_EXPORT lean_object* l_List_filterAuxM___at_Lean_Elab_Tactic_pruneSolvedGoals___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -401,6 +434,7 @@ static lean_object* l_Lean_Elab_Term_reportUnsolvedGoals___closed__1; static lean_object* l_Lean_Elab_Tactic_mkTacticAttribute___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Tactic_withTacticInfoContext___spec__2___rarg___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_instMonadExceptExceptionTacticM___closed__2; +static lean_object* l_Lean_Elab_Tactic_evalTactic___lambda__4___closed__4; lean_object* lean_array_get_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalTactic_handleEx___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_Tactic_ensureHasNoMVars___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -409,15 +443,16 @@ uint8_t lean_nat_dec_le(lean_object*, lean_object*); extern lean_object* l_Lean_Elab_unsupportedSyntaxExceptionId; LEAN_EXPORT lean_object* l_Lean_Elab_admitGoal(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_evalTactic_handleEx___spec__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3(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_EXPORT lean_object* l_Lean_Elab_Tactic_setGoals(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_instMonadTacticM___closed__5; lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_liftMetaTactic1(lean_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_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__3___boxed(lean_object**); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_instMonadTacticM___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_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__9; -static lean_object* l_Lean_Elab_Tactic_evalTactic___lambda__3___closed__1; +static lean_object* l_Lean_Elab_Tactic_evalTactic___lambda__4___closed__6; uint8_t l___private_Lean_Message_0__Lean_beqMessageSeverity____x40_Lean_Message___hyg_120_(uint8_t, uint8_t); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__5; +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__1; 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_EXPORT lean_object* l_Lean_Elab_Tactic_saveState(lean_object*); static lean_object* l_Lean_Elab_Tactic_evalTactic___lambda__5___closed__6; @@ -429,10 +464,13 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_liftMetaTactic(lean_object*, lean_ob LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_done(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_run___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_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Tactic_withTacticInfoContext___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Elab_Tactic_evalTactic___spec__11___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_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__3___closed__4; static lean_object* l_Lean_Elab_Tactic_instMonadTacticM___closed__3; LEAN_EXPORT lean_object* l_Lean_MVarId_withContext___at_Lean_Elab_Tactic_run___spec__1(lean_object*); LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Tactic_tagUntaggedGoals___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_List_forIn_loop___at_Lean_Elab_Tactic_evalTactic_handleEx___spec__4___closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Elab_Tactic_evalTactic___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*); extern lean_object* l_Lean_Elab_abortTacticExceptionId; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_setGoals___boxed(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_liftMacroM___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*); @@ -628,70 +666,70 @@ lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l_List_forM___at_Lean_Elab_Term_reportUnsolvedGoals___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_List_forM___at_Lean_Elab_Term_reportUnsolvedGoals___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: { if (lean_obj_tag(x_1) == 0) { -lean_object* x_9; lean_object* x_10; -lean_dec(x_7); -lean_dec(x_6); +lean_object* x_7; lean_object* x_8; lean_dec(x_5); lean_dec(x_4); -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_dec(x_3); +lean_dec(x_2); +x_7 = lean_box(0); +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_7); +lean_ctor_set(x_8, 1, x_6); +return x_8; } else { -lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_11 = lean_ctor_get(x_1, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_1, 1); -lean_inc(x_12); +lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_9 = lean_ctor_get(x_1, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_1, 1); +lean_inc(x_10); lean_dec(x_1); -lean_inc(x_7); -lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_13 = l_Lean_Elab_admitGoal(x_11, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_13) == 0) +lean_inc(x_3); +lean_inc(x_2); +x_11 = l_Lean_Elab_admitGoal(x_9, x_2, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_11) == 0) { -lean_object* x_14; -x_14 = lean_ctor_get(x_13, 1); -lean_inc(x_14); -lean_dec(x_13); -x_1 = x_12; -x_8 = x_14; +lean_object* x_12; +x_12 = lean_ctor_get(x_11, 1); +lean_inc(x_12); +lean_dec(x_11); +x_1 = x_10; +x_6 = x_12; goto _start; } else { -uint8_t x_16; -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); +uint8_t x_14; +lean_dec(x_10); lean_dec(x_5); lean_dec(x_4); -x_16 = !lean_is_exclusive(x_13); -if (x_16 == 0) +lean_dec(x_3); +lean_dec(x_2); +x_14 = !lean_is_exclusive(x_11); +if (x_14 == 0) { -return x_13; +return x_11; } else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; -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(1, 2, 0); -lean_ctor_set(x_19, 0, x_17); -lean_ctor_set(x_19, 1, x_18); -return x_19; +lean_object* x_15; lean_object* x_16; lean_object* x_17; +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_alloc_ctor(1, 2, 0); +lean_ctor_set(x_17, 0, x_15); +lean_ctor_set(x_17, 1, x_16); +return x_17; } } } @@ -757,51 +795,35 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_reportUnsolvedGoals(lean_object* x_1, lean_object* x_2, lean_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_Term_reportUnsolvedGoals(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_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; lean_object* x_17; lean_object* x_18; lean_object* x_19; +lean_object* x_7; 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_15; lean_object* x_16; lean_object* x_17; lean_inc(x_1); -x_9 = l_Lean_Elab_goalsToMessageData(x_1); -x_10 = l_Lean_Elab_Term_reportUnsolvedGoals___closed__5; +x_7 = l_Lean_Elab_goalsToMessageData(x_1); +x_8 = l_Lean_Elab_Term_reportUnsolvedGoals___closed__5; +x_9 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_9, 0, x_8); +lean_ctor_set(x_9, 1, x_7); +x_10 = l_Lean_Elab_Term_reportUnsolvedGoals___closed__7; x_11 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_11, 0, x_10); -lean_ctor_set(x_11, 1, x_9); -x_12 = l_Lean_Elab_Term_reportUnsolvedGoals___closed__7; -x_13 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_13, 0, x_11); -lean_ctor_set(x_13, 1, x_12); -x_14 = l_Lean_Elab_Term_reportUnsolvedGoals___closed__3; -x_15 = lean_alloc_ctor(8, 2, 0); -lean_ctor_set(x_15, 0, x_14); -lean_ctor_set(x_15, 1, x_13); -x_16 = 2; -x_17 = l_Lean_log___at_Lean_Elab_Term_exceptionToSorry___spec__2(x_15, x_16, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -x_18 = lean_ctor_get(x_17, 1); -lean_inc(x_18); -lean_dec(x_17); -x_19 = l_List_forM___at_Lean_Elab_Term_reportUnsolvedGoals___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_18); -return x_19; -} -} -LEAN_EXPORT lean_object* l_List_forM___at_Lean_Elab_Term_reportUnsolvedGoals___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: -{ -lean_object* x_9; -x_9 = l_List_forM___at_Lean_Elab_Term_reportUnsolvedGoals___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_3); -lean_dec(x_2); -return x_9; -} -} -LEAN_EXPORT lean_object* l_Lean_Elab_Term_reportUnsolvedGoals___boxed(lean_object* x_1, lean_object* x_2, 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_Term_reportUnsolvedGoals(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_3); -lean_dec(x_2); -return x_9; +lean_ctor_set(x_11, 0, x_9); +lean_ctor_set(x_11, 1, x_10); +x_12 = l_Lean_Elab_Term_reportUnsolvedGoals___closed__3; +x_13 = lean_alloc_ctor(8, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_11); +x_14 = 2; +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +x_15 = l_Lean_log___at_Lean_Meta_computeSynthOrder___spec__6(x_13, x_14, x_2, x_3, x_4, x_5, x_6); +x_16 = lean_ctor_get(x_15, 1); +lean_inc(x_16); +lean_dec(x_15); +x_17 = l_List_forM___at_Lean_Elab_Term_reportUnsolvedGoals___spec__1(x_1, x_2, x_3, x_4, x_5, x_16); +return x_17; } } static uint8_t _init_l_Lean_Elab_Tactic_Context_recover___default() { @@ -6136,148 +6158,1245 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_evalTactic___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, lean_object* x_7, 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_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Tactic_evalTactic___spec__4___rarg(lean_object* x_1, lean_object* x_2) { _start: { -uint8_t x_14; -x_14 = lean_usize_dec_eq(x_2, x_3); -if (x_14 == 0) -{ -lean_object* x_15; lean_object* x_16; +lean_object* x_3; 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_st_ref_get(x_1, x_2); +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 = lean_ctor_get(x_4, 3); +lean_inc(x_6); lean_dec(x_4); -x_15 = lean_array_uget(x_1, x_2); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); +x_7 = lean_st_ref_take(x_1, x_5); +x_8 = lean_ctor_get(x_7, 0); lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_16 = l_Lean_Elab_Tactic_evalTactic(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_16) == 0) -{ -lean_object* x_17; lean_object* x_18; size_t x_19; size_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); -x_19 = 1; -x_20 = lean_usize_add(x_2, x_19); -x_2 = x_20; -x_4 = x_17; -x_13 = x_18; -goto _start; -} -else +x_9 = lean_ctor_get(x_7, 1); +lean_inc(x_9); +lean_dec(x_7); +x_10 = !lean_is_exclusive(x_8); +if (x_10 == 0) { -uint8_t x_22; -lean_dec(x_12); +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_11 = lean_ctor_get(x_8, 3); 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_22 = !lean_is_exclusive(x_16); -if (x_22 == 0) +x_12 = l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Tactic_withTacticInfoContext___spec__2___rarg___closed__3; +lean_ctor_set(x_8, 3, x_12); +x_13 = lean_st_ref_set(x_1, x_8, x_9); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) { -return x_16; +lean_object* x_15; +x_15 = lean_ctor_get(x_13, 0); +lean_dec(x_15); +lean_ctor_set(x_13, 0, x_6); +return x_13; } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_23 = lean_ctor_get(x_16, 0); -x_24 = lean_ctor_get(x_16, 1); -lean_inc(x_24); -lean_inc(x_23); -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; -} +lean_object* x_16; lean_object* x_17; +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); +lean_dec(x_13); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_6); +lean_ctor_set(x_17, 1, x_16); +return x_17; } } else { -lean_object* x_26; -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); +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_18 = lean_ctor_get(x_8, 0); +x_19 = lean_ctor_get(x_8, 1); +x_20 = lean_ctor_get(x_8, 2); +x_21 = lean_ctor_get(x_8, 4); +x_22 = lean_ctor_get(x_8, 5); +x_23 = lean_ctor_get(x_8, 6); +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_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_13); -return x_26; +x_24 = l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Tactic_withTacticInfoContext___spec__2___rarg___closed__3; +x_25 = lean_alloc_ctor(0, 7, 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_24); +lean_ctor_set(x_25, 4, x_21); +lean_ctor_set(x_25, 5, x_22); +lean_ctor_set(x_25, 6, x_23); +x_26 = lean_st_ref_set(x_1, x_25, x_9); +x_27 = lean_ctor_get(x_26, 1); +lean_inc(x_27); +if (lean_is_exclusive(x_26)) { + lean_ctor_release(x_26, 0); + lean_ctor_release(x_26, 1); + x_28 = x_26; +} else { + lean_dec_ref(x_26); + x_28 = lean_box(0); } +if (lean_is_scalar(x_28)) { + x_29 = lean_alloc_ctor(0, 2, 0); +} else { + x_29 = x_28; } +lean_ctor_set(x_29, 0, x_6); +lean_ctor_set(x_29, 1, x_27); +return x_29; } -LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_evalTactic___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) { -_start: -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_11 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_evalTactic_expandEval___spec__5___closed__2; -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_1); -lean_ctor_set(x_12, 1, x_11); -x_13 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_13, 0, x_12); -lean_ctor_set(x_13, 1, x_10); -return x_13; } } -LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Elab_Tactic_evalTactic___spec__5___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_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Tactic_evalTactic___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) { _start: { -lean_object* x_14; lean_object* x_15; -x_14 = lean_apply_8(x_3, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -x_15 = l_Lean_profileitIOUnsafe___rarg(x_1, x_2, x_14, x_4, x_13); -return x_15; +lean_object* x_8; +x_8 = lean_alloc_closure((void*)(l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Tactic_evalTactic___spec__4___rarg___boxed), 2, 0); +return x_8; } } -LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Elab_Tactic_evalTactic___spec__5(lean_object* x_1) { +static double _init_l_Lean_withSeconds___at_Lean_Elab_Tactic_evalTactic___spec__5___closed__1() { _start: { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_profileitM___at_Lean_Elab_Tactic_evalTactic___spec__5___rarg___boxed), 13, 0); -return x_2; +lean_object* x_1; uint8_t x_2; lean_object* x_3; double x_4; +x_1 = lean_unsigned_to_nat(1000000000u); +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_Lean_Elab_Tactic_evalTactic___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_withSeconds___at_Lean_Elab_Tactic_evalTactic___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_4; lean_object* x_5; -x_4 = lean_box(0); -x_5 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_5, 0, x_4); -lean_ctor_set(x_5, 1, x_3); -return x_5; -} -} -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalTactic___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) { -_start: +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_11 = lean_io_mono_nanos_now(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_apply_9(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +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; -lean_dec(x_4); -x_14 = l_Lean_Elab_Tactic_saveState___rarg(x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; 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_addTrace___at_Lean_Elab_Tactic_evalTactic_handleEx___spec__2___closed__1; -x_18 = l_Lean_Elab_Tactic_evalTactic_expandEval(x_1, x_15, x_2, x_3, x_17, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_16); +x_17 = lean_io_mono_nanos_now(x_16); +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; uint8_t x_21; lean_object* x_22; double x_23; double x_24; double x_25; lean_object* x_26; lean_object* x_27; +x_19 = lean_ctor_get(x_17, 0); +x_20 = lean_nat_sub(x_19, x_12); +lean_dec(x_12); +lean_dec(x_19); +x_21 = 0; +x_22 = lean_unsigned_to_nat(0u); +x_23 = l_Float_ofScientific(x_20, x_21, x_22); +lean_dec(x_20); +x_24 = l_Lean_withSeconds___at_Lean_Elab_Tactic_evalTactic___spec__5___closed__1; +x_25 = lean_float_div(x_23, x_24); +x_26 = lean_box_float(x_25); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_15); +lean_ctor_set(x_27, 1, x_26); +lean_ctor_set(x_17, 0, x_27); +return x_17; +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; double x_33; double x_34; double x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_28 = lean_ctor_get(x_17, 0); +x_29 = lean_ctor_get(x_17, 1); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_17); +x_30 = lean_nat_sub(x_28, x_12); +lean_dec(x_12); +lean_dec(x_28); +x_31 = 0; +x_32 = lean_unsigned_to_nat(0u); +x_33 = l_Float_ofScientific(x_30, x_31, x_32); +lean_dec(x_30); +x_34 = l_Lean_withSeconds___at_Lean_Elab_Tactic_evalTactic___spec__5___closed__1; +x_35 = lean_float_div(x_33, x_34); +x_36 = lean_box_float(x_35); +x_37 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_37, 0, x_15); +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_29); +return x_38; +} +} +else +{ +uint8_t x_39; +lean_dec(x_12); +x_39 = !lean_is_exclusive(x_14); +if (x_39 == 0) +{ +return x_14; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_14, 0); +x_41 = lean_ctor_get(x_14, 1); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_14); +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_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Elab_Tactic_evalTactic___spec__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, 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; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_15 = lean_st_ref_take(x_13, x_14); +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_is_exclusive(x_16); +if (x_18 == 0) +{ +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; uint8_t x_29; +x_19 = lean_ctor_get(x_16, 3); +x_20 = l_Lean_PersistentArray_toArray___rarg(x_19); +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_Util_Trace_0__Lean_addTraceNodeCore___spec__1(x_22, x_23, x_20); +x_25 = lean_alloc_ctor(9, 3, 1); +lean_ctor_set(x_25, 0, x_2); +lean_ctor_set(x_25, 1, x_4); +lean_ctor_set(x_25, 2, x_24); +lean_ctor_set_uint8(x_25, sizeof(void*)*3, x_5); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_3); +lean_ctor_set(x_26, 1, x_25); +x_27 = l_Lean_PersistentArray_push___rarg(x_1, x_26); +lean_ctor_set(x_16, 3, x_27); +x_28 = lean_st_ref_set(x_13, x_16, x_17); +x_29 = !lean_is_exclusive(x_28); +if (x_29 == 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_28, 0, x_31); +return x_28; +} +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); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_32); +return x_34; +} +} +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; size_t x_44; size_t 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_35 = lean_ctor_get(x_16, 0); +x_36 = lean_ctor_get(x_16, 1); +x_37 = lean_ctor_get(x_16, 2); +x_38 = lean_ctor_get(x_16, 3); +x_39 = lean_ctor_get(x_16, 4); +x_40 = lean_ctor_get(x_16, 5); +x_41 = lean_ctor_get(x_16, 6); +lean_inc(x_41); +lean_inc(x_40); +lean_inc(x_39); +lean_inc(x_38); +lean_inc(x_37); +lean_inc(x_36); +lean_inc(x_35); +lean_dec(x_16); +x_42 = l_Lean_PersistentArray_toArray___rarg(x_38); +x_43 = lean_array_get_size(x_42); +x_44 = lean_usize_of_nat(x_43); +lean_dec(x_43); +x_45 = 0; +x_46 = l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___spec__1(x_44, x_45, x_42); +x_47 = lean_alloc_ctor(9, 3, 1); +lean_ctor_set(x_47, 0, x_2); +lean_ctor_set(x_47, 1, x_4); +lean_ctor_set(x_47, 2, x_46); +lean_ctor_set_uint8(x_47, sizeof(void*)*3, x_5); +x_48 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_48, 0, x_3); +lean_ctor_set(x_48, 1, x_47); +x_49 = l_Lean_PersistentArray_push___rarg(x_1, x_48); +x_50 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_50, 0, x_35); +lean_ctor_set(x_50, 1, x_36); +lean_ctor_set(x_50, 2, x_37); +lean_ctor_set(x_50, 3, x_49); +lean_ctor_set(x_50, 4, x_39); +lean_ctor_set(x_50, 5, x_40); +lean_ctor_set(x_50, 6, x_41); +x_51 = lean_st_ref_set(x_13, x_50, x_17); +x_52 = lean_ctor_get(x_51, 1); +lean_inc(x_52); +if (lean_is_exclusive(x_51)) { + lean_ctor_release(x_51, 0); + lean_ctor_release(x_51, 1); + x_53 = x_51; +} else { + lean_dec_ref(x_51); + x_53 = lean_box(0); +} +x_54 = lean_box(0); +if (lean_is_scalar(x_53)) { + x_55 = lean_alloc_ctor(0, 2, 0); +} else { + x_55 = x_53; +} +lean_ctor_set(x_55, 0, x_54); +lean_ctor_set(x_55, 1, x_52); +return x_55; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__6(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, lean_object* x_13, lean_object* x_14) { +_start: +{ +uint8_t x_15; +x_15 = !lean_is_exclusive(x_12); +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; +x_16 = lean_ctor_get(x_12, 5); +x_17 = l_Lean_replaceRef(x_3, x_16); +lean_dec(x_16); +lean_ctor_set(x_12, 5, x_17); +x_18 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_4, x_10, x_11, x_12, x_13, x_14); +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___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Elab_Tactic_evalTactic___spec__7(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_20); +lean_dec(x_12); +return x_21; +} +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; +x_22 = lean_ctor_get(x_12, 0); +x_23 = lean_ctor_get(x_12, 1); +x_24 = lean_ctor_get(x_12, 2); +x_25 = lean_ctor_get(x_12, 3); +x_26 = lean_ctor_get(x_12, 4); +x_27 = lean_ctor_get(x_12, 5); +x_28 = lean_ctor_get(x_12, 6); +x_29 = lean_ctor_get(x_12, 7); +x_30 = lean_ctor_get(x_12, 8); +x_31 = lean_ctor_get(x_12, 9); +x_32 = lean_ctor_get(x_12, 10); +lean_inc(x_32); +lean_inc(x_31); +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_dec(x_12); +x_33 = l_Lean_replaceRef(x_3, x_27); +lean_dec(x_27); +x_34 = lean_alloc_ctor(0, 11, 0); +lean_ctor_set(x_34, 0, x_22); +lean_ctor_set(x_34, 1, x_23); +lean_ctor_set(x_34, 2, x_24); +lean_ctor_set(x_34, 3, x_25); +lean_ctor_set(x_34, 4, x_26); +lean_ctor_set(x_34, 5, x_33); +lean_ctor_set(x_34, 6, x_28); +lean_ctor_set(x_34, 7, x_29); +lean_ctor_set(x_34, 8, x_30); +lean_ctor_set(x_34, 9, x_31); +lean_ctor_set(x_34, 10, x_32); +x_35 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_4, x_10, x_11, x_34, x_13, x_14); +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___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Elab_Tactic_evalTactic___spec__7(x_1, x_2, x_3, x_36, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_34, x_13, x_37); +lean_dec(x_34); +return x_38; +} +} +} +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Tactic_evalTactic___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: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_ctor_get(x_1, 0); +lean_inc(x_11); +x_12 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_10); +return x_12; +} +else +{ +lean_object* x_13; lean_object* x_14; +x_13 = lean_ctor_get(x_1, 0); +lean_inc(x_13); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_10); +return x_14; +} +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___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) { +_start: +{ +lean_object* x_11; +x_11 = lean_apply_9(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; +x_13 = lean_ctor_get(x_11, 0); +x_14 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_11, 0, x_14); +return x_11; +} +else +{ +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_alloc_ctor(1, 1, 0); +lean_ctor_set(x_17, 0, 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 +{ +uint8_t x_19; +x_19 = !lean_is_exclusive(x_11); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_11, 0); +x_21 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set_tag(x_11, 0); +lean_ctor_set(x_11, 0, x_21); +return x_11; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_22 = lean_ctor_get(x_11, 0); +x_23 = lean_ctor_get(x_11, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_11); +x_24 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_24, 0, x_22); +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; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___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, lean_object* x_15, lean_object* x_16) { +_start: +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +lean_inc(x_14); +x_17 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__6(x_1, x_2, x_3, x_6, x_4, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +x_18 = lean_ctor_get(x_17, 1); +lean_inc(x_18); +lean_dec(x_17); +x_19 = l_MonadExcept_ofExcept___at_Lean_Elab_Tactic_evalTactic___spec__8(x_5, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_18); +lean_dec(x_14); +lean_dec(x_10); +lean_dec(x_8); +return x_19; +} +} +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_profiler; +return x_1; +} +} +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__3___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("[", 1); +return x_1; +} +} +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__3___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__3___closed__2; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__3___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("s] ", 3); +return x_1; +} +} +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__3___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__3___closed__4; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, uint8_t x_7, double 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_dec(x_9); +x_19 = lean_ctor_get(x_16, 5); +lean_inc(x_19); +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); +lean_inc(x_10); +lean_inc(x_2); +x_20 = lean_apply_10(x_1, x_2, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +if (lean_obj_tag(x_20) == 0) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +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 = l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__3___closed__1; +x_24 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_6, x_23); +lean_dec(x_6); +if (x_24 == 0) +{ +if (x_7 == 0) +{ +lean_object* x_25; lean_object* x_26; +x_25 = lean_box(0); +x_26 = l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__2(x_3, x_4, x_19, x_5, x_2, x_21, x_25, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_22); +lean_dec(x_17); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_11); +lean_dec(x_2); +return x_26; +} +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; +x_27 = lean_float_to_string(x_8); +x_28 = lean_alloc_ctor(3, 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 = l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__3___closed__3; +x_31 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_29); +x_32 = l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__3___closed__5; +x_33 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +x_34 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_21); +x_35 = l_Lean_Elab_Term_reportUnsolvedGoals___closed__7; +x_36 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +x_37 = lean_box(0); +x_38 = l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__2(x_3, x_4, x_19, x_5, x_2, x_36, x_37, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_22); +lean_dec(x_17); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_11); +lean_dec(x_2); +return x_38; +} +} +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; +x_39 = lean_float_to_string(x_8); +x_40 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_40, 0, x_39); +x_41 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_41, 0, x_40); +x_42 = l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__3___closed__3; +x_43 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_41); +x_44 = l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__3___closed__5; +x_45 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +x_46 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_21); +x_47 = l_Lean_Elab_Term_reportUnsolvedGoals___closed__7; +x_48 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_48, 0, x_46); +lean_ctor_set(x_48, 1, x_47); +x_49 = lean_box(0); +x_50 = l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__2(x_3, x_4, x_19, x_5, x_2, x_48, x_49, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_22); +lean_dec(x_17); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_11); +lean_dec(x_2); +return x_50; +} +} +else +{ +uint8_t x_51; +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_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_51 = !lean_is_exclusive(x_20); +if (x_51 == 0) +{ +return x_20; +} +else +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_52 = lean_ctor_get(x_20, 0); +x_53 = lean_ctor_get(x_20, 1); +lean_inc(x_53); +lean_inc(x_52); +lean_dec(x_20); +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; +} +} +} +} +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__4___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_trace_profiler; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t 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, 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_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +lean_dec(x_7); +x_17 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Tactic_evalTactic___spec__4___rarg(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); +lean_dec(x_17); +x_20 = lean_alloc_closure((void*)(l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__1), 10, 1); +lean_closure_set(x_20, 0, x_1); +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_8); +x_21 = l_Lean_withSeconds___at_Lean_Elab_Tactic_evalTactic___spec__5(x_20, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, 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; uint8_t x_26; lean_object* x_71; uint8_t x_72; +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_22, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_22, 1); +lean_inc(x_25); +lean_dec(x_22); +x_71 = l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__4___closed__1; +x_72 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_5, x_71); +if (x_72 == 0) +{ +uint8_t x_73; +x_73 = 0; +x_26 = x_73; +goto block_70; +} +else +{ +double x_74; double x_75; uint8_t x_76; +x_74 = l_Lean_trace_profiler_threshold_getSecs(x_5); +x_75 = lean_unbox_float(x_25); +x_76 = lean_float_decLt(x_74, x_75); +x_26 = x_76; +goto block_70; +} +block_70: +{ +if (x_6 == 0) +{ +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; +lean_dec(x_25); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +x_27 = lean_st_ref_take(x_15, x_23); +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_is_exclusive(x_28); +if (x_30 == 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_28, 3); +x_32 = l_Lean_PersistentArray_append___rarg(x_18, x_31); +lean_ctor_set(x_28, 3, x_32); +x_33 = lean_st_ref_set(x_15, x_28, x_29); +x_34 = lean_ctor_get(x_33, 1); +lean_inc(x_34); +lean_dec(x_33); +x_35 = l_MonadExcept_ofExcept___at_Lean_Elab_Tactic_evalTactic___spec__8(x_24, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_34); +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); +lean_dec(x_24); +if (lean_obj_tag(x_35) == 0) +{ +uint8_t x_36; +x_36 = !lean_is_exclusive(x_35); +if (x_36 == 0) +{ +return x_35; +} +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; +x_40 = !lean_is_exclusive(x_35); +if (x_40 == 0) +{ +return x_35; +} +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_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; +} +} +} +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; 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_44 = lean_ctor_get(x_28, 0); +x_45 = lean_ctor_get(x_28, 1); +x_46 = lean_ctor_get(x_28, 2); +x_47 = lean_ctor_get(x_28, 3); +x_48 = lean_ctor_get(x_28, 4); +x_49 = lean_ctor_get(x_28, 5); +x_50 = lean_ctor_get(x_28, 6); +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_dec(x_28); +x_51 = l_Lean_PersistentArray_append___rarg(x_18, x_47); +x_52 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_52, 0, x_44); +lean_ctor_set(x_52, 1, x_45); +lean_ctor_set(x_52, 2, x_46); +lean_ctor_set(x_52, 3, x_51); +lean_ctor_set(x_52, 4, x_48); +lean_ctor_set(x_52, 5, x_49); +lean_ctor_set(x_52, 6, x_50); +x_53 = lean_st_ref_set(x_15, x_52, x_29); +x_54 = lean_ctor_get(x_53, 1); +lean_inc(x_54); +lean_dec(x_53); +x_55 = l_MonadExcept_ofExcept___at_Lean_Elab_Tactic_evalTactic___spec__8(x_24, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_54); +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); +lean_dec(x_24); +if (lean_obj_tag(x_55) == 0) +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +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; +} else { + lean_dec_ref(x_55); + x_58 = lean_box(0); +} +if (lean_is_scalar(x_58)) { + x_59 = lean_alloc_ctor(0, 2, 0); +} else { + x_59 = x_58; +} +lean_ctor_set(x_59, 0, x_56); +lean_ctor_set(x_59, 1, x_57); +return x_59; +} +else +{ +lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; +x_60 = lean_ctor_get(x_55, 0); +lean_inc(x_60); +x_61 = lean_ctor_get(x_55, 1); +lean_inc(x_61); +if (lean_is_exclusive(x_55)) { + lean_ctor_release(x_55, 0); + lean_ctor_release(x_55, 1); + x_62 = x_55; +} else { + lean_dec_ref(x_55); + x_62 = lean_box(0); +} +if (lean_is_scalar(x_62)) { + x_63 = lean_alloc_ctor(1, 2, 0); +} else { + x_63 = x_62; +} +lean_ctor_set(x_63, 0, x_60); +lean_ctor_set(x_63, 1, x_61); +return x_63; +} +} +} +else +{ +lean_object* x_64; double x_65; lean_object* x_66; +x_64 = lean_box(0); +x_65 = lean_unbox_float(x_25); +lean_dec(x_25); +x_66 = l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__3(x_2, x_24, x_18, x_3, x_4, x_5, x_26, x_65, x_64, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_23); +return x_66; +} +} +else +{ +lean_object* x_67; double x_68; lean_object* x_69; +x_67 = lean_box(0); +x_68 = lean_unbox_float(x_25); +lean_dec(x_25); +x_69 = l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__3(x_2, x_24, x_18, x_3, x_4, x_5, x_26, x_68, x_67, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_23); +return x_69; +} +} +} +else +{ +uint8_t x_77; +lean_dec(x_18); +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); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +x_77 = !lean_is_exclusive(x_21); +if (x_77 == 0) +{ +return x_21; +} +else +{ +lean_object* x_78; lean_object* x_79; lean_object* x_80; +x_78 = lean_ctor_get(x_21, 0); +x_79 = lean_ctor_get(x_21, 1); +lean_inc(x_79); +lean_inc(x_78); +lean_dec(x_21); +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_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__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) { +_start: +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_14 = lean_ctor_get(x_11, 2); +lean_inc(x_14); +lean_inc(x_1); +x_15 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_evalTactic_handleEx___spec__1(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_unbox(x_16); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_18 = lean_ctor_get(x_15, 1); +lean_inc(x_18); +lean_dec(x_15); +x_19 = l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__4___closed__1; +x_20 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_14, x_19); +if (x_20 == 0) +{ +lean_object* x_21; +lean_dec(x_16); +lean_dec(x_14); +lean_dec(x_2); +lean_dec(x_1); +x_21 = lean_apply_9(x_3, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_18); +if (lean_obj_tag(x_21) == 0) +{ +uint8_t x_22; +x_22 = !lean_is_exclusive(x_21); +if (x_22 == 0) +{ +return x_21; +} +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; +} +} +else +{ +uint8_t x_26; +x_26 = !lean_is_exclusive(x_21); +if (x_26 == 0) +{ +return x_21; +} +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_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 +{ +lean_object* x_30; uint8_t x_31; lean_object* x_32; +x_30 = lean_box(0); +x_31 = lean_unbox(x_16); +lean_dec(x_16); +x_32 = l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__4(x_3, x_2, x_1, x_4, x_14, x_31, x_30, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_18); +return x_32; +} +} +else +{ +lean_object* x_33; lean_object* x_34; uint8_t x_35; lean_object* x_36; +x_33 = lean_ctor_get(x_15, 1); +lean_inc(x_33); +lean_dec(x_15); +x_34 = lean_box(0); +x_35 = lean_unbox(x_16); +lean_dec(x_16); +x_36 = l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__4(x_3, x_2, x_1, x_4, x_14, x_35, x_34, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_33); +return x_36; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_evalTactic___spec__9(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) { +_start: +{ +uint8_t x_14; +x_14 = lean_usize_dec_eq(x_2, x_3); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; +lean_dec(x_4); +x_15 = lean_array_uget(x_1, x_2); +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_16 = l_Lean_Elab_Tactic_evalTactic(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_16) == 0) +{ +lean_object* x_17; lean_object* x_18; size_t x_19; size_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); +x_19 = 1; +x_20 = lean_usize_add(x_2, x_19); +x_2 = x_20; +x_4 = x_17; +x_13 = x_18; +goto _start; +} +else +{ +uint8_t x_22; +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); +x_22 = !lean_is_exclusive(x_16); +if (x_22 == 0) +{ +return x_16; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_16, 0); +x_24 = lean_ctor_get(x_16, 1); +lean_inc(x_24); +lean_inc(x_23); +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; +} +} +} +else +{ +lean_object* x_26; +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); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_4); +lean_ctor_set(x_26, 1, x_13); +return x_26; +} +} +} +LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_evalTactic___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; lean_object* x_12; lean_object* x_13; +x_11 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_evalTactic_expandEval___spec__5___closed__2; +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_1); +lean_ctor_set(x_12, 1, x_11); +x_13 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_10); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Elab_Tactic_evalTactic___spec__11___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_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_apply_8(x_3, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_15 = l_Lean_profileitIOUnsafe___rarg(x_1, x_2, x_14, x_4, x_13); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Elab_Tactic_evalTactic___spec__11(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_profileitM___at_Lean_Elab_Tactic_evalTactic___spec__11___rarg___boxed), 13, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalTactic___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 = lean_box(0); +x_5 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_5, 0, x_4); +lean_ctor_set(x_5, 1, x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalTactic___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; +x_12 = l_Lean_MessageData_ofSyntax(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; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalTactic___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_dec(x_4); +x_14 = l_Lean_Elab_Tactic_saveState___rarg(x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +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_addTrace___at_Lean_Elab_Tactic_evalTactic_handleEx___spec__2___closed__1; +x_18 = l_Lean_Elab_Tactic_evalTactic_expandEval(x_1, x_15, x_2, x_3, x_17, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_16); return x_18; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalTactic___lambda__3___closed__1() { +static lean_object* _init_l_Lean_Elab_Tactic_evalTactic___lambda__4___closed__1() { _start: { lean_object* x_1; @@ -6285,7 +7404,7 @@ x_1 = l_Lean_Elab_Tactic_tacticElabAttribute; return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalTactic___lambda__3___closed__2() { +static lean_object* _init_l_Lean_Elab_Tactic_evalTactic___lambda__4___closed__2() { _start: { lean_object* x_1; @@ -6293,7 +7412,7 @@ x_1 = l_Lean_Elab_macroAttribute; return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalTactic___lambda__3___closed__3() { +static lean_object* _init_l_Lean_Elab_Tactic_evalTactic___lambda__4___closed__3() { _start: { lean_object* x_1; @@ -6301,16 +7420,16 @@ x_1 = lean_mk_string_from_bytes("tactic '", 8); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalTactic___lambda__3___closed__4() { +static lean_object* _init_l_Lean_Elab_Tactic_evalTactic___lambda__4___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic_evalTactic___lambda__3___closed__3; +x_1 = l_Lean_Elab_Tactic_evalTactic___lambda__4___closed__3; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalTactic___lambda__3___closed__5() { +static lean_object* _init_l_Lean_Elab_Tactic_evalTactic___lambda__4___closed__5() { _start: { lean_object* x_1; @@ -6318,148 +7437,100 @@ x_1 = lean_mk_string_from_bytes("' has not been implemented", 26); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_evalTactic___lambda__3___closed__6() { +static lean_object* _init_l_Lean_Elab_Tactic_evalTactic___lambda__4___closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic_evalTactic___lambda__3___closed__5; +x_1 = l_Lean_Elab_Tactic_evalTactic___lambda__4___closed__5; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalTactic___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_EXPORT lean_object* l_Lean_Elab_Tactic_evalTactic___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) { _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; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; -lean_dec(x_2); -x_12 = lean_st_ref_get(x_10, x_11); -x_13 = lean_ctor_get(x_12, 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; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_11 = lean_st_ref_get(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); -x_14 = lean_ctor_get(x_12, 1); +lean_dec(x_11); +x_14 = lean_ctor_get(x_12, 0); lean_inc(x_14); lean_dec(x_12); -x_15 = lean_ctor_get(x_13, 0); -lean_inc(x_15); -lean_dec(x_13); lean_inc(x_1); -x_16 = l_Lean_Syntax_getKind(x_1); -x_17 = l_Lean_Elab_Tactic_evalTactic___lambda__3___closed__1; -lean_inc(x_16); -x_18 = l_Lean_KeyedDeclsAttribute_getEntries___rarg(x_17, x_15, x_16); -x_19 = lean_st_ref_get(x_10, x_14); -x_20 = lean_ctor_get(x_19, 0); +x_15 = l_Lean_Syntax_getKind(x_1); +x_16 = l_Lean_Elab_Tactic_evalTactic___lambda__4___closed__1; +lean_inc(x_15); +x_17 = l_Lean_KeyedDeclsAttribute_getEntries___rarg(x_16, x_14, x_15); +x_18 = lean_st_ref_get(x_9, 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); -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); -lean_inc(x_22); -lean_dec(x_20); -x_23 = l_Lean_Elab_Tactic_evalTactic___lambda__3___closed__2; -lean_inc(x_16); -x_24 = l_Lean_KeyedDeclsAttribute_getEntries___rarg(x_23, x_22, x_16); -x_25 = l_List_isEmpty___rarg(x_18); -if (x_25 == 0) -{ -lean_object* x_26; lean_object* x_27; -lean_dec(x_16); -x_26 = lean_box(0); -x_27 = l_Lean_Elab_Tactic_evalTactic___lambda__2(x_1, x_24, x_18, x_26, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_21); -return x_27; -} -else -{ -uint8_t x_28; -x_28 = l_List_isEmpty___rarg(x_24); -if (x_28 == 0) -{ -lean_object* x_29; lean_object* x_30; -lean_dec(x_16); -x_29 = lean_box(0); -x_30 = l_Lean_Elab_Tactic_evalTactic___lambda__2(x_1, x_24, x_18, x_29, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_21); -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; uint8_t x_37; -lean_dec(x_24); -lean_dec(x_18); -x_31 = l_Lean_MessageData_ofName(x_16); -x_32 = l_Lean_Elab_Tactic_evalTactic___lambda__3___closed__4; -x_33 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_33, 0, x_32); -lean_ctor_set(x_33, 1, x_31); -x_34 = l_Lean_Elab_Tactic_evalTactic___lambda__3___closed__6; -x_35 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -x_36 = l_Lean_throwErrorAt___at_Lean_Elab_Tactic_evalTactic___spec__1(x_1, x_35, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_21); -x_37 = !lean_is_exclusive(x_36); -if (x_37 == 0) -{ -return x_36; -} -else +x_22 = l_Lean_Elab_Tactic_evalTactic___lambda__4___closed__2; +lean_inc(x_15); +x_23 = l_Lean_KeyedDeclsAttribute_getEntries___rarg(x_22, x_21, x_15); +x_24 = l_List_isEmpty___rarg(x_17); +if (x_24 == 0) { -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_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_object* x_25; lean_object* x_26; +lean_dec(x_15); +x_25 = lean_box(0); +x_26 = l_Lean_Elab_Tactic_evalTactic___lambda__3(x_1, x_23, x_17, x_25, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_20); +return x_26; } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalTactic___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) { -_start: +else { -lean_object* x_12; lean_object* x_13; uint8_t x_14; -lean_inc(x_1); -x_12 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Tactic_evalTactic_handleEx___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_9, x_10, x_11); -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_unbox(x_13); -lean_dec(x_13); -if (x_14 == 0) +uint8_t x_27; +x_27 = l_List_isEmpty___rarg(x_23); +if (x_27 == 0) { -lean_object* x_15; lean_object* x_16; lean_object* x_17; -lean_dec(x_1); -x_15 = lean_ctor_get(x_12, 1); -lean_inc(x_15); -lean_dec(x_12); -x_16 = lean_box(0); -x_17 = l_Lean_Elab_Tactic_evalTactic___lambda__3(x_8, x_16, x_2, x_3, x_4, x_5, x_6, x_7, x_9, x_10, x_15); -return x_17; +lean_object* x_28; lean_object* x_29; +lean_dec(x_15); +x_28 = lean_box(0); +x_29 = l_Lean_Elab_Tactic_evalTactic___lambda__3(x_1, x_23, x_17, x_28, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_20); +return x_29; } 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; -x_18 = lean_ctor_get(x_12, 1); -lean_inc(x_18); -lean_dec(x_12); -lean_inc(x_8); -x_19 = l_Lean_MessageData_ofSyntax(x_8); -x_20 = l_Lean_Elab_Term_reportUnsolvedGoals___closed__7; -x_21 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_21, 0, x_20); -lean_ctor_set(x_21, 1, x_19); -x_22 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_22, 0, x_21); -lean_ctor_set(x_22, 1, x_20); -x_23 = l_Lean_addTrace___at_Lean_Elab_Tactic_evalTactic_handleEx___spec__2(x_1, x_22, x_2, x_3, x_4, x_5, x_6, x_7, x_9, x_10, x_18); -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_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_dec(x_23); -x_26 = l_Lean_Elab_Tactic_evalTactic___lambda__3(x_8, x_24, x_2, x_3, x_4, x_5, x_6, x_7, x_9, x_10, x_25); -return x_26; +lean_dec(x_17); +x_30 = l_Lean_MessageData_ofName(x_15); +x_31 = l_Lean_Elab_Tactic_evalTactic___lambda__4___closed__4; +x_32 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_30); +x_33 = l_Lean_Elab_Tactic_evalTactic___lambda__4___closed__6; +x_34 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +x_35 = l_Lean_throwErrorAt___at_Lean_Elab_Tactic_evalTactic___spec__1(x_1, x_34, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_20); +x_36 = !lean_is_exclusive(x_35); +if (x_36 == 0) +{ +return x_35; +} +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(1, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +return x_39; +} +} } } } @@ -6610,116 +7681,125 @@ 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_evalTactic___lambda__5___closed__5; -x_34 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_evalTactic___lambda__4), 11, 8); -lean_closure_set(x_34, 0, x_33); -lean_closure_set(x_34, 1, x_2); -lean_closure_set(x_34, 2, x_3); -lean_closure_set(x_34, 3, x_4); -lean_closure_set(x_34, 4, x_5); -lean_closure_set(x_34, 5, x_6); -lean_closure_set(x_34, 6, x_7); -lean_closure_set(x_34, 7, x_1); -x_35 = l_Lean_Core_withFreshMacroScope___rarg(x_34, x_27, x_9, x_10); -return x_35; +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_inc(x_1); +x_33 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_evalTactic___lambda__2___boxed), 11, 1); +lean_closure_set(x_33, 0, x_1); +x_34 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_evalTactic___lambda__4), 10, 1); +lean_closure_set(x_34, 0, x_1); +x_35 = l_Lean_Elab_Tactic_evalTactic___lambda__5___closed__5; +x_36 = 1; +x_37 = lean_box(x_36); +x_38 = lean_alloc_closure((void*)(l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___boxed), 13, 10); +lean_closure_set(x_38, 0, x_35); +lean_closure_set(x_38, 1, x_33); +lean_closure_set(x_38, 2, x_34); +lean_closure_set(x_38, 3, x_37); +lean_closure_set(x_38, 4, x_2); +lean_closure_set(x_38, 5, x_3); +lean_closure_set(x_38, 6, x_4); +lean_closure_set(x_38, 7, x_5); +lean_closure_set(x_38, 8, x_6); +lean_closure_set(x_38, 9, x_7); +x_39 = l_Lean_Core_withFreshMacroScope___rarg(x_38, x_27, x_9, x_10); +return x_39; } else { -lean_object* x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; -x_36 = l_Lean_Syntax_getArgs(x_1); +lean_object* x_40; lean_object* x_41; lean_object* x_42; uint8_t x_43; +x_40 = l_Lean_Syntax_getArgs(x_1); lean_dec(x_1); -x_37 = lean_array_get_size(x_36); -x_38 = lean_unsigned_to_nat(0u); -x_39 = lean_nat_dec_lt(x_38, x_37); -if (x_39 == 0) +x_41 = lean_array_get_size(x_40); +x_42 = lean_unsigned_to_nat(0u); +x_43 = lean_nat_dec_lt(x_42, x_41); +if (x_43 == 0) { -lean_object* x_40; lean_object* x_41; -lean_dec(x_37); -lean_dec(x_36); +lean_object* x_44; lean_object* x_45; +lean_dec(x_41); +lean_dec(x_40); 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_40 = l_Lean_Elab_Tactic_evalTactic___lambda__5___closed__1; -x_41 = l_Lean_Core_withFreshMacroScope___rarg(x_40, x_27, x_9, x_10); -return x_41; +x_44 = l_Lean_Elab_Tactic_evalTactic___lambda__5___closed__1; +x_45 = l_Lean_Core_withFreshMacroScope___rarg(x_44, x_27, x_9, x_10); +return x_45; } else { -uint8_t x_42; -x_42 = lean_nat_dec_le(x_37, x_37); -if (x_42 == 0) +uint8_t x_46; +x_46 = lean_nat_dec_le(x_41, x_41); +if (x_46 == 0) { -lean_object* x_43; lean_object* x_44; -lean_dec(x_37); -lean_dec(x_36); +lean_object* x_47; lean_object* x_48; +lean_dec(x_41); +lean_dec(x_40); 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_43 = l_Lean_Elab_Tactic_evalTactic___lambda__5___closed__1; -x_44 = l_Lean_Core_withFreshMacroScope___rarg(x_43, x_27, x_9, x_10); -return x_44; +x_47 = l_Lean_Elab_Tactic_evalTactic___lambda__5___closed__1; +x_48 = l_Lean_Core_withFreshMacroScope___rarg(x_47, x_27, x_9, x_10); +return x_48; } else { -size_t x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_45 = lean_usize_of_nat(x_37); -lean_dec(x_37); -x_46 = lean_box(0); -x_47 = l_Lean_Elab_Tactic_evalTactic___lambda__5___boxed__const__1; -x_48 = lean_box_usize(x_45); -x_49 = lean_alloc_closure((void*)(l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_evalTactic___spec__3___boxed), 13, 10); -lean_closure_set(x_49, 0, x_36); -lean_closure_set(x_49, 1, x_47); -lean_closure_set(x_49, 2, x_48); -lean_closure_set(x_49, 3, x_46); -lean_closure_set(x_49, 4, x_2); -lean_closure_set(x_49, 5, x_3); -lean_closure_set(x_49, 6, x_4); -lean_closure_set(x_49, 7, x_5); -lean_closure_set(x_49, 8, x_6); -lean_closure_set(x_49, 9, x_7); -x_50 = l_Lean_Core_withFreshMacroScope___rarg(x_49, x_27, x_9, x_10); -return x_50; +size_t x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_49 = lean_usize_of_nat(x_41); +lean_dec(x_41); +x_50 = lean_box(0); +x_51 = l_Lean_Elab_Tactic_evalTactic___lambda__5___boxed__const__1; +x_52 = lean_box_usize(x_49); +x_53 = lean_alloc_closure((void*)(l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_evalTactic___spec__9___boxed), 13, 10); +lean_closure_set(x_53, 0, x_40); +lean_closure_set(x_53, 1, x_51); +lean_closure_set(x_53, 2, x_52); +lean_closure_set(x_53, 3, x_50); +lean_closure_set(x_53, 4, x_2); +lean_closure_set(x_53, 5, x_3); +lean_closure_set(x_53, 6, x_4); +lean_closure_set(x_53, 7, x_5); +lean_closure_set(x_53, 8, x_6); +lean_closure_set(x_53, 9, x_7); +x_54 = l_Lean_Core_withFreshMacroScope___rarg(x_53, x_27, x_9, x_10); +return x_54; } } } } default: { -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 = l_Lean_MessageData_ofSyntax(x_1); -x_52 = l_Lean_indentD(x_51); -x_53 = l_Lean_Elab_Tactic_evalTactic___lambda__5___closed__7; -x_54 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_54, 0, x_53); -lean_ctor_set(x_54, 1, x_52); -x_55 = l_Lean_Elab_Term_reportUnsolvedGoals___closed__7; -x_56 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_56, 0, x_54); -lean_ctor_set(x_56, 1, x_55); -x_57 = lean_alloc_closure((void*)(l_Lean_throwError___at_Lean_Elab_Tactic_evalTactic_throwExs___spec__2___boxed), 10, 7); -lean_closure_set(x_57, 0, x_56); -lean_closure_set(x_57, 1, x_2); -lean_closure_set(x_57, 2, x_3); -lean_closure_set(x_57, 3, x_4); -lean_closure_set(x_57, 4, x_5); -lean_closure_set(x_57, 5, x_6); -lean_closure_set(x_57, 6, x_7); -x_58 = l_Lean_Core_withFreshMacroScope___rarg(x_57, x_27, x_9, x_10); -return x_58; -} -} -} -else -{ -lean_object* x_59; +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_55 = l_Lean_MessageData_ofSyntax(x_1); +x_56 = l_Lean_indentD(x_55); +x_57 = l_Lean_Elab_Tactic_evalTactic___lambda__5___closed__7; +x_58 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_58, 0, x_57); +lean_ctor_set(x_58, 1, x_56); +x_59 = l_Lean_Elab_Term_reportUnsolvedGoals___closed__7; +x_60 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_60, 0, x_58); +lean_ctor_set(x_60, 1, x_59); +x_61 = lean_alloc_closure((void*)(l_Lean_throwError___at_Lean_Elab_Tactic_evalTactic_throwExs___spec__2___boxed), 10, 7); +lean_closure_set(x_61, 0, x_60); +lean_closure_set(x_61, 1, x_2); +lean_closure_set(x_61, 2, x_3); +lean_closure_set(x_61, 3, x_4); +lean_closure_set(x_61, 4, x_5); +lean_closure_set(x_61, 5, x_6); +lean_closure_set(x_61, 6, x_7); +x_62 = l_Lean_Core_withFreshMacroScope___rarg(x_61, x_27, x_9, x_10); +return x_62; +} +} +} +else +{ +lean_object* x_63; lean_dec(x_22); lean_dec(x_21); lean_dec(x_20); @@ -6731,7 +7811,7 @@ lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_1); -x_59 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_evalTactic___spec__4(x_23, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_63 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_evalTactic___spec__10(x_23, 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); @@ -6740,23 +7820,27 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -return x_59; +return x_63; } } 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; uint8_t x_73; -x_60 = lean_ctor_get(x_8, 0); -x_61 = lean_ctor_get(x_8, 1); -x_62 = lean_ctor_get(x_8, 2); -x_63 = lean_ctor_get(x_8, 3); -x_64 = lean_ctor_get(x_8, 4); -x_65 = lean_ctor_get(x_8, 5); -x_66 = lean_ctor_get(x_8, 6); -x_67 = lean_ctor_get(x_8, 7); -x_68 = lean_ctor_get(x_8, 8); -x_69 = lean_ctor_get(x_8, 9); -x_70 = lean_ctor_get(x_8, 10); +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; +x_64 = lean_ctor_get(x_8, 0); +x_65 = lean_ctor_get(x_8, 1); +x_66 = lean_ctor_get(x_8, 2); +x_67 = lean_ctor_get(x_8, 3); +x_68 = lean_ctor_get(x_8, 4); +x_69 = lean_ctor_get(x_8, 5); +x_70 = lean_ctor_get(x_8, 6); +x_71 = lean_ctor_get(x_8, 7); +x_72 = lean_ctor_get(x_8, 8); +x_73 = lean_ctor_get(x_8, 9); +x_74 = lean_ctor_get(x_8, 10); +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); @@ -6764,211 +7848,216 @@ lean_inc(x_67); lean_inc(x_66); lean_inc(x_65); lean_inc(x_64); -lean_inc(x_63); -lean_inc(x_62); -lean_inc(x_61); -lean_inc(x_60); lean_dec(x_8); -x_71 = l_Lean_replaceRef(x_1, x_65); -lean_dec(x_65); +x_75 = l_Lean_replaceRef(x_1, x_69); +lean_dec(x_69); +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_75); lean_inc(x_68); lean_inc(x_67); lean_inc(x_66); -lean_inc(x_71); +lean_inc(x_65); lean_inc(x_64); -lean_inc(x_63); -lean_inc(x_62); -lean_inc(x_61); -lean_inc(x_60); -x_72 = lean_alloc_ctor(0, 11, 0); -lean_ctor_set(x_72, 0, x_60); -lean_ctor_set(x_72, 1, x_61); -lean_ctor_set(x_72, 2, x_62); -lean_ctor_set(x_72, 3, x_63); -lean_ctor_set(x_72, 4, x_64); -lean_ctor_set(x_72, 5, x_71); -lean_ctor_set(x_72, 6, x_66); -lean_ctor_set(x_72, 7, x_67); -lean_ctor_set(x_72, 8, x_68); -lean_ctor_set(x_72, 9, x_69); -lean_ctor_set(x_72, 10, x_70); -x_73 = lean_nat_dec_eq(x_63, x_64); -if (x_73 == 0) -{ -lean_object* x_74; lean_object* x_75; lean_object* x_76; -lean_dec(x_72); -x_74 = lean_unsigned_to_nat(1u); -x_75 = lean_nat_add(x_63, x_74); -lean_dec(x_63); x_76 = lean_alloc_ctor(0, 11, 0); -lean_ctor_set(x_76, 0, x_60); -lean_ctor_set(x_76, 1, x_61); -lean_ctor_set(x_76, 2, x_62); -lean_ctor_set(x_76, 3, x_75); -lean_ctor_set(x_76, 4, x_64); -lean_ctor_set(x_76, 5, x_71); -lean_ctor_set(x_76, 6, x_66); -lean_ctor_set(x_76, 7, x_67); -lean_ctor_set(x_76, 8, x_68); -lean_ctor_set(x_76, 9, x_69); -lean_ctor_set(x_76, 10, x_70); +lean_ctor_set(x_76, 0, x_64); +lean_ctor_set(x_76, 1, x_65); +lean_ctor_set(x_76, 2, x_66); +lean_ctor_set(x_76, 3, x_67); +lean_ctor_set(x_76, 4, x_68); +lean_ctor_set(x_76, 5, x_75); +lean_ctor_set(x_76, 6, x_70); +lean_ctor_set(x_76, 7, x_71); +lean_ctor_set(x_76, 8, x_72); +lean_ctor_set(x_76, 9, x_73); +lean_ctor_set(x_76, 10, x_74); +x_77 = lean_nat_dec_eq(x_67, x_68); +if (x_77 == 0) +{ +lean_object* x_78; lean_object* x_79; lean_object* x_80; +lean_dec(x_76); +x_78 = lean_unsigned_to_nat(1u); +x_79 = lean_nat_add(x_67, x_78); +lean_dec(x_67); +x_80 = lean_alloc_ctor(0, 11, 0); +lean_ctor_set(x_80, 0, x_64); +lean_ctor_set(x_80, 1, x_65); +lean_ctor_set(x_80, 2, x_66); +lean_ctor_set(x_80, 3, x_79); +lean_ctor_set(x_80, 4, x_68); +lean_ctor_set(x_80, 5, x_75); +lean_ctor_set(x_80, 6, x_70); +lean_ctor_set(x_80, 7, x_71); +lean_ctor_set(x_80, 8, x_72); +lean_ctor_set(x_80, 9, x_73); +lean_ctor_set(x_80, 10, x_74); switch (lean_obj_tag(x_1)) { case 0: { -lean_object* x_77; lean_object* x_78; +lean_object* x_81; lean_object* x_82; 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_77 = l_Lean_Elab_Tactic_evalTactic___lambda__5___closed__1; -x_78 = l_Lean_Core_withFreshMacroScope___rarg(x_77, x_76, x_9, x_10); -return x_78; +x_81 = l_Lean_Elab_Tactic_evalTactic___lambda__5___closed__1; +x_82 = l_Lean_Core_withFreshMacroScope___rarg(x_81, x_80, x_9, x_10); +return x_82; } 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_Elab_Tactic_evalTactic___lambda__5___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_evalTactic___lambda__5___closed__5; -x_83 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_evalTactic___lambda__4), 11, 8); -lean_closure_set(x_83, 0, x_82); -lean_closure_set(x_83, 1, x_2); -lean_closure_set(x_83, 2, x_3); -lean_closure_set(x_83, 3, x_4); -lean_closure_set(x_83, 4, x_5); -lean_closure_set(x_83, 5, x_6); -lean_closure_set(x_83, 6, x_7); -lean_closure_set(x_83, 7, x_1); -x_84 = l_Lean_Core_withFreshMacroScope___rarg(x_83, x_76, x_9, x_10); -return x_84; -} -else -{ -lean_object* x_85; lean_object* x_86; lean_object* x_87; uint8_t x_88; -x_85 = l_Lean_Syntax_getArgs(x_1); +lean_object* x_83; lean_object* x_84; uint8_t x_85; +x_83 = lean_ctor_get(x_1, 1); +lean_inc(x_83); +x_84 = l_Lean_Elab_Tactic_evalTactic___lambda__5___closed__3; +x_85 = lean_name_eq(x_83, x_84); +lean_dec(x_83); +if (x_85 == 0) +{ +lean_object* x_86; lean_object* x_87; lean_object* x_88; uint8_t x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; +lean_inc(x_1); +x_86 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_evalTactic___lambda__2___boxed), 11, 1); +lean_closure_set(x_86, 0, x_1); +x_87 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_evalTactic___lambda__4), 10, 1); +lean_closure_set(x_87, 0, x_1); +x_88 = l_Lean_Elab_Tactic_evalTactic___lambda__5___closed__5; +x_89 = 1; +x_90 = lean_box(x_89); +x_91 = lean_alloc_closure((void*)(l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___boxed), 13, 10); +lean_closure_set(x_91, 0, x_88); +lean_closure_set(x_91, 1, x_86); +lean_closure_set(x_91, 2, x_87); +lean_closure_set(x_91, 3, x_90); +lean_closure_set(x_91, 4, x_2); +lean_closure_set(x_91, 5, x_3); +lean_closure_set(x_91, 6, x_4); +lean_closure_set(x_91, 7, x_5); +lean_closure_set(x_91, 8, x_6); +lean_closure_set(x_91, 9, x_7); +x_92 = l_Lean_Core_withFreshMacroScope___rarg(x_91, x_80, x_9, x_10); +return x_92; +} +else +{ +lean_object* x_93; lean_object* x_94; lean_object* x_95; uint8_t x_96; +x_93 = l_Lean_Syntax_getArgs(x_1); lean_dec(x_1); -x_86 = lean_array_get_size(x_85); -x_87 = lean_unsigned_to_nat(0u); -x_88 = lean_nat_dec_lt(x_87, x_86); -if (x_88 == 0) +x_94 = lean_array_get_size(x_93); +x_95 = lean_unsigned_to_nat(0u); +x_96 = lean_nat_dec_lt(x_95, x_94); +if (x_96 == 0) { -lean_object* x_89; lean_object* x_90; -lean_dec(x_86); -lean_dec(x_85); +lean_object* x_97; lean_object* x_98; +lean_dec(x_94); +lean_dec(x_93); 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_89 = l_Lean_Elab_Tactic_evalTactic___lambda__5___closed__1; -x_90 = l_Lean_Core_withFreshMacroScope___rarg(x_89, x_76, x_9, x_10); -return x_90; +x_97 = l_Lean_Elab_Tactic_evalTactic___lambda__5___closed__1; +x_98 = l_Lean_Core_withFreshMacroScope___rarg(x_97, x_80, x_9, x_10); +return x_98; } else { -uint8_t x_91; -x_91 = lean_nat_dec_le(x_86, x_86); -if (x_91 == 0) +uint8_t x_99; +x_99 = lean_nat_dec_le(x_94, x_94); +if (x_99 == 0) { -lean_object* x_92; lean_object* x_93; -lean_dec(x_86); -lean_dec(x_85); +lean_object* x_100; lean_object* x_101; +lean_dec(x_94); +lean_dec(x_93); 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_92 = l_Lean_Elab_Tactic_evalTactic___lambda__5___closed__1; -x_93 = l_Lean_Core_withFreshMacroScope___rarg(x_92, x_76, x_9, x_10); -return x_93; +x_100 = l_Lean_Elab_Tactic_evalTactic___lambda__5___closed__1; +x_101 = l_Lean_Core_withFreshMacroScope___rarg(x_100, x_80, x_9, x_10); +return x_101; } else { -size_t x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; -x_94 = lean_usize_of_nat(x_86); -lean_dec(x_86); -x_95 = lean_box(0); -x_96 = l_Lean_Elab_Tactic_evalTactic___lambda__5___boxed__const__1; -x_97 = lean_box_usize(x_94); -x_98 = lean_alloc_closure((void*)(l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_evalTactic___spec__3___boxed), 13, 10); -lean_closure_set(x_98, 0, x_85); -lean_closure_set(x_98, 1, x_96); -lean_closure_set(x_98, 2, x_97); -lean_closure_set(x_98, 3, x_95); -lean_closure_set(x_98, 4, x_2); -lean_closure_set(x_98, 5, x_3); -lean_closure_set(x_98, 6, x_4); -lean_closure_set(x_98, 7, x_5); -lean_closure_set(x_98, 8, x_6); -lean_closure_set(x_98, 9, x_7); -x_99 = l_Lean_Core_withFreshMacroScope___rarg(x_98, x_76, x_9, x_10); -return x_99; +size_t x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; +x_102 = lean_usize_of_nat(x_94); +lean_dec(x_94); +x_103 = lean_box(0); +x_104 = l_Lean_Elab_Tactic_evalTactic___lambda__5___boxed__const__1; +x_105 = lean_box_usize(x_102); +x_106 = lean_alloc_closure((void*)(l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_evalTactic___spec__9___boxed), 13, 10); +lean_closure_set(x_106, 0, x_93); +lean_closure_set(x_106, 1, x_104); +lean_closure_set(x_106, 2, x_105); +lean_closure_set(x_106, 3, x_103); +lean_closure_set(x_106, 4, x_2); +lean_closure_set(x_106, 5, x_3); +lean_closure_set(x_106, 6, x_4); +lean_closure_set(x_106, 7, x_5); +lean_closure_set(x_106, 8, x_6); +lean_closure_set(x_106, 9, x_7); +x_107 = l_Lean_Core_withFreshMacroScope___rarg(x_106, x_80, x_9, x_10); +return x_107; } } } } default: { -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_100 = l_Lean_MessageData_ofSyntax(x_1); -x_101 = l_Lean_indentD(x_100); -x_102 = l_Lean_Elab_Tactic_evalTactic___lambda__5___closed__7; -x_103 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_103, 0, x_102); -lean_ctor_set(x_103, 1, x_101); -x_104 = l_Lean_Elab_Term_reportUnsolvedGoals___closed__7; -x_105 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_105, 0, x_103); -lean_ctor_set(x_105, 1, x_104); -x_106 = lean_alloc_closure((void*)(l_Lean_throwError___at_Lean_Elab_Tactic_evalTactic_throwExs___spec__2___boxed), 10, 7); -lean_closure_set(x_106, 0, x_105); -lean_closure_set(x_106, 1, x_2); -lean_closure_set(x_106, 2, x_3); -lean_closure_set(x_106, 3, x_4); -lean_closure_set(x_106, 4, x_5); -lean_closure_set(x_106, 5, x_6); -lean_closure_set(x_106, 6, x_7); -x_107 = l_Lean_Core_withFreshMacroScope___rarg(x_106, x_76, x_9, x_10); -return x_107; -} -} -} -else -{ -lean_object* x_108; +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_108 = l_Lean_MessageData_ofSyntax(x_1); +x_109 = l_Lean_indentD(x_108); +x_110 = l_Lean_Elab_Tactic_evalTactic___lambda__5___closed__7; +x_111 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_111, 0, x_110); +lean_ctor_set(x_111, 1, x_109); +x_112 = l_Lean_Elab_Term_reportUnsolvedGoals___closed__7; +x_113 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_113, 0, x_111); +lean_ctor_set(x_113, 1, x_112); +x_114 = lean_alloc_closure((void*)(l_Lean_throwError___at_Lean_Elab_Tactic_evalTactic_throwExs___spec__2___boxed), 10, 7); +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_7); +x_115 = l_Lean_Core_withFreshMacroScope___rarg(x_114, x_80, x_9, x_10); +return x_115; +} +} +} +else +{ +lean_object* x_116; +lean_dec(x_74); +lean_dec(x_73); +lean_dec(x_72); +lean_dec(x_71); lean_dec(x_70); -lean_dec(x_69); lean_dec(x_68); lean_dec(x_67); lean_dec(x_66); +lean_dec(x_65); lean_dec(x_64); -lean_dec(x_63); -lean_dec(x_62); -lean_dec(x_61); -lean_dec(x_60); lean_dec(x_1); -x_108 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_evalTactic___spec__4(x_71, x_2, x_3, x_4, x_5, x_6, x_7, x_72, x_9, x_10); +x_116 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_evalTactic___spec__10(x_75, x_2, x_3, x_4, x_5, x_6, x_7, x_76, x_9, x_10); lean_dec(x_9); -lean_dec(x_72); +lean_dec(x_76); 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_108; +return x_116; } } } @@ -6992,7 +8081,7 @@ x_12 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_evalTactic___lambda__5), 10 lean_closure_set(x_12, 0, x_1); x_13 = l_Lean_Syntax_getKind(x_1); x_14 = l_Lean_Elab_Tactic_evalTactic___closed__1; -x_15 = l_Lean_profileitM___at_Lean_Elab_Tactic_evalTactic___spec__5___rarg(x_14, x_11, x_12, x_13, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_15 = l_Lean_profileitM___at_Lean_Elab_Tactic_evalTactic___spec__11___rarg(x_14, x_11, x_12, x_13, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_11); return x_15; } @@ -7113,7 +8202,154 @@ lean_dec(x_2); return x_11; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_evalTactic___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) { +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Tactic_evalTactic___spec__4___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Tactic_evalTactic___spec__4___rarg(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Tactic_evalTactic___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) { +_start: +{ +lean_object* x_8; +x_8 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Tactic_evalTactic___spec__4(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___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Elab_Tactic_evalTactic___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) { +_start: +{ +uint8_t x_15; lean_object* x_16; +x_15 = lean_unbox(x_5); +lean_dec(x_5); +x_16 = l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Elab_Tactic_evalTactic___spec__7(x_1, x_2, x_3, x_4, x_15, 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); +return x_16; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Tactic_evalTactic___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) { +_start: +{ +uint8_t x_15; lean_object* x_16; +x_15 = lean_unbox(x_5); +lean_dec(x_5); +x_16 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__6(x_1, x_2, x_3, x_4, x_15, 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_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +return x_16; +} +} +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Tactic_evalTactic___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_MonadExcept_ofExcept___at_Lean_Elab_Tactic_evalTactic___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); +lean_dec(x_1); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___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) { +_start: +{ +uint8_t x_17; lean_object* x_18; +x_17 = lean_unbox(x_4); +lean_dec(x_4); +x_18 = l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__2(x_1, x_2, x_3, x_17, 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_7); +lean_dec(x_5); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__3___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: +{ +uint8_t x_19; uint8_t x_20; double x_21; lean_object* x_22; +x_19 = lean_unbox(x_5); +lean_dec(x_5); +x_20 = lean_unbox(x_7); +lean_dec(x_7); +x_21 = lean_unbox_float(x_8); +lean_dec(x_8); +x_22 = l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__3(x_1, x_2, x_3, x_4, x_19, x_6, x_20, x_21, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +return x_22; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___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, lean_object* x_15, lean_object* x_16) { +_start: +{ +uint8_t x_17; uint8_t x_18; lean_object* x_19; +x_17 = lean_unbox(x_4); +lean_dec(x_4); +x_18 = lean_unbox(x_6); +lean_dec(x_6); +x_19 = l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__4(x_1, x_2, x_3, x_17, x_5, x_18, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +return x_19; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___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: +{ +uint8_t x_14; lean_object* x_15; +x_14 = lean_unbox(x_4); +lean_dec(x_4); +x_15 = l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3(x_1, x_2, x_3, x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_evalTactic___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; @@ -7121,16 +8357,16 @@ x_14 = lean_unbox_usize(x_2); lean_dec(x_2); x_15 = lean_unbox_usize(x_3); lean_dec(x_3); -x_16 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_evalTactic___spec__3(x_1, x_14, x_15, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +x_16 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_evalTactic___spec__9(x_1, x_14, x_15, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); lean_dec(x_1); return x_16; } } -LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_evalTactic___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_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_evalTactic___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_11; -x_11 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_evalTactic___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_11 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_evalTactic___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); @@ -7142,11 +8378,11 @@ lean_dec(x_2); return x_11; } } -LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Elab_Tactic_evalTactic___spec__5___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, 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_profileitM___at_Lean_Elab_Tactic_evalTactic___spec__11___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, 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_profileitM___at_Lean_Elab_Tactic_evalTactic___spec__5___rarg(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_Lean_profileitM___at_Lean_Elab_Tactic_evalTactic___spec__11___rarg(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_2); lean_dec(x_1); return x_14; @@ -7162,6 +8398,23 @@ lean_dec(x_1); return x_4; } } +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalTactic___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_Lean_Elab_Tactic_evalTactic___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_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_Lean_throwError___at_Lean_Elab_Tactic_throwNoGoalsToBeSolved___spec__1___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: { @@ -7314,7 +8567,7 @@ if (x_14 == 0) { lean_object* x_15; lean_free_object(x_10); -x_15 = l_Lean_Elab_Term_reportUnsolvedGoals(x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_13); +x_15 = l_Lean_Elab_Term_reportUnsolvedGoals(x_12, x_5, x_6, x_7, x_8, x_13); if (lean_obj_tag(x_15) == 0) { lean_object* x_16; lean_object* x_17; @@ -7372,7 +8625,7 @@ x_25 = l_List_isEmpty___rarg(x_23); if (x_25 == 0) { lean_object* x_26; -x_26 = l_Lean_Elab_Term_reportUnsolvedGoals(x_23, x_3, x_4, x_5, x_6, x_7, x_8, x_24); +x_26 = l_Lean_Elab_Term_reportUnsolvedGoals(x_23, x_5, x_6, x_7, x_8, x_24); if (lean_obj_tag(x_26) == 0) { lean_object* x_27; lean_object* x_28; @@ -7869,7 +9122,7 @@ else { lean_object* x_267; uint8_t x_268; x_267 = l_Lean_logAt___at_Lean_Elab_Tactic_closeUsingOrAdmit___spec__2___closed__1; -x_268 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_264, x_267); +x_268 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_264, x_267); if (x_268 == 0) { x_14 = x_3; @@ -13204,7 +14457,7 @@ lean_dec(x_1); return x_5; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__1() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -13214,7 +14467,7 @@ x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__2() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -13224,27 +14477,27 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__3() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__2; +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__2; x_2 = l_Lean_Elab_Tactic_mkTacticAttribute___closed__8; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__4() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__3; +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__3; x_2 = l_Lean_Elab_Term_reportUnsolvedGoals___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__5() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__5() { _start: { lean_object* x_1; @@ -13252,17 +14505,17 @@ x_1 = lean_mk_string_from_bytes("initFn", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__6() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__4; -x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__5; +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__4; +x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__7() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__7() { _start: { lean_object* x_1; @@ -13270,47 +14523,47 @@ x_1 = lean_mk_string_from_bytes("_@", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__8() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__6; -x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__7; +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__6; +x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__7; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__9() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__8; +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__8; x_2 = l_Lean_Elab_Tactic_mkTacticAttribute___closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__10() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__9; +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__9; x_2 = l_Lean_Elab_Tactic_mkTacticAttribute___closed__8; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__11() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__10; +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__10; x_2 = l_Lean_Elab_Term_reportUnsolvedGoals___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__12() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__12() { _start: { lean_object* x_1; @@ -13318,17 +14571,17 @@ x_1 = lean_mk_string_from_bytes("Basic", 5); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__13() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__11; -x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__12; +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__11; +x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__12; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__14() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__14() { _start: { lean_object* x_1; @@ -13336,54 +14589,54 @@ x_1 = lean_mk_string_from_bytes("_hyg", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__15() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__13; -x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__14; +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__13; +x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__14; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__16() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__15; -x_2 = lean_unsigned_to_nat(10065u); +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__15; +x_2 = lean_unsigned_to_nat(9985u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__1; +x_2 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__1; x_3 = 0; -x_4 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__16; +x_4 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__16; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } } -static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10108____closed__1() { +static lean_object* _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10028____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__15; -x_2 = lean_unsigned_to_nat(10108u); +x_1 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__15; +x_2 = lean_unsigned_to_nat(10028u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10108_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10028_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; x_2 = l_List_forIn_loop___at_Lean_Elab_Tactic_evalTactic_handleEx___spec__4___closed__2; x_3 = 0; -x_4 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10108____closed__1; +x_4 = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10028____closed__1; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } @@ -13491,18 +14744,31 @@ l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_evalTactic_expandEval___spec__5_ lean_mark_persistent(l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_evalTactic_expandEval___spec__5___closed__2); l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalTactic_expandEval___spec__6___rarg___closed__1 = _init_l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalTactic_expandEval___spec__6___rarg___closed__1(); lean_mark_persistent(l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalTactic_expandEval___spec__6___rarg___closed__1); -l_Lean_Elab_Tactic_evalTactic___lambda__3___closed__1 = _init_l_Lean_Elab_Tactic_evalTactic___lambda__3___closed__1(); -lean_mark_persistent(l_Lean_Elab_Tactic_evalTactic___lambda__3___closed__1); -l_Lean_Elab_Tactic_evalTactic___lambda__3___closed__2 = _init_l_Lean_Elab_Tactic_evalTactic___lambda__3___closed__2(); -lean_mark_persistent(l_Lean_Elab_Tactic_evalTactic___lambda__3___closed__2); -l_Lean_Elab_Tactic_evalTactic___lambda__3___closed__3 = _init_l_Lean_Elab_Tactic_evalTactic___lambda__3___closed__3(); -lean_mark_persistent(l_Lean_Elab_Tactic_evalTactic___lambda__3___closed__3); -l_Lean_Elab_Tactic_evalTactic___lambda__3___closed__4 = _init_l_Lean_Elab_Tactic_evalTactic___lambda__3___closed__4(); -lean_mark_persistent(l_Lean_Elab_Tactic_evalTactic___lambda__3___closed__4); -l_Lean_Elab_Tactic_evalTactic___lambda__3___closed__5 = _init_l_Lean_Elab_Tactic_evalTactic___lambda__3___closed__5(); -lean_mark_persistent(l_Lean_Elab_Tactic_evalTactic___lambda__3___closed__5); -l_Lean_Elab_Tactic_evalTactic___lambda__3___closed__6 = _init_l_Lean_Elab_Tactic_evalTactic___lambda__3___closed__6(); -lean_mark_persistent(l_Lean_Elab_Tactic_evalTactic___lambda__3___closed__6); +l_Lean_withSeconds___at_Lean_Elab_Tactic_evalTactic___spec__5___closed__1 = _init_l_Lean_withSeconds___at_Lean_Elab_Tactic_evalTactic___spec__5___closed__1(); +l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__3___closed__1 = _init_l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__3___closed__1); +l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__3___closed__2 = _init_l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__3___closed__2); +l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__3___closed__3 = _init_l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__3___closed__3(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__3___closed__3); +l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__3___closed__4 = _init_l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__3___closed__4(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__3___closed__4); +l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__3___closed__5 = _init_l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__3___closed__5(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__3___closed__5); +l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__4___closed__1 = _init_l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__4___closed__1); +l_Lean_Elab_Tactic_evalTactic___lambda__4___closed__1 = _init_l_Lean_Elab_Tactic_evalTactic___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_Elab_Tactic_evalTactic___lambda__4___closed__1); +l_Lean_Elab_Tactic_evalTactic___lambda__4___closed__2 = _init_l_Lean_Elab_Tactic_evalTactic___lambda__4___closed__2(); +lean_mark_persistent(l_Lean_Elab_Tactic_evalTactic___lambda__4___closed__2); +l_Lean_Elab_Tactic_evalTactic___lambda__4___closed__3 = _init_l_Lean_Elab_Tactic_evalTactic___lambda__4___closed__3(); +lean_mark_persistent(l_Lean_Elab_Tactic_evalTactic___lambda__4___closed__3); +l_Lean_Elab_Tactic_evalTactic___lambda__4___closed__4 = _init_l_Lean_Elab_Tactic_evalTactic___lambda__4___closed__4(); +lean_mark_persistent(l_Lean_Elab_Tactic_evalTactic___lambda__4___closed__4); +l_Lean_Elab_Tactic_evalTactic___lambda__4___closed__5 = _init_l_Lean_Elab_Tactic_evalTactic___lambda__4___closed__5(); +lean_mark_persistent(l_Lean_Elab_Tactic_evalTactic___lambda__4___closed__5); +l_Lean_Elab_Tactic_evalTactic___lambda__4___closed__6 = _init_l_Lean_Elab_Tactic_evalTactic___lambda__4___closed__6(); +lean_mark_persistent(l_Lean_Elab_Tactic_evalTactic___lambda__4___closed__6); l_Lean_Elab_Tactic_evalTactic___lambda__5___closed__1 = _init_l_Lean_Elab_Tactic_evalTactic___lambda__5___closed__1(); lean_mark_persistent(l_Lean_Elab_Tactic_evalTactic___lambda__5___closed__1); l_Lean_Elab_Tactic_evalTactic___lambda__5___closed__2 = _init_l_Lean_Elab_Tactic_evalTactic___lambda__5___closed__2(); @@ -13575,44 +14841,44 @@ l_Lean_Elab_Tactic_getNameOfIdent_x27___closed__2 = _init_l_Lean_Elab_Tactic_get lean_mark_persistent(l_Lean_Elab_Tactic_getNameOfIdent_x27___closed__2); l_Lean_Elab_Tactic_withCaseRef___rarg___closed__1 = _init_l_Lean_Elab_Tactic_withCaseRef___rarg___closed__1(); lean_mark_persistent(l_Lean_Elab_Tactic_withCaseRef___rarg___closed__1); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__1 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__1(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__1); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__2 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__2(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__2); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__3 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__3(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__3); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__4 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__4(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__4); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__5 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__5(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__5); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__6 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__6(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__6); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__7 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__7(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__7); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__8 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__8(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__8); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__9 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__9(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__9); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__10 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__10(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__10); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__11 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__11(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__11); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__12 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__12(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__12); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__13 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__13(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__13); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__14 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__14(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__14); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__15 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__15(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__15); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__16 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__16(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065____closed__16); -res = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10065_(lean_io_mk_world()); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__1 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__1(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__1); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__2 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__2(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__2); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__3 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__3(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__3); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__4 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__4(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__4); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__5 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__5(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__5); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__6 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__6(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__6); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__7 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__7(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__7); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__8 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__8(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__8); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__9 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__9(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__9); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__10 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__10(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__10); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__11 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__11(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__11); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__12 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__12(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__12); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__13 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__13(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__13); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__14 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__14(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__14); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__15 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__15(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__15); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__16 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__16(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985____closed__16); +res = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_9985_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10108____closed__1 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10108____closed__1(); -lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10108____closed__1); -res = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10108_(lean_io_mk_world()); +l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10028____closed__1 = _init_l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10028____closed__1(); +lean_mark_persistent(l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10028____closed__1); +res = l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_10028_(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/BuiltinTactic.c b/stage0/stdlib/Lean/Elab/Tactic/BuiltinTactic.c index a9c8a19d21a9..d64654d28193 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/BuiltinTactic.c +++ b/stage0/stdlib/Lean/Elab/Tactic/BuiltinTactic.c @@ -191,7 +191,6 @@ static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalSleep_declRange___closed LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Tactic_evalAnyGoals___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___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*); static lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_Tactic_evalOpen___spec__16___closed__2; -lean_object* l_Lean_Option_get___at_Std_Format_pretty_x27___spec__1(lean_object*, lean_object*); lean_object* l_Lean_replaceRef(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalWithAnnotateState___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalClear___closed__3; @@ -475,6 +474,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalRotateRight___closed__1; static lean_object* l_Lean_Elab_throwIllFormedSyntax___at_Lean_Elab_Tactic_evalTraceMessage___spec__1___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalContradiction___closed__1; lean_object* l_Lean_ScopedEnvExtension_popScope___rarg(lean_object*, lean_object*); +lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalClear_declRange___closed__7; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalIntroMatch_declRange___closed__1; lean_object* l_Lean_Elab_Term_withoutErrToSorryImp___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -11824,7 +11824,7 @@ lean_dec(x_21); x_22 = lean_ctor_get(x_8, 2); lean_dec(x_22); x_23 = l_Lean_Elab_Tactic_elabSetOption___closed__1; -x_24 = l_Lean_Option_get___at_Std_Format_pretty_x27___spec__1(x_16, x_23); +x_24 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_16, x_23); lean_ctor_set(x_8, 4, x_24); lean_ctor_set(x_8, 2, x_16); x_25 = l_Lean_Elab_Tactic_evalTactic(x_19, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_17); @@ -11853,7 +11853,7 @@ lean_inc(x_27); lean_inc(x_26); lean_dec(x_8); x_35 = l_Lean_Elab_Tactic_elabSetOption___closed__1; -x_36 = l_Lean_Option_get___at_Std_Format_pretty_x27___spec__1(x_16, x_35); +x_36 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_16, x_35); x_37 = lean_alloc_ctor(0, 11, 0); lean_ctor_set(x_37, 0, x_26); lean_ctor_set(x_37, 1, x_27); diff --git a/stage0/stdlib/Lean/Elab/Tactic/Cache.c b/stage0/stdlib/Lean/Elab/Tactic/Cache.c index 664672f6e5e1..391fb4cb3221 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Cache.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Cache.c @@ -14,7 +14,6 @@ extern "C" { #endif static lean_object* l___private_Lean_Elab_Tactic_Cache_0__Lean_Elab_Tactic_dbg__cache_x27___closed__4; -uint8_t l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(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___regBuiltin_Lean_Elab_Tactic_evalCheckpoint_declRange(lean_object*); static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Cache___hyg_20____closed__1; @@ -73,6 +72,7 @@ LEAN_EXPORT uint8_t l___private_Lean_Elab_Tactic_Cache_0__Lean_Elab_Tactic_equiv static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Cache___hyg_20____closed__4; static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Cache___hyg_20____closed__9; static lean_object* l___private_Lean_Elab_Tactic_Cache_0__Lean_Elab_Tactic_dbg__cache_x27___closed__5; +uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); lean_object* lean_usize_to_nat(size_t); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalCheckpoint___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -277,7 +277,7 @@ lean_object* x_11; lean_object* x_12; uint8_t x_13; x_11 = lean_ctor_get(x_8, 2); lean_inc(x_11); x_12 = l___private_Lean_Elab_Tactic_Cache_0__Lean_Elab_Tactic_dbg__cache___closed__1; -x_13 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_11, x_12); +x_13 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_11, x_12); lean_dec(x_11); if (x_13 == 0) { @@ -541,7 +541,7 @@ lean_object* x_14; lean_object* x_15; uint8_t x_16; x_14 = lean_ctor_get(x_11, 2); lean_inc(x_14); x_15 = l___private_Lean_Elab_Tactic_Cache_0__Lean_Elab_Tactic_dbg__cache___closed__1; -x_16 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_14, x_15); +x_16 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_14, x_15); lean_dec(x_14); if (x_16 == 0) { diff --git a/stage0/stdlib/Lean/Elab/Tactic/Conv/Congr.c b/stage0/stdlib/Lean/Elab/Tactic/Conv/Congr.c index 53901541fc61..c64e5520fbf0 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Conv/Congr.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Conv/Congr.c @@ -171,7 +171,6 @@ static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalRhs___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalSkip___closed__4; lean_object* lean_array_to_list(lean_object*, lean_object*); lean_object* l_Lean_Meta_mkEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalExt___closed__4; static lean_object* l___private_Lean_Elab_Tactic_Conv_Congr_0__Lean_Elab_Tactic_Conv_selectIdx___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalRhs___closed__3; @@ -262,6 +261,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalCongr___closed__1; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Tactic_Conv_congr___spec__3___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Conv_Congr_0__Lean_Elab_Tactic_Conv_isImplies(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_filterMap___at_Lean_Elab_Tactic_Conv_evalCongr___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_congr___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___private_Lean_Elab_Tactic_Conv_Congr_0__Lean_Elab_Tactic_Conv_extCore___lambda__3___closed__1; @@ -4198,7 +4198,7 @@ lean_inc(x_15); x_16 = lean_ctor_get(x_13, 1); lean_inc(x_16); lean_dec(x_13); -x_17 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_15, x_4, x_5, x_6, x_7, x_14); +x_17 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_15, x_4, x_5, x_6, x_7, x_14); x_18 = lean_ctor_get(x_17, 0); lean_inc(x_18); x_19 = lean_ctor_get(x_17, 1); @@ -8682,7 +8682,7 @@ lean_inc(x_11); x_12 = lean_ctor_get(x_9, 1); lean_inc(x_12); lean_dec(x_9); -x_13 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_11, x_3, x_4, x_5, x_6, x_10); +x_13 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_11, x_3, x_4, x_5, x_6, x_10); x_14 = lean_ctor_get(x_13, 0); lean_inc(x_14); x_15 = lean_ctor_get(x_13, 1); diff --git a/stage0/stdlib/Lean/Elab/Tactic/Location.c b/stage0/stdlib/Lean/Elab/Tactic/Location.c index a10622ecf9e2..faddc0679966 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Location.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Location.c @@ -19,18 +19,21 @@ lean_object* l_Lean_Elab_Tactic_getFVarId(lean_object*, lean_object*, lean_objec static lean_object* l_Lean_Elab_Tactic_expandOptLocation___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_expandOptLocation(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_expandLocation(lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_withLocation___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_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_withLocation___spec__1(lean_object*, lean_object*, size_t, size_t, uint8_t, 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_object* l_Lean_Syntax_getArgs(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_withLocation___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* l_Lean_Elab_Tactic_getMainGoal(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_expandOptLocation___closed__2; +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_withLocation___spec__1___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*); static lean_object* l_Lean_Elab_Tactic_expandLocation___closed__1; size_t lean_usize_of_nat(lean_object*); lean_object* l_Array_reverse___rarg(lean_object*); LEAN_EXPORT 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*); static lean_object* l_Lean_Elab_Tactic_expandLocation___closed__3; lean_object* l_Lean_Syntax_getKind(lean_object*); +lean_object* l_Lean_FVarId_getDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Tactic_withMainContext___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_Tactic_expandOptLocation___boxed(lean_object*); static lean_object* l_Lean_Elab_Tactic_expandLocation___closed__4; @@ -38,11 +41,13 @@ uint8_t lean_name_eq(lean_object*, lean_object*); lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_withLocation___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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_withLocation___spec__2(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*); +uint8_t l_Lean_LocalDecl_isImplementationDetail(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_withLocation___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_Lean_Elab_Tactic_expandLocation___closed__2; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_withLocation___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*); static lean_object* l_Lean_Elab_Tactic_expandLocation___closed__5; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_withLocation___spec__1___lambda__1___closed__1; uint8_t l_Lean_Syntax_isNone(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_withLocation___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*); size_t lean_usize_add(size_t, size_t); @@ -209,6 +214,82 @@ lean_dec(x_1); return x_2; } } +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_withLocation___spec__1___lambda__1___closed__1() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; +x_1 = 1; +x_2 = lean_box(x_1); +x_3 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_3, 0, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_withLocation___spec__1___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, 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; +x_14 = lean_apply_1(x_1, x_2); +x_15 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_withMainContext___rarg), 10, 1); +lean_closure_set(x_15, 0, x_14); +x_16 = l_Lean_Elab_Tactic_tryTactic___rarg(x_15, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +if (x_3 == 0) +{ +uint8_t x_17; +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; +x_18 = lean_ctor_get(x_16, 0); +x_19 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_16, 0, x_19); +return x_16; +} +else +{ +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_inc(x_20); +lean_dec(x_16); +x_22 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_22, 0, 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_24; +x_24 = !lean_is_exclusive(x_16); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; +x_25 = lean_ctor_get(x_16, 0); +lean_dec(x_25); +x_26 = l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_withLocation___spec__1___lambda__1___closed__1; +lean_ctor_set(x_16, 0, x_26); +return x_16; +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_16, 1); +lean_inc(x_27); +lean_dec(x_16); +x_28 = l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_withLocation___spec__1___lambda__1___closed__1; +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_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_withLocation___spec__1(lean_object* x_1, lean_object* x_2, size_t x_3, size_t 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) { _start: { @@ -234,12 +315,25 @@ return x_17; } else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +lean_object* x_18; lean_object* x_19; x_18 = lean_array_uget(x_2, x_4); -lean_inc(x_1); -x_19 = lean_apply_1(x_1, x_18); -x_20 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_withMainContext___rarg), 10, 1); -lean_closure_set(x_20, 0, x_19); +lean_inc(x_10); +lean_inc(x_18); +x_19 = l_Lean_FVarId_getDecl(x_18, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; lean_object* x_21; uint8_t x_22; +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_Lean_LocalDecl_isImplementationDetail(x_20); +lean_dec(x_20); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; size_t x_28; size_t x_29; uint8_t x_30; +x_23 = lean_box(0); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); @@ -248,39 +342,69 @@ lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_21 = l_Lean_Elab_Tactic_tryTactic___rarg(x_20, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (x_5 == 0) -{ -lean_object* x_22; lean_object* x_23; size_t x_24; size_t x_25; uint8_t x_26; -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 = 1; -x_25 = lean_usize_add(x_4, x_24); -x_26 = lean_unbox(x_22); -lean_dec(x_22); -x_4 = x_25; -x_5 = x_26; -x_14 = x_23; +lean_inc(x_1); +x_24 = l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_withLocation___spec__1___lambda__1(x_1, x_18, x_5, x_23, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_21); +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_25, 0); +lean_inc(x_27); +lean_dec(x_25); +x_28 = 1; +x_29 = lean_usize_add(x_4, x_28); +x_30 = lean_unbox(x_27); +lean_dec(x_27); +x_4 = x_29; +x_5 = x_30; +x_14 = x_26; goto _start; } else { -lean_object* x_28; size_t x_29; size_t x_30; uint8_t x_31; -x_28 = lean_ctor_get(x_21, 1); -lean_inc(x_28); -lean_dec(x_21); -x_29 = 1; -x_30 = lean_usize_add(x_4, x_29); -x_31 = 1; -x_4 = x_30; -x_5 = x_31; -x_14 = x_28; +size_t x_32; size_t x_33; +lean_dec(x_18); +x_32 = 1; +x_33 = lean_usize_add(x_4, x_32); +x_4 = x_33; +x_14 = x_21; goto _start; } } +else +{ +uint8_t x_35; +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_1); +x_35 = !lean_is_exclusive(x_19); +if (x_35 == 0) +{ +return x_19; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_19, 0); +x_37 = lean_ctor_get(x_19, 1); +lean_inc(x_37); +lean_inc(x_36); +lean_dec(x_19); +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; +} +} +} } } LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_withLocation___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) { @@ -432,7 +556,7 @@ return x_28; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_withLocation___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, 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; size_t x_17; size_t x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; +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; x_13 = lean_ctor_get(x_8, 1); lean_inc(x_13); x_14 = l_Lean_LocalContext_getFVarIds(x_13); @@ -451,6 +575,9 @@ lean_inc(x_5); lean_inc(x_4); x_19 = l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_withLocation___spec__1(x_1, x_15, x_17, x_18, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); lean_dec(x_15); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; uint8_t x_21; x_20 = lean_ctor_get(x_19, 0); lean_inc(x_20); x_21 = lean_unbox(x_20); @@ -549,6 +676,38 @@ return x_36; } } } +else +{ +uint8_t x_37; +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_37 = !lean_is_exclusive(x_19); +if (x_37 == 0) +{ +return x_19; +} +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_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_EXPORT lean_object* l_Lean_Elab_Tactic_withLocation(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, 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: @@ -747,6 +906,17 @@ return x_45; } } } +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_withLocation___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) { +_start: +{ +uint8_t x_14; lean_object* x_15; +x_14 = lean_unbox(x_3); +lean_dec(x_3); +x_15 = l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_withLocation___spec__1___lambda__1(x_1, x_2, x_14, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_4); +return x_15; +} +} LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_withLocation___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: { @@ -825,6 +995,8 @@ l_Lean_Elab_Tactic_expandOptLocation___closed__1 = _init_l_Lean_Elab_Tactic_expa lean_mark_persistent(l_Lean_Elab_Tactic_expandOptLocation___closed__1); l_Lean_Elab_Tactic_expandOptLocation___closed__2 = _init_l_Lean_Elab_Tactic_expandOptLocation___closed__2(); lean_mark_persistent(l_Lean_Elab_Tactic_expandOptLocation___closed__2); +l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_withLocation___spec__1___lambda__1___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_withLocation___spec__1___lambda__1___closed__1(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_withLocation___spec__1___lambda__1___closed__1); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lean/Elab/Tactic/Meta.c b/stage0/stdlib/Lean/Elab/Tactic/Meta.c index 683f9e383389..af7c41a100e2 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Meta.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Meta.c @@ -39,7 +39,6 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_runTactic___sp lean_object* lean_st_ref_get(lean_object*, lean_object*); static lean_object* l_Lean_instantiateLCtxMVars___at_Lean_Elab_runTactic___spec__2___closed__5; lean_object* l_Lean_LocalContext_mkLetDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t); -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_runTactic(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Tactic_pruneSolvedGoals(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_usize_to_nat(size_t); @@ -48,6 +47,7 @@ static lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArra uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instantiateLCtxMVars___at_Lean_Elab_runTactic___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_TermElabM_run___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_runTactic___spec__7___boxed(lean_object*, lean_object*, 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_instantiateLCtxMVars___at_Lean_Elab_runTactic___spec__2___closed__7; @@ -141,7 +141,7 @@ lean_inc(x_18); x_19 = lean_ctor_get_uint8(x_15, sizeof(void*)*4); x_20 = lean_ctor_get_uint8(x_15, sizeof(void*)*4 + 1); lean_dec(x_15); -x_21 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_18, x_5, x_6, x_7, x_8, x_9); +x_21 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_18, x_5, x_6, x_7, x_8, x_9); x_22 = lean_ctor_get(x_21, 0); lean_inc(x_22); x_23 = lean_ctor_get(x_21, 1); @@ -169,13 +169,13 @@ lean_inc(x_31); x_32 = lean_ctor_get_uint8(x_15, sizeof(void*)*5); x_33 = lean_ctor_get_uint8(x_15, sizeof(void*)*5 + 1); lean_dec(x_15); -x_34 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_30, x_5, x_6, x_7, x_8, x_9); +x_34 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_30, x_5, x_6, x_7, x_8, x_9); 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 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_31, x_5, x_6, x_7, x_8, x_36); +x_37 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_31, x_5, x_6, x_7, x_8, x_36); x_38 = lean_ctor_get(x_37, 0); lean_inc(x_38); x_39 = lean_ctor_get(x_37, 1); @@ -971,7 +971,7 @@ lean_inc(x_16); x_17 = lean_ctor_get(x_15, 1); lean_inc(x_17); lean_dec(x_15); -x_18 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_14, x_2, x_3, x_4, x_5, x_17); +x_18 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_14, x_2, x_3, x_4, x_5, x_17); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); @@ -1204,7 +1204,7 @@ lean_inc(x_83); x_84 = lean_ctor_get(x_82, 1); lean_inc(x_84); lean_dec(x_82); -x_85 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_76, x_2, x_3, x_4, x_5, x_84); +x_85 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_76, x_2, x_3, x_4, x_5, x_84); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); diff --git a/stage0/stdlib/Lean/Elab/Tactic/Simp.c b/stage0/stdlib/Lean/Elab/Tactic/Simp.c index 75aed51a3bd4..f82d9ebd6837 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Simp.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Simp.c @@ -15,7 +15,6 @@ extern "C" { #endif LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_evalSimp(lean_object*); lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); -uint8_t l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(lean_object*, lean_object*); lean_object* l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Init_Util_0__outOfBounds___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_resolveGlobalConst___at_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*); @@ -297,6 +296,7 @@ static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_elabSimpArgs_ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabSimpArgs___lambda__2(lean_object*, uint8_t, uint8_t, 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_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_node2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_mkSimpContext___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*); +uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_simpLocation___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_Tactic_initFn____x40_Lean_Elab_Tactic_Simp___hyg_4750____closed__7; static lean_object* l_Lean_Elab_Tactic_elabDSimpConfigCore___closed__2; @@ -13345,7 +13345,7 @@ x_26 = lean_ctor_get(x_23, 1); x_27 = lean_ctor_get(x_8, 2); lean_inc(x_27); x_28 = l_Lean_Elab_Tactic_evalSimp___closed__1; -x_29 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_27, x_28); +x_29 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_27, x_28); lean_dec(x_27); if (x_29 == 0) { @@ -13379,7 +13379,7 @@ lean_dec(x_23); x_34 = lean_ctor_get(x_8, 2); lean_inc(x_34); x_35 = l_Lean_Elab_Tactic_evalSimp___closed__1; -x_36 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_34, x_35); +x_36 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_34, x_35); lean_dec(x_34); if (x_36 == 0) { @@ -13643,7 +13643,7 @@ lean_object* x_13; lean_object* x_14; uint8_t x_15; x_13 = lean_ctor_get(x_10, 2); lean_inc(x_13); x_14 = l_Lean_Elab_Tactic_evalSimp___closed__1; -x_15 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_13, x_14); +x_15 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_13, x_14); lean_dec(x_13); if (x_15 == 0) { @@ -14148,7 +14148,7 @@ lean_object* x_12; lean_object* x_13; uint8_t x_14; x_12 = lean_ctor_get(x_9, 2); lean_inc(x_12); x_13 = l_Lean_Elab_Tactic_evalSimp___closed__1; -x_14 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_12, x_13); +x_14 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_12, x_13); lean_dec(x_12); if (x_14 == 0) { diff --git a/stage0/stdlib/Lean/Elab/Term.c b/stage0/stdlib/Lean/Elab/Term.c index 84a6296e5ebe..4f1d443ab7cc 100644 --- a/stage0/stdlib/Lean/Elab/Term.c +++ b/stage0/stdlib/Lean/Elab/Term.c @@ -17,11 +17,9 @@ lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTr LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__6___rarg(lean_object*); static lean_object* l_Lean_Elab_throwAbortCommand___at_Lean_Elab_Term_ensureNoUnassignedMVars___spec__1___rarg___closed__2; static lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_Term_0__Lean_Elab_Term_isLambdaWithImplicit___spec__1___closed__2; -static lean_object* l_Lean_withOptProfile___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__16___closed__1; LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at_Lean_Elab_Term_synthesizeInstMVarCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); static lean_object* l_Lean_logAt___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__2___closed__1; -uint8_t l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_hasNoImplicitLambdaAnnotation___closed__1; static lean_object* l_Lean_Elab_Term_resolveId_x3f___lambda__2___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_withoutModifyingStateWithInfoAndMessages(lean_object*, lean_object*); @@ -34,7 +32,7 @@ static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_postponeElabTer LEAN_EXPORT lean_object* l_Lean_Elab_Term_instToStringSyntheticMVarKind___boxed(lean_object*); static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__1___closed__4; LEAN_EXPORT 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*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_collectUnassignedMVars_go___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_collectUnassignedMVars_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_EXPORT lean_object* l_Lean_Elab_Term_Context_autoBoundImplicitForbidden___default___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_ContextInfo_saveNoFileMap___at_Lean_Elab_Term_withoutModifyingElabMetaStateWithInfo___spec__4___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_isLetRecAuxMVar___lambda__2___closed__2; @@ -58,6 +56,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabUsingE static lean_object* l_Lean_Elab_Term_synthesizeInstMVarCore___lambda__3___closed__1; static lean_object* l_Lean_Elab_Term_mkAuxName___closed__1; 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*); +extern lean_object* l_Lean_profiler; lean_object* l_Lean_Core_getMaxHeartbeats(lean_object*); uint8_t l_Lean_Elab_isAbortExceptionId(lean_object*); static lean_object* l_Lean_Elab_Term_tryPostponeIfHasMVars___closed__2; @@ -74,7 +73,6 @@ LEAN_EXPORT lean_object* l_Lean_getDelayedMVarAssignment_x3f___at_Lean_Elab_Term LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__13(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_getMVarErrorInfo_x3f___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_collectUnassignedMVars___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_initFn____x40_Lean_Elab_Term___hyg_21681____closed__7; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_mkConst___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Term_expandDeclId___spec__6___closed__2; static lean_object* l_Lean_Elab_Term_instMetaEvalTermElabM___rarg___closed__2; @@ -96,7 +94,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_isLetRecAuxMVar(lean_object*, lean_obj LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Elab_Term_resolveName___spec__2(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_mkExplicitBinder___closed__14; LEAN_EXPORT lean_object* l_Lean_Elab_setInfoState___at_Lean_Elab_Term_SavedState_restore___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__21(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_DependsOn_dep_visitApp___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__20___boxed(lean_object*, lean_object*, lean_object*); 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*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__10(lean_object*, lean_object*, lean_object*); @@ -118,7 +115,6 @@ LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Elab_Term_resolveLocalName static lean_object* l_Lean_Elab_Term_mkExplicitBinder___closed__6; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_expandDeclId___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_Array_forInUnsafe_loop___at___private_Lean_Elab_Term_0__Lean_Elab_Term_decorateErrorMessageWithLambdaImplicitVars___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_withTraceNode___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__14___closed__2; lean_object* l_Lean_PersistentArray_toArray___rarg(lean_object*); static lean_object* l_Lean_Elab_Term_getFVarLocalDecl_x21___closed__2; static lean_object* l_Lean_Elab_Term_mkSaveInfoAnnotation___closed__1; @@ -165,7 +161,6 @@ uint64_t lean_uint64_mix_hash(uint64_t, uint64_t); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__41(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Elab_Term_registerMVarErrorInfo___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_List_elem___at_Lean_NameHashSet_insert___spec__2(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__21___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_elabTerm(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_resolveName_process___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_mkAuxName___at_Lean_Elab_Term_mkAuxName___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -179,28 +174,30 @@ lean_object* l_Lean_FileMap_toPosition(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___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_mkTypeMismatchError___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_ContextInfo_saveNoFileMap___at_Lean_Elab_Term_withoutModifyingElabMetaStateWithInfo___spec__4___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_mkCoe___lambda__1___closed__3; static lean_object* l_Lean_Elab_Term_synthesizeInstMVarCore___lambda__2___closed__3; lean_object* l_Lean_Meta_mkSyntheticSorry(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_elabImplicitLambdaAux___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_Term_0__Lean_Elab_Term_postponeElabTermCore___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_throwError___at_Lean_Elab_Term_throwTypeMismatchError___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_List_filterTR_loop___at_Lean_Elab_Term_resolveId_x3f___spec__2(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__15; +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623____closed__6; static lean_object* l_Lean_Elab_Term_MVarErrorInfo_logError___closed__7; static lean_object* l_Lean_Elab_Term_MVarErrorInfo_logError___closed__6; static lean_object* l_Lean_Elab_Term_mkCoe___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_instMonadTermElabM___closed__3; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__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_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_resolveName(lean_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_addCompletionInfo___at_Lean_Elab_Term_addDotCompletionInfo___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_MVarErrorInfo_logError___closed__8; LEAN_EXPORT lean_object* l_Lean_PersistentArray_findSomeRevM_x3f___at_Lean_Elab_Term_resolveLocalName___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getDelayedMVarAssignment_x3f___at_Lean_Elab_Term_isLetRecAuxMVar___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_mkCoe___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_mkCoe___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_forInUnsafe_loop___at_Lean_Elab_Term_expandDeclId___spec__4(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_getResetInfoTrees___at_Lean_Elab_Term_withoutModifyingElabMetaStateWithInfo___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__15___rarg___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__6; +double l_Lean_trace_profiler_threshold_getSecs(lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getId(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*); @@ -217,6 +214,7 @@ lean_object* l_Lean_Meta_trySynthInstance(lean_object*, lean_object*, lean_objec lean_object* l_Lean_MessageData_ofList(lean_object*); LEAN_EXPORT uint8_t l_Lean_Elab_Term_blockImplicitLambda(lean_object*); lean_object* l_Lean_PersistentArray_push___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Term_mkCoe___spec__4(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_array_push(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Term_0__Lean_Elab_Term_withoutModifyingStateWithInfoAndMessagesImpl___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_Elab_Term_registerSyntheticMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -239,11 +237,11 @@ LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_Term_0_ lean_object* l_ReaderT_instFunctorReaderT___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapM___at_Lean_Elab_Term_withoutModifyingElabMetaStateWithInfo___spec__5___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_ensureType___lambda__1___closed__1; -LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__22___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_array(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_traceAtCmdPos___boxed(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_addAutoBoundImplicits_go___spec__32(lean_object*, lean_object*, size_t, size_t); static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_isTypeAscription___closed__2; +static lean_object* l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwAutoBoundImplicitLocal___at_Lean_Elab_Term_resolveName_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_Elab_Term_instToStringMVarErrorKind___boxed(lean_object*); @@ -252,6 +250,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_resolveLocalName___lambda__1___boxed(l LEAN_EXPORT uint8_t l___private_Lean_Elab_Term_0__Lean_Elab_Term_isExplicit(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_TermElabM_toIO___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Term_expandDeclId___spec__2___closed__2; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__15; static lean_object* l_Lean_Elab_Term_mkTermElabAttributeUnsafe___closed__4; static lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Term_expandDeclId___spec__2___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_instMonadTermElabM; @@ -260,13 +259,11 @@ lean_object* l_Lean_Syntax_getArgs(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Elab_Term_withAuxDecl___spec__1___rarg___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_Term_0__Lean_Elab_Term_elabUsingElabFnsAux___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_withoutModifyingStateWithInfoAndMessages___rarg___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_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__12; LEAN_EXPORT lean_object* l_Lean_Elab_Term_hasNoImplicitLambdaAnnotation___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_findSomeRevMAux___at_Lean_Elab_Term_resolveLocalName___spec__4(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapMAux___at_Lean_Elab_Term_withoutModifyingElabMetaStateWithInfo___spec__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_mkAuxName___closed__2; lean_object* l_Lean_Meta_getMVarsAtDecl(lean_object*, lean_object*, 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*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_exceptionToSorry(lean_object*, lean_object*, 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_Lean_mkHashMap___at_Lean_Elab_Term_containsPendingMVar___spec__1(lean_object*); @@ -282,7 +279,6 @@ static lean_object* l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_ LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___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_Term_0__Lean_Elab_Term_decorateErrorMessageWithLambdaImplicitVars___closed__2; lean_object* l_Lean_Syntax_getPos_x3f(lean_object*, uint8_t); -LEAN_EXPORT lean_object* l_Lean_getBoolOption___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__17___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_bvar___override(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_expandDeclId___at_Lean_Elab_Term_expandDeclId___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_array_fget(lean_object*, lean_object*); @@ -303,24 +299,26 @@ lean_object* l_Lean_RBNode_insert___at_Lean_MVarIdSet_insert___spec__1(lean_obje LEAN_EXPORT lean_object* l_Lean_Elab_getResetInfoTrees___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__46(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_instMetaEvalTermElabM___rarg___closed__27; +static lean_object* l_Lean_Elab_Term_mkCoe___lambda__2___closed__1; static lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Term_expandDeclId___spec__6___lambda__2___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabImplicitLambda_loop___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_getResetInfoTrees___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_ppGoal(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_ReaderT_instMonadReaderT___rarg(lean_object*); LEAN_EXPORT 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_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__4___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___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_RBNode_forIn_visit___at_Lean_Elab_Term_logUnassignedUsingErrorInfos___spec__6___at_Lean_Elab_Term_logUnassignedUsingErrorInfos___spec__7(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_Array_forInUnsafe_loop___at_Lean_Elab_Term_expandDeclId___spec__4___closed__3; static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_3358____closed__1; lean_object* l_Lean_Expr_fvarId_x21(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__38___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_mkCoe___closed__3; static lean_object* l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__11; LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_useImplicitLambda___lambda__4(lean_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_logException___at_Lean_Elab_Term_exceptionToSorry___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_applyResult(lean_object*); lean_object* lean_environment_find(lean_object*, lean_object*); +uint8_t lean_float_decLt(double, double); static lean_object* l_Lean_Elab_Term_MVarErrorInfo_logError___closed__11; LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__22___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkConst___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -331,6 +329,7 @@ lean_object* l_Lean_removeLeadingSpaces(lean_object*); static lean_object* l_Lean_Elab_Term_mkExplicitBinder___closed__9; static lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_Term_0__Lean_Elab_Term_isLambdaWithImplicit___spec__1___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Elab_Term_withAuxDecl___spec__1___rarg(lean_object*, uint8_t, 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_Term_mkCoe___lambda__1___closed__5; LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Elab_Term_getMVarErrorInfo_x3f___spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_synthesizeInstMVarCore___lambda__2___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Term_liftLevelM___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -340,11 +339,9 @@ LEAN_EXPORT lean_object* l_Lean_HashMapImp_expand___at_Lean_Elab_Term_ContainsPe lean_object* l_Lean_mkIdentFrom(lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_List_foldlM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___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___private_Lean_Elab_Term_0__Lean_Elab_Term_dropTermParens___closed__1; -static lean_object* l_Lean_withTraceNode___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__14___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_ContextInfo_save___at_Lean_Elab_Term_withoutModifyingElabMetaStateWithInfo___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_instMonadBacktrackSavedStateTermElabM___closed__3; static lean_object* l_Lean_Elab_Term_instMetaEvalTermElabM___rarg___closed__21; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__11; LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_useImplicitLambda(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Linter_getLinterValue(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -362,7 +359,6 @@ static lean_object* l_Lean_Elab_Term_MVarErrorInfo_logError___closed__3; LEAN_EXPORT lean_object* l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___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_mkExplicitBinder___closed__16; lean_object* lean_io_get_num_heartbeats(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681_(lean_object*); uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); lean_object* l_Lean_Syntax_identComponents(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_logUnassignedUsingErrorInfos___closed__2; @@ -401,10 +397,9 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_instMonadMacroAdapterTermElabM___lambd lean_object* l_Lean_Elab_expandMacroImpl_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_string_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_instInhabitedTermElabM(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkCoe___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_Term_mkCoe___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_liftMacroM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__1___lambda__3___boxed(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_isLambdaWithImplicit___boxed(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__15(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Expr_instBEqExpr; uint8_t l_Lean_Syntax_isAtom(lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__49(lean_object*, lean_object*, size_t, size_t); @@ -425,12 +420,12 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabImplic static lean_object* l_Lean_Elab_Term_mkTermElabAttributeUnsafe___closed__12; static lean_object* l_Lean_Elab_Term_instMonadMacroAdapterTermElabM___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_State_letRecsToLift___default; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__10; static lean_object* l_Lean_Elab_Term_levelMVarToParam___closed__1; lean_object* l_instBEqProd___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Tactic_hashCacheKey____x40_Lean_Elab_Term___hyg_756____boxed(lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyM___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__29(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__52___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_Util_Trace_0__Lean_addTraceNodeCore___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___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_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabUsingElabFnsAux___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 uint8_t l_Lean_Elab_Term_hasNoImplicitLambdaAnnotation(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_expandDeclId___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -475,7 +470,6 @@ static lean_object* l_Lean_Elab_Term_synthesizeInstMVarCore___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Term_TermElabM_run(lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Elab_Term_0__Lean_Elab_Term_isTypeAscription(lean_object*); lean_object* l_instHashableArray___rarg___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__3; LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___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_logAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__8___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_Term_0__Lean_Elab_Term_isHole___closed__5; @@ -497,7 +491,6 @@ lean_object* l_ReaderT_instApplicativeReaderT___rarg(lean_object*); static lean_object* l_Lean_Elab_Term_instMonadMacroAdapterTermElabM___closed__3; LEAN_EXPORT uint8_t l_Lean_Elab_Term_Context_mayPostpone___default; static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabUsingElabFnsAux___lambda__1___closed__1; -static lean_object* l_Lean_Elab_Term_mkCoe___closed__5; LEAN_EXPORT lean_object* l_Lean_Option_set___at_Lean_Elab_Term_withoutMacroStackAtErr___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_levelMVarToParam___lambda__1___boxed(lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); @@ -523,7 +516,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_ static lean_object* l_Lean_Elab_Term_instMetaEvalTermElabM___rarg___closed__16; LEAN_EXPORT lean_object* l_Lean_log___at_Lean_Elab_Term_exceptionToSorry___spec__2(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_throwErrorIfErrors(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__14___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*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__51(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_instMetaEvalTermElabM___rarg___closed__10; static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabUsingElabFns___closed__4; @@ -531,17 +523,18 @@ static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_isExplicitApp__ static lean_object* l_Lean_Elab_Term_instToStringSyntheticMVarKind___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__8(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_Lean_PersistentArray_anyM___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__13(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__14; static lean_object* l_Lean_Elab_Term_isLetRecAuxMVar___closed__2; lean_object* l_Lean_throwError___at_Lean_Expr_abstractRangeM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1___closed__2; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___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_local_ctx_find(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_registerMVarErrorImplicitArgInfo(lean_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_MVarId_isAssigned___at_Lean_Elab_Term_synthesizeInstMVarCore___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_withTraceNode___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__14___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_List_foldl___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_ensureNoUnassignedMVars(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_synthesizeInstMVarCore___closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___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_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623____closed__7; lean_object* lean_st_ref_take(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_getLetRecsToLift(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_saveContext(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -568,6 +561,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_addDotComp static lean_object* l_Lean_Elab_Term_instMetaEvalTermElabM___rarg___closed__6; LEAN_EXPORT lean_object* l_List_foldlM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___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_Term_0__Lean_Elab_Term_withoutModifyingStateWithInfoAndMessagesImpl___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_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__4(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_EXPORT lean_object* l_Lean_Elab_throwPostpone___at_Lean_Elab_Term_tryPostpone___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Term_expandDeclId___spec__6___lambda__1___closed__2; 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*); @@ -595,10 +589,12 @@ static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___l LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__14(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_instInhabitedCache; static lean_object* l_Lean_Elab_Term_mkTypeMismatchError___closed__6; +lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_throwAbortCommand___at_Lean_Elab_Term_ensureNoUnassignedMVars___spec__1___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_withoutErrToSorryImp___rarg(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_traceAtCmdPos___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_initFn____x40_Lean_Elab_Term___hyg_21623____closed__1; LEAN_EXPORT 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_Lean_PersistentArray_anyMAux___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__6___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_instMonadBacktrackSavedStateTermElabM; @@ -609,6 +605,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_addAutoBoundImplicits_go___boxed(lean_ LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabUsingElabFnsAux___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___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Term_mkCoe___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_log___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__9(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_exprToSyntax(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_addAutoBoundImplicits_go___spec__51___closed__1; @@ -630,7 +627,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Term LEAN_EXPORT lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Term_expandDeclId___spec__2___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*); static lean_object* l_Lean_Elab_Term_mkExplicitBinder___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Term_resolveLocalName_go(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__14___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_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_Lean_Elab_Term_instMetaEvalTermElabM___rarg___closed__29; static lean_object* l_Lean_Elab_Term_instMetaEvalTermElabM___rarg___closed__9; @@ -639,8 +635,9 @@ LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at_Lean_Elab_Term_addAuto LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__52(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_withoutMacroStackAtErr___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Elab_Term_Context_saveRecAppSyntax___default; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__4; lean_object* l_Lean_Syntax_getKind(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_collectUnassignedMVars_go(lean_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_collectUnassignedMVars_go(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_isTacticOrPostponedHole_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_NameMap_contains___rarg(lean_object*, lean_object*); lean_object* l_Lean_PersistentArray_append___rarg(lean_object*, lean_object*); @@ -655,6 +652,7 @@ static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabImplicitLam LEAN_EXPORT lean_object* l_Lean_Elab_Term_addAutoBoundImplicits___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_resolveId_x3f___lambda__2___closed__4; LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___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_initFn____x40_Lean_Elab_Term___hyg_21623____closed__9; LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_decorateErrorMessageWithLambdaImplicitVars(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_withAuxDecl(lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Elab_Term_logUnassignedUsingErrorInfos___spec__6(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -663,13 +661,13 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_addTermInfo___lambda__1___boxed(lean_o lean_object* l_Lean_FVarId_getDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Cache_pre___default___closed__1; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Term_logUnassignedUsingErrorInfos___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681____closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkSaveInfoAnnotation(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_instMetaEvalTermElabM___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_withoutPostponing(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabImplicitLambdaAux___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___private_Lean_Elab_Term_0__Lean_Elab_Term_postponeElabTermCore___closed__3; LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapMAux___at___private_Lean_Elab_Term_0__Lean_Elab_Term_withoutModifyingStateWithInfoAndMessagesImpl___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___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Term_mkCoe___spec__2___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_synthesizeInstMVarCore___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_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_mkSyntheticSorryFor___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_replacePrefix(lean_object*, lean_object*, lean_object*); @@ -686,17 +684,17 @@ static lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Term_expandDeclId___sp lean_object* lean_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_withPushMacroExpansionStack(lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMap_insert___at_Lean_Elab_Term_ContainsPendingMVar_visit___spec__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___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_EXPORT lean_object* l_Lean_Elab_Term_throwMVarError(lean_object*); static lean_object* l_Lean_Elab_Term_mkTypeMismatchError___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Term_MVarErrorInfo_logError_appendExtra___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_mkCoe___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkFreshIdent___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkFreshBinderName___rarg(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_registerCustomErrorIfMVar___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_profileitM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__11___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_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*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_TermElabM_run_x27___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___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_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___boxed__const__1; LEAN_EXPORT lean_object* l_List_forM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___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_mkFreshBinderName___rarg___closed__1; @@ -722,7 +720,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_addDotCompletionInfo(lean_object*, lea LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__17(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__43(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_mkConst___closed__1; -static lean_object* l_Lean_Elab_Term_mkCoe___closed__7; static lean_object* l_Lean_Elab_Term_instMonadBacktrackSavedStateTermElabM___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_termElabAttribute; LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabImplicitLambdaAux___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -747,10 +744,12 @@ lean_object* l_Lean_RBNode_insert___at_Lean_MVarIdMap_insert___spec__1___rarg(le LEAN_EXPORT lean_object* l_Lean_PersistentArray_findSomeRevMAux___at_Lean_Elab_Term_resolveLocalName___spec__9___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Elab_ContextInfo_saveNoFileMap___at_Lean_Elab_Term_withoutModifyingElabMetaStateWithInfo___spec__4___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_getDelayedMVarAssignment_x3f___at_Lean_Elab_Term_ContainsPendingMVar_visit___spec__10___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_Term___hyg_21623____closed__3; lean_object* l_Lean_printTraces___at_Lean_Core_instMetaEvalCoreM___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MessageLog_forM___at_Lean_Elab_Term_instMetaEvalTermElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Expr_instHashableExpr; LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoLeaf___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__2(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_21623____closed__8; LEAN_EXPORT lean_object* l_Lean_Elab_Term_instMonadTermElabM___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_Lean_Elab_Term_throwErrorIfErrors___closed__1; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_synthesizeInstMVarCore___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -779,11 +778,13 @@ LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_addDotComp LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_tryPostponeIfHasMVars___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_Context_macroStack___default; LEAN_EXPORT lean_object* l_Lean_throwError___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*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__3; static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabUsingElabFns___closed__2; LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapMAux___at___private_Lean_Elab_Term_0__Lean_Elab_Term_withoutModifyingStateWithInfoAndMessagesImpl___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_Lean_Elab_Term_getMVarErrorInfo_x3f(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_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 uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_Term_collectUnassignedMVars_go___spec__2(lean_object*, lean_object*, size_t, size_t); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Term_mkCoe___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_Lean_Option_set___at_Lean_Elab_Term_withoutMacroStackAtErr___spec__1(lean_object*, lean_object*, uint8_t); uint64_t l_Lean_Expr_hash(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -798,13 +799,13 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_tryPostponeIfMVar(lean_object*, lean_o lean_object* l_Lean_Elab_expandDeclIdCore(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_tryPostponeIfHasMVars_x3f___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___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__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_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__15___rarg(lean_object*, lean_object*); lean_object* l_List_lengthTRAux___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_resolveId_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_Elab_Term_mkNoImplicitLambdaAnnotation(lean_object*); lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MetavarContext_levelMVarToParam(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_commitIfNoErrors_x3f___rarg(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_18351____closed__7; static lean_object* l_Lean_Elab_Term_synthesizeInstMVarCore___closed__3; LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Elab_Term_isLetRecAuxMVar___spec__4(lean_object*, lean_object*); lean_object* l_Lean_Core_withFreshMacroScope___rarg(lean_object*, lean_object*, lean_object*, lean_object*); @@ -831,21 +832,22 @@ LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_v lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_getLevelNames___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_isHole___closed__3; +static lean_object* l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__3___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Term_addTermInfo_x27___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_mkPatternWithRef(lean_object*, lean_object*); extern lean_object* l_Lean_warningAsError; static lean_object* l_Lean_Elab_Term_tryPostponeIfHasMVars___closed__1; extern lean_object* l_Lean_Elab_pp_macroStack; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__9; LEAN_EXPORT lean_object* l_Lean_Elab_Term_instMonadTermElabM___lambda__1(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_Term_0__Lean_Elab_Term_isNoImplicitLambda(lean_object*); static lean_object* l_Lean_Elab_Term_mkTermElabAttributeUnsafe___closed__10; LEAN_EXPORT lean_object* l_Lean_Elab_Term_isSaveInfoAnnotation_x3f___boxed(lean_object*); lean_object* l_Lean_Syntax_node2(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_mkCoe___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_withoutModifyingElabMetaStateWithInfo___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_EXPORT lean_object* l_Lean_mkConstWithLevelParams___at_Lean_Elab_Term_expandDeclId___spec__7___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_Term_0__Lean_Elab_Term_applyAttributesCore___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*); +uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623____closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_applyResult___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_instMetaEvalTermElabM___rarg___closed__30; LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Elab_Term_resolveLocalName___spec__1___boxed(lean_object*, lean_object*); @@ -855,6 +857,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_blockImplicitLambda___boxed(lean_objec LEAN_EXPORT lean_object* l_Lean_Elab_Term_Context_sectionFVars___default; LEAN_EXPORT lean_object* l_Lean_Elab_Term_instMetaEvalTermElabM___rarg(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__48(lean_object*, lean_object*, size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351_(lean_object*); lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_isLambdaWithImplicit___closed__2; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_expandDeclId___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -880,15 +883,12 @@ lean_object* l_Lean_Meta_saveState___rarg(lean_object*, lean_object*, lean_objec LEAN_EXPORT lean_object* l_Lean_Elab_Term_getLetRecsToLift___boxed(lean_object*); uint8_t l_Lean_Syntax_isIdent(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabImplicitLambdaAux(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_withTraceNode___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__14(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_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__10; LEAN_EXPORT 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*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_resolveId_x3f___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_postponeElabTerm(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__46___boxed(lean_object*, lean_object*); lean_object* l_Lean_LocalDecl_fvarId(lean_object*); LEAN_EXPORT lean_object* l_List_filterMap___at_Lean_Elab_Term_resolveName___spec__1(lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_isTacticBlock___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_postponeElabTerm___lambda__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_Elab_abortTermExceptionId; @@ -903,22 +903,24 @@ uint8_t l_Lean_LocalDecl_isAuxDecl(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_resolveLocalName_go___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Term_0__Lean_Elab_Term_withoutModifyingStateWithInfoAndMessagesImpl___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_EXPORT lean_object* l_Lean_resolveGlobalName___at_Lean_Elab_Term_resolveLocalName_loop___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withOptProfile___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__16(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at_Lean_Elab_Term_ContainsPendingMVar_visit___spec__8(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_liftLevelM(lean_object*); extern lean_object* l_Lean_Meta_instMonadMetaM; LEAN_EXPORT lean_object* l_Lean_logAt___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__2(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_Term_getSyntheticMVarDecl_x3f___boxed(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_addTraceNodeCore___at_Lean_Elab_Term_mkCoe___spec__5(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___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts(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_ContextInfo_saveNoFileMap___at_Lean_Elab_Term_withoutModifyingElabMetaStateWithInfo___spec__4___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_withAutoBoundImplicitForbiddenPred___rarg(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_instToStringMVarErrorKind(lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__6; static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_dropTermParens___closed__2; static lean_object* l_Lean_Elab_Term_instInhabitedMVarErrorInfo___closed__1; static lean_object* l_Lean_Elab_Term_instInhabitedMVarErrorKind___closed__1; lean_object* l_Lean_MacroScopesView_review(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_setInfoState___at_Lean_Elab_Term_SavedState_restore___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_throwMaxRecDepthAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__5___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_mkCoe___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_Term_0__Lean_Elab_Term_useImplicitLambda___lambda__1(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_mkElabAttribute___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -926,8 +928,8 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabUsingE LEAN_EXPORT uint8_t l___private_Lean_Elab_Term_0__Lean_Elab_Term_isTacticBlock(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__34(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_exprToSyntax___lambda__1___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__20(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_Lean_Elab_Term_containsPendingMVar___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__3___closed__2; LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Elab_Term_logUnassignedUsingErrorInfos___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*); uint8_t l_Lean_PersistentHashMap_contains___at_Lean_MVarId_isAssigned___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_Context_declName_x3f___default; @@ -944,7 +946,6 @@ LEAN_EXPORT lean_object* l_Lean_addDocString_x27___at_Lean_Elab_Term_expandDeclI static lean_object* l_Lean_Elab_Term_resolveName_process___closed__1; LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyM___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__37(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_instMetaEvalTermElabM___rarg___closed__4; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__4; lean_object* l_Array_append___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_resolveName___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_applyAttributesAt___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -957,11 +958,13 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_isMVarApp___boxed(lean_object*, lean_o lean_object* l_Lean_MessageData_ofExpr(lean_object*); 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_EXPORT lean_object* l_Lean_profileitM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__11(lean_object*); +static lean_object* l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__3___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Term_addAutoBoundImplicits_go(lean_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_findSomeRevM_x3f_find___at_Lean_Elab_Term_resolveLocalName___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_isTypeAscription___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_3358_(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__10___lambda__2(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_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Term_mkCoe___spec__6___boxed(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_DependsOn_dep_visitMain___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__35___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_Context_autoBoundImplicits___default; lean_object* l_List_tail_x21___rarg(lean_object*); @@ -971,10 +974,10 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_getDeclName_x3f___boxed(lean_object*, LEAN_EXPORT lean_object* l_Lean_Elab_Term_withAutoBoundImplicit___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___closed__1; LEAN_EXPORT uint8_t l_Lean_Elab_Term_Context_implicitLambda___default; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__5; static lean_object* l_Lean_Elab_Term_MVarErrorInfo_logError___closed__2; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_throwTypeMismatchError___spec__1(lean_object*); static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_isHole___closed__1; +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623____closed__4; uint8_t l_Lean_PersistentArray_anyM___at_Lean_MessageLog_hasErrors___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoTree___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); @@ -987,7 +990,6 @@ static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_isTacticBlock__ LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapMAux___at_Lean_Elab_Term_withoutModifyingElabMetaStateWithInfo___spec__11(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_MVarErrorInfo_logError_appendExtra(lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_setElabConfig(lean_object*); -static lean_object* l_Lean_withTraceNode___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__14___closed__3; static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_decorateErrorMessageWithLambdaImplicitVars___closed__4; LEAN_EXPORT lean_object* l_Lean_log___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__9___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_registerMVarErrorHoleInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -999,8 +1001,6 @@ static lean_object* l_Lean_Elab_Term_instMonadBacktrackSavedStateTermElabM___clo extern lean_object* l_Lean_Linter_deprecatedAttr; LEAN_EXPORT lean_object* l_Lean_Elab_Term_instMonadMacroAdapterTermElabM___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_isSaveInfoAnnotation_x3f(lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__14; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__8; LEAN_EXPORT lean_object* l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_withoutModifyingElabMetaStateWithInfo___spec__2___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwAbortTerm___at_Lean_Elab_Term_throwMVarError___spec__2___boxed(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_Term_expandDeclId___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_object*); @@ -1013,9 +1013,10 @@ LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_Elab_Term_addTermInfo___spec__ static lean_object* l_Lean_Elab_Term_instMetaEvalTermElabM___rarg___closed__1; LEAN_EXPORT lean_object* l_Array_findSomeRevM_x3f_find___at_Lean_Elab_Term_resolveLocalName___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_synthesizeInstMVarCore___lambda__3___closed__5; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___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_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Term_expandDeclId___spec__6___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_throwTypeMismatchError(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkCoe___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_Elab_Term_mkCoe___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_mkFreshIdent___rarg___lambda__1(lean_object*, lean_object*, uint8_t, lean_object*); static lean_object* l_Lean_Elab_Term_mkCoe___closed__1; LEAN_EXPORT lean_object* l_Array_findSomeRevM_x3f_find___at_Lean_Elab_Term_resolveLocalName___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1037,8 +1038,8 @@ LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Elab_Term_getSyntheticMVar LEAN_EXPORT uint8_t l_Lean_Elab_Term_Context_autoBoundImplicitForbidden___default(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLevel___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__1___closed__1; -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681____closed__4; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__4(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623_(lean_object*); static lean_object* l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__8; static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_isExplicit___closed__2; static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__6___rarg___closed__1; @@ -1058,24 +1059,21 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_synthesizeInstMVa LEAN_EXPORT uint8_t l_Lean_Elab_Term_LVal_isFieldName(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_throwMVarError___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_commitIfDidNotPostpone(lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681____closed__3; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__36(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Term_traceAtCmdPos___spec__1(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_RBNode_find___at_Lean_Elab_Term_getSyntheticMVarDecl_x3f___spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabUsingElabFns___closed__3; -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681____closed__8; LEAN_EXPORT lean_object* l_Lean_Elab_withSaveInfoContext___at_Lean_Elab_Term_withoutModifyingElabMetaStateWithInfo___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_instMetaEvalTermElabM___rarg___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_Term_0__Lean_Elab_Term_isTacticBlock___closed__5; LEAN_EXPORT lean_object* l_Lean_addDocString___at_Lean_Elab_Term_expandDeclId___spec__9___boxed(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_getFVarLocalDecl_x21___spec__1___closed__1; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_Term_addAutoBoundImplicits_x27___spec__1(lean_object*, size_t, size_t); -static double l_Lean_withSeconds___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__18___closed__1; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_synthesizeInstMVarCore___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_getBoolOption___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__17(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_instMetaEvalTermElabM___rarg___lambda__2___closed__1; LEAN_EXPORT 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*); +static lean_object* l_Lean_Elab_Term_mkCoe___lambda__1___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Term_ensureType(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_Term_0__Lean_Elab_Term_applyAttributesCore___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*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__7(lean_object*, lean_object*, size_t, size_t); @@ -1102,6 +1100,7 @@ LEAN_EXPORT lean_object* l_Lean_addDocString___at_Lean_Elab_Term_expandDeclId___ LEAN_EXPORT lean_object* l_Lean_Elab_Term_instMonadMacroAdapterTermElabM___lambda__3(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_withoutModifyingElabMetaStateWithInfo___spec__12___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_Meta_withLocalDecl___at_Lean_Elab_Term_withAuxDecl___spec__1(lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__5; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_withoutModifyingElabMetaStateWithInfo___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* l_Lean_Name_mkStr2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__11___boxed(lean_object*, lean_object*, lean_object*); @@ -1111,8 +1110,8 @@ static lean_object* l_Lean_Elab_Term_mkTermElabAttributeUnsafe___closed__6; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_expandDeclId___spec__4___closed__4; lean_object* l_Lean_ParametricAttribute_getParam_x3f___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_saveState___rarg___boxed(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_18409____closed__2; lean_object* l_Lean_indentExpr(lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___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*); uint8_t l_Lean_BinderInfo_isExplicit(uint8_t); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Term_0__Lean_Elab_Term_withoutModifyingStateWithInfoAndMessagesImpl___spec__9(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_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*); @@ -1132,15 +1131,13 @@ static lean_object* l_Lean_Elab_Term_instToStringMVarErrorKind___closed__2; lean_object* l_Lean_TagDeclarationExtension_tag(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Term_0__Lean_Elab_Term_withoutModifyingStateWithInfoAndMessagesImpl___spec__10(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_checkNotAlreadyDeclared___at_Lean_Elab_Term_expandDeclId___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_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409_(lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__13; LEAN_EXPORT lean_object* l_Lean_Elab_getResetInfoTrees___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__5___rarg(lean_object*, lean_object*); lean_object* l_Lean_Meta_mkHasTypeButIsExpectedMsg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Term_logUnassignedUsingErrorInfos___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_postponeElabTermCore(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_expandDeclId(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_initFn____x40_Lean_Elab_Term___hyg_21681____closed__6; -static lean_object* l_Lean_withOptProfile___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__16___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_instInhabitedMVarErrorInfo; LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_isExplicit___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_MVarErrorInfo_logError_addArgName(lean_object*, lean_object*, lean_object*); @@ -1149,14 +1146,16 @@ LEAN_EXPORT lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Term_expandDeclId static lean_object* l_Lean_Elab_Term_mkCoe___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_getDelayedMVarRoot___at_Lean_Elab_Term_isLetRecAuxMVar___spec__1(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_Term_0__Lean_Elab_Term_applyAttributesCore___spec__10___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_KVMap_getBool(lean_object*, lean_object*, uint8_t); LEAN_EXPORT 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*); static lean_object* l_Lean_Elab_Term_instMetaEvalTermElabM___rarg___closed__18; 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_localDeclDependsOn___at_Lean_Elab_Term_addAutoBoundImplicits_go___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*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__11; LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_isHole___boxed(lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623____closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Term_commitIfDidNotPostpone___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_instInhabitedLetRecToLift___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_checkIfShadowingStructureField___at_Lean_Elab_Term_expandDeclId___spec__3(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_synthesizeInstMVarCore___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1176,14 +1175,12 @@ LEAN_EXPORT uint8_t l_Lean_Elab_Term_Context_autoBoundImplicit___default; static lean_object* l_Lean_Elab_Term_exprToSyntax___lambda__1___closed__2; static lean_object* l_Lean_Elab_Term_instMetaEvalTermElabM___rarg___closed__15; LEAN_EXPORT lean_object* l_Array_contains___at_Lean_Elab_Term_logUnassignedUsingErrorInfos___spec__1___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681____closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_ensureType___lambda__1(lean_object*, 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*); lean_object* lean_nat_sub(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Elab_Term_Context_isNoncomputableSection___default; LEAN_EXPORT lean_object* l_Lean_getDelayedMVarAssignment_x3f___at_Lean_Elab_Term_isLetRecAuxMVar___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Elab_Term_0__Lean_Elab_Term_isLambdaWithImplicit(lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__7; static lean_object* l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__5___closed__1; LEAN_EXPORT lean_object* l_List_foldr___at_Lean_Elab_Term_isLetRecAuxMVar___spec__3___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_getAppFn(lean_object*); @@ -1202,11 +1199,11 @@ LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_v static lean_object* l_Lean_Elab_Term_instMonadMacroAdapterTermElabM___closed__1; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__51___closed__2; static lean_object* l_Lean_Elab_Term_instMonadTermElabM___closed__1; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__6; LEAN_EXPORT lean_object* l_Lean_HashMapImp_moveEntries___at_Lean_Elab_Term_ContainsPendingMVar_visit___spec__4(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_instInhabitedSyntheticMVarKind; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___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 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*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_withoutAutoBoundImplicit___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Elab_relaxedAutoImplicit; lean_object* l_Lean_Meta_InfoCacheKey_instHashableInfoCacheKey___boxed(lean_object*); @@ -1241,7 +1238,6 @@ static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_postponeElabTer extern lean_object* l_Lean_Elab_postponeExceptionId; LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at_Lean_Elab_Term_ContainsPendingMVar_visit___spec__8___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_contains___at_Lean_Elab_Term_collectUnassignedMVars_go___spec__1___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_protectedExt; LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkCoe(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_tryPostponeIfHasMVars___closed__6; @@ -1251,6 +1247,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_useImplici LEAN_EXPORT lean_object* l_Lean_Elab_expandDeclId___at_Lean_Elab_Term_expandDeclId___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_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_resolveName_x27___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Term_expandDeclId___spec__2___closed__1; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__9; static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_isNoImplicitLambda___closed__1; lean_object* l_List_reverse___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_withInfoTreeContext___at_Lean_Elab_Term_withMacroExpansion___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1280,6 +1277,7 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_instMetaE extern lean_object* l_Lean_Elab_autoBoundImplicitExceptionId; LEAN_EXPORT lean_object* l_Lean_Elab_Term_synthesizeInstMVarCore___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_Name_quickCmp(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__14(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__1; static lean_object* l_Lean_Elab_Term_mkTermElabAttributeUnsafe___closed__7; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_instMetaEvalTermElabM___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*); @@ -1301,10 +1299,10 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_addAutoBoundImpli uint8_t l_Lean_Expr_hasFVar(lean_object*); uint64_t l___private_Lean_Expr_0__Lean_hashMVarId____x40_Lean_Expr___hyg_2302_(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__11(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681____closed__2; LEAN_EXPORT lean_object* l_Lean_logAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__8(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_Term_0__Lean_Elab_Term_elabTermAux___lambda__7___closed__1; LEAN_EXPORT uint8_t l_Lean_Elab_Term_addAutoBoundImplicits_go___lambda__1(lean_object*); +static lean_object* l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__3___closed__4; static lean_object* l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__5___closed__2; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_Term_logUnassignedUsingErrorInfos___spec__2(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Elab_Term_resolveName___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*); @@ -1319,6 +1317,7 @@ LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at_Lean_Elab_Term_addAuto static lean_object* l_Lean_Elab_Term_mkExplicitBinder___closed__13; LEAN_EXPORT lean_object* l_Lean_Elab_Term_addTermInfo(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*); static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_postponeElabTermCore___closed__1; +extern lean_object* l_Lean_trace_profiler; 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*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_withoutErrToSorry___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_State_levelNames___default; @@ -1337,10 +1336,10 @@ LEAN_EXPORT lean_object* l_Lean_Elab_withSaveInfoContext___at_Lean_Elab_Term_wit lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapMAux___at___private_Lean_Elab_Term_0__Lean_Elab_Term_withoutModifyingStateWithInfoAndMessagesImpl___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_getSyntheticMVarDecl_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkCoe___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_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__13; +LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkCoe___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___private_Lean_Elab_Term_0__Lean_Elab_Term_mkSyntheticSorryFor(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_observing___rarg___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_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__2; lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT 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_anyMUnsafe_any___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__32___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -1354,14 +1353,16 @@ LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Term_addAutoBound static lean_object* l_Lean_Elab_throwAbortCommand___at_Lean_Elab_Term_ensureNoUnassignedMVars___spec__1___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__21___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_mkExplicitBinder___closed__5; +static lean_object* l_Lean_Elab_Term_mkCoe___lambda__1___closed__6; lean_object* l_Lean_Meta_mkFreshTypeMVar(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_universeConstraintsCheckpoint(lean_object*); static lean_object* l_Lean_Elab_Term_instToStringMVarErrorKind___closed__1; LEAN_EXPORT 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_Lean_isTracingEnabledFor___at_Lean_Elab_Term_traceAtCmdPos___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); -static lean_object* l_Lean_withTraceNode___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__14___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_withInfoTreeContext___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__4(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_18351____closed__12; +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Term_mkCoe___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l___private_Lean_Attributes_0__Lean_beqAttributeApplicationTime____x40_Lean_Attributes___hyg_23_(uint8_t, uint8_t); static lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Term_expandDeclId___spec__6___closed__1; static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__1___closed__3; @@ -1378,18 +1379,17 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkFreshBinderName___rarg___lambda__2(l lean_object* lean_array_get_size(lean_object*); lean_object* l_Lean_LocalDecl_toExpr(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_instMonadMacroAdapterTermElabM___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Term_mkCoe___spec__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_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*); static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_decorateErrorMessageWithLambdaImplicitVars___closed__1; static lean_object* l_Lean_Elab_Term_withAutoBoundImplicit___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_logException___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__22(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_instToStringLVal(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwAbortTerm___at_Lean_Elab_Term_ensureType___spec__1(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_elabUsingElabFnsAux___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_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__13; lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_withAutoBoundImplicitForbiddenPred(lean_object*); -static lean_object* l_Lean_Elab_Term_mkCoe___closed__6; static lean_object* l_Lean_Elab_Term_Context_autoBoundImplicits___default___closed__3; static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_decorateErrorMessageWithLambdaImplicitVars___closed__3; uint8_t lean_nat_dec_le(lean_object*, lean_object*); @@ -1429,6 +1429,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkFreshIdent___rarg(lean_object*, lean LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkConstWithLevelParams___at_Lean_Elab_Term_expandDeclId___spec__7(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_Term_0__Lean_Elab_Term_elabTermAux___spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__8; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__19(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoLeaf___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_instBEqArray___rarg___boxed(lean_object*, lean_object*, lean_object*); @@ -1436,6 +1437,7 @@ lean_object* l_List_filterTR_loop___at_Lean_resolveGlobalConstCore___spec__1(lea LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoTree___at_Lean_Elab_Term_addDotCompletionInfo___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_isTacticBlock___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabImplicitLambda(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static double l_Lean_withSeconds___at_Lean_Elab_Term_mkCoe___spec__3___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_MVarErrorInfo_logError(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isFVar(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_exprToSyntax___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1451,6 +1453,7 @@ static lean_object* l_Lean_Elab_Term_mkTermElabAttributeUnsafe___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Term_expandDeclId___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_EXPORT lean_object* l_Lean_Elab_withInfoContext_x27___at_Lean_Elab_Term_withInfoContext_x27___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_getStructureFieldsFlattened(lean_object*, lean_object*, uint8_t); +LEAN_EXPORT lean_object* l_Lean_withSeconds___at_Lean_Elab_Term_mkCoe___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_mvarId_x21(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabImplicitLambda_loop___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_Elab_throwPostpone___at_Lean_Elab_Term_tryPostpone___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1464,11 +1467,12 @@ lean_object* l_String_toSubstring_x27(lean_object*); static lean_object* l_Lean_Elab_Term_instInhabitedLetRecToLift___closed__9; static lean_object* l_Lean_Elab_Term_synthesizeInstMVarCore___lambda__3___closed__4; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_withoutModifyingElabMetaStateWithInfo___spec__9(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_mkCoe___lambda__2___closed__2; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); lean_object* l_Lean_MessageData_ofName(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_withSaveInfoContext___at___private_Lean_Elab_Term_0__Lean_Elab_Term_withoutModifyingStateWithInfoAndMessagesImpl___spec__1___rarg(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_getResetTraces___at_Lean_Elab_Term_mkCoe___spec__2___rarg___boxed(lean_object*, lean_object*); lean_object* lean_expr_instantiate1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withSeconds___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__18(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_expandDeclId___at_Lean_Elab_Term_expandDeclId___spec__1___closed__1; LEAN_EXPORT 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*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkFreshBinderName___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); @@ -1476,7 +1480,6 @@ static lean_object* l_Lean_Elab_Term_mkExplicitBinder___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_tryPostpone(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__38(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_findSomeRevM_x3f_find___at_Lean_Elab_Term_resolveLocalName___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__19(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_Lean_Elab_Term_removeSaveInfoAnnotation(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Term_0__Lean_Elab_Term_withoutModifyingStateWithInfoAndMessagesImpl___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_exprToSyntax___lambda__1___closed__3; @@ -1494,6 +1497,7 @@ LEAN_EXPORT uint8_t l_Lean_Elab_Term_Context_ignoreTCFailures___default; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__47___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_collectUnassignedMVars_go___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_Term_0__Lean_Elab_Term_isLambdaWithImplicit___spec__1___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___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_Lean_Elab_Term_isMVarApp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_applyVisibility___at_Lean_Elab_Term_expandDeclId___spec__5___closed__1; lean_object* l_Lean_Meta_whnfR(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1506,10 +1510,11 @@ LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Elab_Term_resolveName_x27___ LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_useImplicitLambda___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_withAutoBoundImplicit_loop___rarg___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_Term_0__Lean_Elab_Term_isLambdaWithImplicit___closed__1; -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681____closed__9; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_throwTypeMismatchError___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, 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*); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Elab_Term_mkCoe___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* l___private_Lean_Data_HashMap_0__Lean_numBucketsForCapacity(lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___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_ContainsPendingMVar_visit___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_hasExprMVar(lean_object*); lean_object* l_Lean_Meta_processPostponed(uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -3595,7 +3600,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddEr lean_object* x_10; lean_object* x_11; uint8_t x_12; x_10 = lean_ctor_get(x_7, 2); x_11 = l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1___closed__1; -x_12 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_10, x_11); +x_12 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_10, x_11); if (x_12 == 0) { lean_object* x_13; @@ -10892,7 +10897,7 @@ else { lean_object* x_265; uint8_t x_266; x_265 = l_Lean_logAt___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__2___closed__1; -x_266 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_262, x_265); +x_266 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_262, x_265); if (x_266 == 0) { x_12 = x_3; @@ -15207,7 +15212,7 @@ else { lean_object* x_265; uint8_t x_266; x_265 = l_Lean_logAt___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__2___closed__1; -x_266 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_262, x_265); +x_266 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_262, x_265); if (x_266 == 0) { x_12 = x_3; @@ -22030,159 +22035,193 @@ lean_dec(x_3); return x_12; } } -static lean_object* _init_l_Lean_Elab_Term_mkCoe___lambda__1___closed__1() { +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Term_mkCoe___spec__2___rarg(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("failed", 6); -return x_1; +lean_object* x_3; 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_st_ref_get(x_1, x_2); +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 = lean_ctor_get(x_4, 3); +lean_inc(x_6); +lean_dec(x_4); +x_7 = lean_st_ref_take(x_1, x_5); +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 = !lean_is_exclusive(x_8); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_11 = lean_ctor_get(x_8, 3); +lean_dec(x_11); +x_12 = l_Lean_Elab_Term_Context_autoBoundImplicits___default___closed__3; +lean_ctor_set(x_8, 3, x_12); +x_13 = lean_st_ref_set(x_1, x_8, x_9); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +lean_object* x_15; +x_15 = lean_ctor_get(x_13, 0); +lean_dec(x_15); +lean_ctor_set(x_13, 0, x_6); +return x_13; } +else +{ +lean_object* x_16; lean_object* x_17; +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); +lean_dec(x_13); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_6); +lean_ctor_set(x_17, 1, x_16); +return x_17; } -static lean_object* _init_l_Lean_Elab_Term_mkCoe___lambda__1___closed__2() { +} +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; +x_18 = lean_ctor_get(x_8, 0); +x_19 = lean_ctor_get(x_8, 1); +x_20 = lean_ctor_get(x_8, 2); +x_21 = lean_ctor_get(x_8, 4); +x_22 = lean_ctor_get(x_8, 5); +x_23 = lean_ctor_get(x_8, 6); +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_dec(x_8); +x_24 = l_Lean_Elab_Term_Context_autoBoundImplicits___default___closed__3; +x_25 = lean_alloc_ctor(0, 7, 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_24); +lean_ctor_set(x_25, 4, x_21); +lean_ctor_set(x_25, 5, x_22); +lean_ctor_set(x_25, 6, x_23); +x_26 = lean_st_ref_set(x_1, x_25, x_9); +x_27 = lean_ctor_get(x_26, 1); +lean_inc(x_27); +if (lean_is_exclusive(x_26)) { + lean_ctor_release(x_26, 0); + lean_ctor_release(x_26, 1); + x_28 = x_26; +} else { + lean_dec_ref(x_26); + x_28 = lean_box(0); +} +if (lean_is_scalar(x_28)) { + x_29 = lean_alloc_ctor(0, 2, 0); +} else { + x_29 = x_28; +} +lean_ctor_set(x_29, 0, x_6); +lean_ctor_set(x_29, 1, x_27); +return x_29; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Term_mkCoe___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; lean_object* x_2; -x_1 = l_Lean_Elab_Term_mkCoe___lambda__1___closed__1; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +lean_object* x_6; +x_6 = lean_alloc_closure((void*)(l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Term_mkCoe___spec__2___rarg___boxed), 2, 0); +return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkCoe___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) { +static double _init_l_Lean_withSeconds___at_Lean_Elab_Term_mkCoe___spec__3___closed__1() { _start: { -lean_object* x_12; +lean_object* x_1; uint8_t x_2; lean_object* x_3; double x_4; +x_1 = lean_unsigned_to_nat(1000000000u); +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_Lean_withSeconds___at_Lean_Elab_Term_mkCoe___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_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_9 = lean_io_mono_nanos_now(x_8); +x_10 = lean_ctor_get(x_9, 0); lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_2); -lean_inc(x_1); -x_12 = l_Lean_Meta_coerce_x3f(x_1, x_2, x_7, x_8, x_9, x_10, x_11); +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +lean_dec(x_9); +x_12 = lean_apply_7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_11); if (lean_obj_tag(x_12) == 0) { -lean_object* x_13; +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); -switch (lean_obj_tag(x_13)) { -case 0: -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); x_14 = lean_ctor_get(x_12, 1); lean_inc(x_14); lean_dec(x_12); -x_15 = l_Lean_Elab_Term_mkCoe___lambda__1___closed__2; -x_16 = l_Lean_throwError___at_Lean_Expr_abstractRangeM___spec__1(x_15, x_7, x_8, x_9, x_10, x_14); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -return x_16; -} -case 1: +x_15 = lean_io_mono_nanos_now(x_14); +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 0) { -uint8_t x_17; +lean_object* x_17; lean_object* x_18; uint8_t x_19; lean_object* x_20; double x_21; double x_22; double x_23; lean_object* x_24; lean_object* x_25; +x_17 = lean_ctor_get(x_15, 0); +x_18 = lean_nat_sub(x_17, x_10); lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_17 = !lean_is_exclusive(x_12); -if (x_17 == 0) -{ -lean_object* x_18; lean_object* x_19; -x_18 = lean_ctor_get(x_12, 0); +lean_dec(x_17); +x_19 = 0; +x_20 = lean_unsigned_to_nat(0u); +x_21 = l_Float_ofScientific(x_18, x_19, x_20); lean_dec(x_18); -x_19 = lean_ctor_get(x_13, 0); -lean_inc(x_19); -lean_dec(x_13); -lean_ctor_set(x_12, 0, x_19); -return x_12; +x_22 = l_Lean_withSeconds___at_Lean_Elab_Term_mkCoe___spec__3___closed__1; +x_23 = lean_float_div(x_21, x_22); +x_24 = lean_box_float(x_23); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_13); +lean_ctor_set(x_25, 1, x_24); +lean_ctor_set(x_15, 0, x_25); +return x_15; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_12, 1); -lean_inc(x_20); -lean_dec(x_12); -x_21 = lean_ctor_get(x_13, 0); -lean_inc(x_21); -lean_dec(x_13); -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; -} -} -default: -{ -lean_object* x_23; lean_object* x_24; uint8_t x_25; lean_object* x_26; 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_23 = lean_ctor_get(x_12, 1); -lean_inc(x_23); -lean_dec(x_12); -lean_inc(x_2); -x_24 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_24, 0, x_2); -x_25 = 2; -x_26 = lean_box(0); -lean_inc(x_7); -x_27 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkFreshExprMVarImpl(x_24, x_25, x_26, x_7, x_8, x_9, x_10, x_23); -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_28); -x_30 = l_Lean_Expr_mvarId_x21(x_28); -x_31 = lean_alloc_ctor(1, 4, 0); -lean_ctor_set(x_31, 0, x_3); -lean_ctor_set(x_31, 1, x_2); -lean_ctor_set(x_31, 2, x_1); -lean_ctor_set(x_31, 3, x_4); -x_32 = l_Lean_Elab_Term_registerSyntheticMVarWithCurrRef(x_30, x_31, x_5, x_6, x_7, x_8, x_9, x_10, x_29); +lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; lean_object* x_30; double x_31; double x_32; double x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_26 = lean_ctor_get(x_15, 0); +x_27 = lean_ctor_get(x_15, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_15); +x_28 = lean_nat_sub(x_26, x_10); lean_dec(x_10); -lean_dec(x_8); -lean_dec(x_7); -x_33 = !lean_is_exclusive(x_32); -if (x_33 == 0) -{ -lean_object* x_34; -x_34 = lean_ctor_get(x_32, 0); -lean_dec(x_34); -lean_ctor_set(x_32, 0, x_28); -return x_32; -} -else -{ -lean_object* x_35; lean_object* x_36; -x_35 = lean_ctor_get(x_32, 1); -lean_inc(x_35); -lean_dec(x_32); +lean_dec(x_26); +x_29 = 0; +x_30 = lean_unsigned_to_nat(0u); +x_31 = l_Float_ofScientific(x_28, x_29, x_30); +lean_dec(x_28); +x_32 = l_Lean_withSeconds___at_Lean_Elab_Term_mkCoe___spec__3___closed__1; +x_33 = lean_float_div(x_31, x_32); +x_34 = lean_box_float(x_33); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_13); +lean_ctor_set(x_35, 1, x_34); x_36 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_36, 0, x_28); -lean_ctor_set(x_36, 1, x_35); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_27); return x_36; } } -} -} else { uint8_t x_37; lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); x_37 = !lean_is_exclusive(x_12); if (x_37 == 0) { @@ -22204,1377 +22243,1451 @@ return x_40; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkCoe___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_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Elab_Term_mkCoe___spec__5(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_inc(x_4); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_13 = lean_alloc_closure((void*)(l_Lean_Elab_Term_mkCoe___lambda__1___boxed), 11, 4); -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_closure_set(x_13, 3, x_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); -x_14 = l_Lean_Elab_Term_withoutMacroStackAtErr___rarg(x_13, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_14) == 0) -{ -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_3); -lean_dec(x_2); -lean_dec(x_1); -return x_14; -} -else -{ -lean_object* x_15; -x_15 = lean_ctor_get(x_14, 0); +lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_13 = lean_st_ref_take(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); -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_14, 1); -lean_inc(x_16); -lean_dec(x_14); -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_1); -x_18 = lean_infer_type(x_1, x_8, x_9, x_10, x_11, x_16); -if (lean_obj_tag(x_18) == 0) -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -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_alloc_ctor(1, 1, 0); -lean_ctor_set(x_21, 0, x_17); -x_22 = l_Lean_Elab_Term_throwTypeMismatchError___rarg(x_3, x_2, x_19, x_1, x_4, x_21, x_6, x_7, x_8, x_9, x_10, x_11, x_20); -return x_22; -} -else -{ -uint8_t x_23; -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_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_23 = !lean_is_exclusive(x_18); -if (x_23 == 0) +lean_dec(x_13); +x_16 = !lean_is_exclusive(x_14); +if (x_16 == 0) { -return x_18; -} -else +lean_object* x_17; lean_object* x_18; lean_object* x_19; size_t x_20; size_t 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_17 = lean_ctor_get(x_14, 3); +x_18 = l_Lean_PersistentArray_toArray___rarg(x_17); +x_19 = lean_array_get_size(x_18); +x_20 = lean_usize_of_nat(x_19); +lean_dec(x_19); +x_21 = 0; +x_22 = l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___spec__1(x_20, x_21, x_18); +x_23 = lean_alloc_ctor(9, 3, 1); +lean_ctor_set(x_23, 0, x_2); +lean_ctor_set(x_23, 1, x_4); +lean_ctor_set(x_23, 2, x_22); +lean_ctor_set_uint8(x_23, sizeof(void*)*3, x_5); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_3); +lean_ctor_set(x_24, 1, x_23); +x_25 = l_Lean_PersistentArray_push___rarg(x_1, x_24); +lean_ctor_set(x_14, 3, x_25); +x_26 = lean_st_ref_set(x_11, x_14, x_15); +x_27 = !lean_is_exclusive(x_26); +if (x_27 == 0) { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_18, 0); -x_25 = lean_ctor_get(x_18, 1); -lean_inc(x_25); -lean_inc(x_24); -lean_dec(x_18); -x_26 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_26, 0, x_24); -lean_ctor_set(x_26, 1, x_25); +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; } -} -} else { -lean_object* x_27; lean_object* x_28; -lean_dec(x_15); -x_27 = lean_ctor_get(x_14, 1); -lean_inc(x_27); -lean_dec(x_14); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_1); -x_28 = lean_infer_type(x_1, x_8, x_9, x_10, x_11, x_27); -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, 0); -lean_inc(x_29); -x_30 = lean_ctor_get(x_28, 1); +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_28); +lean_dec(x_26); x_31 = lean_box(0); -x_32 = l_Lean_Elab_Term_throwTypeMismatchError___rarg(x_3, x_2, x_29, x_1, x_4, x_31, x_6, x_7, x_8, x_9, x_10, x_11, 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_30); return x_32; } -else -{ -uint8_t x_33; -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_3); -lean_dec(x_2); -lean_dec(x_1); -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_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; 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_33 = lean_ctor_get(x_14, 0); +x_34 = lean_ctor_get(x_14, 1); +x_35 = lean_ctor_get(x_14, 2); +x_36 = lean_ctor_get(x_14, 3); +x_37 = lean_ctor_get(x_14, 4); +x_38 = lean_ctor_get(x_14, 5); +x_39 = lean_ctor_get(x_14, 6); +lean_inc(x_39); +lean_inc(x_38); +lean_inc(x_37); +lean_inc(x_36); 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_inc(x_33); +lean_dec(x_14); +x_40 = l_Lean_PersistentArray_toArray___rarg(x_36); +x_41 = lean_array_get_size(x_40); +x_42 = lean_usize_of_nat(x_41); +lean_dec(x_41); +x_43 = 0; +x_44 = l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___spec__1(x_42, x_43, x_40); +x_45 = lean_alloc_ctor(9, 3, 1); +lean_ctor_set(x_45, 0, x_2); +lean_ctor_set(x_45, 1, x_4); +lean_ctor_set(x_45, 2, x_44); +lean_ctor_set_uint8(x_45, sizeof(void*)*3, x_5); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_3); +lean_ctor_set(x_46, 1, x_45); +x_47 = l_Lean_PersistentArray_push___rarg(x_1, x_46); +x_48 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_48, 0, x_33); +lean_ctor_set(x_48, 1, x_34); +lean_ctor_set(x_48, 2, x_35); +lean_ctor_set(x_48, 3, x_47); +lean_ctor_set(x_48, 4, x_37); +lean_ctor_set(x_48, 5, x_38); +lean_ctor_set(x_48, 6, x_39); +x_49 = lean_st_ref_set(x_11, x_48, x_15); +x_50 = lean_ctor_get(x_49, 1); +lean_inc(x_50); +if (lean_is_exclusive(x_49)) { + lean_ctor_release(x_49, 0); + lean_ctor_release(x_49, 1); + x_51 = x_49; +} else { + lean_dec_ref(x_49); + x_51 = lean_box(0); } +x_52 = lean_box(0); +if (lean_is_scalar(x_51)) { + x_53 = lean_alloc_ctor(0, 2, 0); +} else { + x_53 = x_51; } +lean_ctor_set(x_53, 0, x_52); +lean_ctor_set(x_53, 1, x_50); +return x_53; } } } -static lean_object* _init_l_Lean_Elab_Term_mkCoe___closed__1() { +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Term_mkCoe___spec__4(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_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_mkTermElabAttributeUnsafe___closed__9; -x_2 = l_Lean_Elab_Term_instToStringSyntheticMVarKind___closed__2; -x_3 = l_Lean_Name_mkStr2(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Elab_Term_mkCoe___closed__2() { -_start: +uint8_t x_13; +x_13 = !lean_is_exclusive(x_10); +if (x_13 == 0) { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("adding coercion for ", 20); -return x_1; -} +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_10, 5); +x_15 = l_Lean_replaceRef(x_3, x_14); +lean_dec(x_14); +lean_ctor_set(x_10, 5, x_15); +x_16 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_4, x_8, x_9, x_10, x_11, x_12); +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___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Elab_Term_mkCoe___spec__5(x_1, x_2, x_3, x_17, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_18); +lean_dec(x_10); +return x_19; } -static lean_object* _init_l_Lean_Elab_Term_mkCoe___closed__3() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_mkCoe___closed__2; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +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_20 = lean_ctor_get(x_10, 0); +x_21 = lean_ctor_get(x_10, 1); +x_22 = lean_ctor_get(x_10, 2); +x_23 = lean_ctor_get(x_10, 3); +x_24 = lean_ctor_get(x_10, 4); +x_25 = lean_ctor_get(x_10, 5); +x_26 = lean_ctor_get(x_10, 6); +x_27 = lean_ctor_get(x_10, 7); +x_28 = lean_ctor_get(x_10, 8); +x_29 = lean_ctor_get(x_10, 9); +x_30 = lean_ctor_get(x_10, 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_10); +x_31 = l_Lean_replaceRef(x_3, x_25); +lean_dec(x_25); +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_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_4, x_8, x_9, x_32, x_11, x_12); +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 = l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Elab_Term_mkCoe___spec__5(x_1, x_2, x_3, x_34, x_5, x_6, x_7, x_8, x_9, x_32, x_11, x_35); +lean_dec(x_32); +return x_36; +} +} +} +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Term_mkCoe___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: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_9; lean_object* x_10; +x_9 = lean_ctor_get(x_1, 0); +lean_inc(x_9); +x_10 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_8); +return x_10; +} +else +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_ctor_get(x_1, 0); +lean_inc(x_11); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_8); +return x_12; +} +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___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) { +_start: +{ +lean_object* x_9; +x_9 = lean_apply_7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_9) == 0) +{ +uint8_t x_10; +x_10 = !lean_is_exclusive(x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_ctor_get(x_9, 0); +x_12 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_9, 0, x_12); +return x_9; +} +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_13 = lean_ctor_get(x_9, 0); +x_14 = lean_ctor_get(x_9, 1); +lean_inc(x_14); +lean_inc(x_13); +lean_dec(x_9); +x_15 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_15, 0, 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 +{ +uint8_t x_17; +x_17 = !lean_is_exclusive(x_9); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; +x_18 = lean_ctor_get(x_9, 0); +x_19 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set_tag(x_9, 0); +lean_ctor_set(x_9, 0, x_19); +return x_9; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_20 = lean_ctor_get(x_9, 0); +x_21 = lean_ctor_get(x_9, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_9); +x_22 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_22, 0, 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; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___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; lean_object* x_16; lean_object* x_17; +lean_inc(x_12); +x_15 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Term_mkCoe___spec__4(x_1, x_2, x_3, x_6, x_4, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +x_16 = lean_ctor_get(x_15, 1); +lean_inc(x_16); +lean_dec(x_15); +x_17 = l_MonadExcept_ofExcept___at_Lean_Elab_Term_mkCoe___spec__6(x_5, x_8, x_9, x_10, x_11, x_12, x_13, x_16); +lean_dec(x_12); +lean_dec(x_8); +return x_17; } } -static lean_object* _init_l_Lean_Elab_Term_mkCoe___closed__4() { +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__3___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes(" : ", 3); +x_1 = l_Lean_profiler; return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_mkCoe___closed__5() { +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__3___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("[", 1); +return x_1; +} +} +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__3___closed__3() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_mkCoe___closed__4; +x_1 = l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__3___closed__2; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Term_mkCoe___closed__6() { +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__3___closed__4() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes(" =?= ", 5); +x_1 = lean_mk_string_from_bytes("s] ", 3); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_mkCoe___closed__7() { +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__3___closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_mkCoe___closed__6; +x_1 = l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__3___closed__4; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkCoe(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_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_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, uint8_t x_7, double 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_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_12 = l_Lean_Elab_Term_mkCoe___closed__1; -x_13 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Term_traceAtCmdPos___spec__1(x_12, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -x_14 = lean_ctor_get(x_13, 0); +lean_object* x_17; lean_object* x_18; +lean_dec(x_9); +x_17 = lean_ctor_get(x_14, 5); +lean_inc(x_17); +lean_inc(x_15); lean_inc(x_14); -x_15 = lean_unbox(x_14); -lean_dec(x_14); -if (x_15 == 0) +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_2); +x_18 = lean_apply_8(x_1, x_2, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_18) == 0) { -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_13, 1); -lean_inc(x_16); +lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +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_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__3___closed__1; +x_22 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_6, x_21); +lean_dec(x_6); +if (x_22 == 0) +{ +if (x_7 == 0) +{ +lean_object* x_23; lean_object* x_24; +x_23 = lean_box(0); +x_24 = l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__2(x_3, x_4, x_17, x_5, x_2, x_19, x_23, x_10, x_11, x_12, x_13, x_14, x_15, x_20); +lean_dec(x_15); lean_dec(x_13); -x_17 = lean_box(0); -x_18 = l_Lean_Elab_Term_mkCoe___lambda__2(x_2, x_1, x_4, x_3, x_17, x_5, x_6, x_7, x_8, x_9, x_10, x_16); -return x_18; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_2); +return x_24; } else { -lean_object* x_19; lean_object* x_20; -x_19 = lean_ctor_get(x_13, 1); -lean_inc(x_19); -lean_dec(x_13); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_2); -x_20 = lean_infer_type(x_2, x_7, x_8, x_9, x_10, x_19); -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; -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_2); -x_23 = l_Lean_MessageData_ofExpr(x_2); -x_24 = l_Lean_Elab_Term_mkCoe___closed__3; -x_25 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_25, 0, x_24); -lean_ctor_set(x_25, 1, x_23); -x_26 = l_Lean_Elab_Term_mkCoe___closed__5; -x_27 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_27, 0, x_25); -lean_ctor_set(x_27, 1, x_26); -x_28 = l_Lean_MessageData_ofExpr(x_21); +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; 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_25 = lean_float_to_string(x_8); +x_26 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_26, 0, x_25); +x_27 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_27, 0, x_26); +x_28 = l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__3___closed__3; x_29 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_29, 0, x_27); -lean_ctor_set(x_29, 1, x_28); -x_30 = l_Lean_Elab_Term_mkCoe___closed__7; +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_27); +x_30 = l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__3___closed__5; x_31 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_31, 0, x_29); lean_ctor_set(x_31, 1, x_30); -lean_inc(x_1); -x_32 = l_Lean_MessageData_ofExpr(x_1); -x_33 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_33, 0, x_31); -lean_ctor_set(x_33, 1, x_32); -x_34 = l_Lean_Elab_Term_MVarErrorInfo_logError___closed__4; -x_35 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -x_36 = l_Lean_addTrace___at_Lean_Elab_Term_traceAtCmdPos___spec__2(x_12, x_35, x_5, x_6, x_7, x_8, x_9, x_10, x_22); -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_mkCoe___lambda__2(x_2, x_1, x_4, x_3, x_37, x_5, x_6, x_7, x_8, x_9, x_10, x_38); -lean_dec(x_37); -return x_39; +x_32 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_19); +x_33 = l_Lean_Elab_Term_MVarErrorInfo_logError___closed__4; +x_34 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +x_35 = lean_box(0); +x_36 = l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__2(x_3, x_4, x_17, x_5, x_2, x_34, x_35, x_10, x_11, x_12, x_13, x_14, x_15, x_20); +lean_dec(x_15); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_2); +return x_36; +} } else { -uint8_t x_40; +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; 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_37 = lean_float_to_string(x_8); +x_38 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_38, 0, x_37); +x_39 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_39, 0, x_38); +x_40 = l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__3___closed__3; +x_41 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_39); +x_42 = l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__3___closed__5; +x_43 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_43, 0, x_41); +lean_ctor_set(x_43, 1, x_42); +x_44 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_44, 0, x_43); +lean_ctor_set(x_44, 1, x_19); +x_45 = l_Lean_Elab_Term_MVarErrorInfo_logError___closed__4; +x_46 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +x_47 = lean_box(0); +x_48 = l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__2(x_3, x_4, x_17, x_5, x_2, x_46, x_47, x_10, x_11, x_12, x_13, x_14, x_15, x_20); +lean_dec(x_15); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_2); +return x_48; +} +} +else +{ +uint8_t x_49; +lean_dec(x_17); +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); -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_20); -if (x_40 == 0) +x_49 = !lean_is_exclusive(x_18); +if (x_49 == 0) { -return x_20; +return x_18; } else { -lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_41 = lean_ctor_get(x_20, 0); -x_42 = lean_ctor_get(x_20, 1); -lean_inc(x_42); -lean_inc(x_41); -lean_dec(x_20); -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_object* x_50; lean_object* x_51; lean_object* x_52; +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_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_EXPORT lean_object* l_Lean_Elab_Term_mkCoe___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_mkCoe___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_6); -lean_dec(x_5); -return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkCoe___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) { +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__4___closed__1() { _start: { -lean_object* x_13; -x_13 = l_Lean_Elab_Term_mkCoe___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_5); -return x_13; +lean_object* x_1; +x_1 = l_Lean_trace_profiler; +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_ensureHasType(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_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_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t 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, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { -if (lean_obj_tag(x_1) == 0) -{ -lean_object* x_12; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_2); -lean_ctor_set(x_12, 1, x_11); -return x_12; -} -else -{ -lean_object* x_13; lean_object* x_14; -x_13 = lean_ctor_get(x_1, 0); -lean_inc(x_13); -lean_dec(x_1); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_2); -x_14 = lean_infer_type(x_2, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_14) == 0) -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -x_16 = lean_ctor_get(x_14, 1); +x_15 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Term_mkCoe___spec__2___rarg(x_13, x_14); +x_16 = lean_ctor_get(x_15, 0); lean_inc(x_16); -lean_dec(x_14); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_18 = lean_alloc_closure((void*)(l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__1), 8, 1); +lean_closure_set(x_18, 0, x_1); +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_13); -x_17 = l_Lean_Meta_isExprDefEq(x_15, x_13, x_7, x_8, x_9, x_10, x_16); -if (lean_obj_tag(x_17) == 0) -{ -lean_object* x_18; uint8_t x_19; -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -x_19 = lean_unbox(x_18); -lean_dec(x_18); -if (x_19 == 0) +x_19 = l_Lean_withSeconds___at_Lean_Elab_Term_mkCoe___spec__3(x_18, x_8, x_9, x_10, x_11, x_12, x_13, x_17); +if (lean_obj_tag(x_19) == 0) { -lean_object* x_20; lean_object* x_21; -x_20 = lean_ctor_get(x_17, 1); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; lean_object* x_69; uint8_t x_70; +x_20 = lean_ctor_get(x_19, 0); lean_inc(x_20); -lean_dec(x_17); -x_21 = l_Lean_Elab_Term_mkCoe(x_13, x_2, x_4, x_3, x_5, x_6, x_7, x_8, x_9, x_10, x_20); -return x_21; +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); +x_23 = lean_ctor_get(x_20, 1); +lean_inc(x_23); +lean_dec(x_20); +x_69 = l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__4___closed__1; +x_70 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_5, x_69); +if (x_70 == 0) +{ +uint8_t x_71; +x_71 = 0; +x_24 = x_71; +goto block_68; } else { -uint8_t x_22; +double x_72; double x_73; uint8_t x_74; +x_72 = l_Lean_trace_profiler_threshold_getSecs(x_5); +x_73 = lean_unbox_float(x_23); +x_74 = lean_float_decLt(x_72, x_73); +x_24 = x_74; +goto block_68; +} +block_68: +{ +if (x_6 == 0) +{ +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; +lean_dec(x_23); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +x_25 = lean_st_ref_take(x_13, x_21); +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_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; +x_29 = lean_ctor_get(x_26, 3); +x_30 = l_Lean_PersistentArray_append___rarg(x_16, x_29); +lean_ctor_set(x_26, 3, x_30); +x_31 = lean_st_ref_set(x_13, x_26, x_27); +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +lean_dec(x_31); +x_33 = l_MonadExcept_ofExcept___at_Lean_Elab_Term_mkCoe___spec__6(x_22, x_8, x_9, x_10, x_11, x_12, x_13, x_32); 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); -lean_dec(x_3); -x_22 = !lean_is_exclusive(x_17); -if (x_22 == 0) +lean_dec(x_22); +if (lean_obj_tag(x_33) == 0) { -lean_object* x_23; -x_23 = lean_ctor_get(x_17, 0); -lean_dec(x_23); -lean_ctor_set(x_17, 0, x_2); -return x_17; +uint8_t x_34; +x_34 = !lean_is_exclusive(x_33); +if (x_34 == 0) +{ +return x_33; } else { -lean_object* x_24; lean_object* x_25; -x_24 = lean_ctor_get(x_17, 1); -lean_inc(x_24); -lean_dec(x_17); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_2); -lean_ctor_set(x_25, 1, x_24); -return x_25; -} +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_33, 0); +x_36 = lean_ctor_get(x_33, 1); +lean_inc(x_36); +lean_inc(x_35); +lean_dec(x_33); +x_37 = lean_alloc_ctor(0, 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_26; -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); -x_26 = !lean_is_exclusive(x_17); -if (x_26 == 0) +uint8_t x_38; +x_38 = !lean_is_exclusive(x_33); +if (x_38 == 0) { -return x_17; +return x_33; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_17, 0); -x_28 = lean_ctor_get(x_17, 1); -lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_17); -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_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_33, 0); +x_40 = lean_ctor_get(x_33, 1); +lean_inc(x_40); +lean_inc(x_39); +lean_dec(x_33); +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_30; +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_42 = lean_ctor_get(x_26, 0); +x_43 = lean_ctor_get(x_26, 1); +x_44 = lean_ctor_get(x_26, 2); +x_45 = lean_ctor_get(x_26, 3); +x_46 = lean_ctor_get(x_26, 4); +x_47 = lean_ctor_get(x_26, 5); +x_48 = lean_ctor_get(x_26, 6); +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_inc(x_42); +lean_dec(x_26); +x_49 = l_Lean_PersistentArray_append___rarg(x_16, x_45); +x_50 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_50, 0, x_42); +lean_ctor_set(x_50, 1, x_43); +lean_ctor_set(x_50, 2, x_44); +lean_ctor_set(x_50, 3, x_49); +lean_ctor_set(x_50, 4, x_46); +lean_ctor_set(x_50, 5, x_47); +lean_ctor_set(x_50, 6, x_48); +x_51 = lean_st_ref_set(x_13, x_50, x_27); +x_52 = lean_ctor_get(x_51, 1); +lean_inc(x_52); +lean_dec(x_51); +x_53 = l_MonadExcept_ofExcept___at_Lean_Elab_Term_mkCoe___spec__6(x_22, x_8, x_9, x_10, x_11, x_12, x_13, x_52); 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); -lean_dec(x_3); -lean_dec(x_2); -x_30 = !lean_is_exclusive(x_14); -if (x_30 == 0) +lean_dec(x_22); +if (lean_obj_tag(x_53) == 0) { -return x_14; +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_54 = lean_ctor_get(x_53, 0); +lean_inc(x_54); +x_55 = lean_ctor_get(x_53, 1); +lean_inc(x_55); +if (lean_is_exclusive(x_53)) { + lean_ctor_release(x_53, 0); + lean_ctor_release(x_53, 1); + x_56 = x_53; +} else { + lean_dec_ref(x_53); + x_56 = lean_box(0); +} +if (lean_is_scalar(x_56)) { + x_57 = lean_alloc_ctor(0, 2, 0); +} else { + x_57 = x_56; +} +lean_ctor_set(x_57, 0, x_54); +lean_ctor_set(x_57, 1, x_55); +return x_57; } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_14, 0); -x_32 = lean_ctor_get(x_14, 1); -lean_inc(x_32); -lean_inc(x_31); -lean_dec(x_14); -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_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_58 = lean_ctor_get(x_53, 0); +lean_inc(x_58); +x_59 = lean_ctor_get(x_53, 1); +lean_inc(x_59); +if (lean_is_exclusive(x_53)) { + lean_ctor_release(x_53, 0); + lean_ctor_release(x_53, 1); + x_60 = x_53; +} else { + lean_dec_ref(x_53); + x_60 = lean_box(0); } +if (lean_is_scalar(x_60)) { + x_61 = lean_alloc_ctor(1, 2, 0); +} else { + x_61 = x_60; } +lean_ctor_set(x_61, 0, x_58); +lean_ctor_set(x_61, 1, x_59); +return x_61; } } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_mkSyntheticSorryFor___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: +else { -lean_object* x_9; -x_9 = l_Lean_Meta_mkSyntheticSorry(x_1, x_4, x_5, x_6, x_7, x_8); -return x_9; +lean_object* x_62; double x_63; lean_object* x_64; +x_62 = lean_box(0); +x_63 = lean_unbox_float(x_23); +lean_dec(x_23); +x_64 = l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__3(x_2, x_22, x_16, x_3, x_4, x_5, x_24, x_63, x_62, x_8, x_9, x_10, x_11, x_12, x_13, x_21); +return x_64; } } -static lean_object* _init_l___private_Lean_Elab_Term_0__Lean_Elab_Term_mkSyntheticSorryFor___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Elab_Term_0__Lean_Elab_Term_mkSyntheticSorryFor___lambda__1___boxed), 8, 0); -return x_1; +lean_object* x_65; double x_66; lean_object* x_67; +x_65 = lean_box(0); +x_66 = lean_unbox_float(x_23); +lean_dec(x_23); +x_67 = l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__3(x_2, x_22, x_16, x_3, x_4, x_5, x_24, x_66, x_65, x_8, x_9, x_10, x_11, x_12, x_13, x_21); +return x_67; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_mkSyntheticSorryFor(lean_object* x_1, lean_object* x_2, 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_Term_0__Lean_Elab_Term_mkSyntheticSorryFor___closed__1; -if (lean_obj_tag(x_1) == 0) +} +else { -uint8_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_10 = 0; -x_11 = lean_box(0); -lean_inc(x_4); -x_12 = l_Lean_Meta_mkFreshTypeMVar(x_10, x_11, 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); +uint8_t x_75; +lean_dec(x_16); +lean_dec(x_13); lean_dec(x_12); -x_15 = lean_apply_8(x_9, x_13, x_2, x_3, x_4, x_5, x_6, x_7, x_14); -return x_15; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +x_75 = !lean_is_exclusive(x_19); +if (x_75 == 0) +{ +return x_19; } else { -lean_object* x_16; lean_object* x_17; -x_16 = lean_ctor_get(x_1, 0); -lean_inc(x_16); -lean_dec(x_1); -x_17 = lean_apply_8(x_9, x_16, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -return x_17; -} +lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_76 = lean_ctor_get(x_19, 0); +x_77 = lean_ctor_get(x_19, 1); +lean_inc(x_77); +lean_inc(x_76); +lean_dec(x_19); +x_78 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_78, 0, x_76); +lean_ctor_set(x_78, 1, x_77); +return x_78; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_mkSyntheticSorryFor___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___private_Lean_Elab_Term_0__Lean_Elab_Term_mkSyntheticSorryFor___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_3); -lean_dec(x_2); -return x_9; } } -LEAN_EXPORT lean_object* l_Lean_log___at_Lean_Elab_Term_exceptionToSorry___spec__2(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_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___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) { _start: { -lean_object* x_10; lean_object* x_11; -x_10 = lean_ctor_get(x_7, 5); -x_11 = l_Lean_logAt___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__2(x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_11; -} -} -LEAN_EXPORT lean_object* l_Lean_Elab_logException___at_Lean_Elab_Term_exceptionToSorry___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_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_12 = lean_ctor_get(x_9, 2); +lean_inc(x_12); +lean_inc(x_1); +x_13 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Term_traceAtCmdPos___spec__1(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_unbox(x_14); +if (x_15 == 0) { -if (lean_obj_tag(x_1) == 0) +lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); +lean_dec(x_13); +x_17 = l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__4___closed__1; +x_18 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_12, x_17); +if (x_18 == 0) { -lean_object* x_9; lean_object* x_10; uint8_t x_11; lean_object* x_12; -x_9 = lean_ctor_get(x_1, 0); -lean_inc(x_9); -x_10 = lean_ctor_get(x_1, 1); -lean_inc(x_10); -lean_dec(x_1); -x_11 = 2; -x_12 = l_Lean_logAt___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__2(x_9, x_10, x_11, 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_object* x_19; +lean_dec(x_14); +lean_dec(x_12); lean_dec(x_2); -lean_dec(x_9); -return x_12; -} -else -{ -uint8_t x_13; -x_13 = !lean_is_exclusive(x_1); -if (x_13 == 0) +lean_dec(x_1); +x_19 = lean_apply_7(x_3, x_5, x_6, x_7, x_8, x_9, x_10, x_16); +if (lean_obj_tag(x_19) == 0) { -lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_14 = lean_ctor_get(x_1, 0); -x_15 = lean_ctor_get(x_1, 1); -lean_dec(x_15); -x_16 = l_Lean_Elab_isAbortExceptionId(x_14); -if (x_16 == 0) +uint8_t x_20; +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) { -lean_object* x_17; lean_object* x_18; -x_17 = lean_ctor_get(x_6, 5); -lean_inc(x_17); -x_18 = l_Lean_InternalExceptionId_getName(x_14, x_8); -lean_dec(x_14); -if (lean_obj_tag(x_18) == 0) +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; uint8_t x_26; lean_object* x_27; -lean_dec(x_17); -lean_free_object(x_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); -lean_dec(x_18); -x_21 = l_Lean_MessageData_ofName(x_19); -x_22 = l_Lean_Elab_logException___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__7___closed__2; -x_23 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_23, 0, x_22); -lean_ctor_set(x_23, 1, x_21); -x_24 = l_Lean_Elab_Term_MVarErrorInfo_logError___closed__4; -x_25 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_25, 0, x_23); -lean_ctor_set(x_25, 1, x_24); -x_26 = 2; -x_27 = l_Lean_log___at_Lean_Elab_Term_exceptionToSorry___spec__2(x_25, x_26, x_2, x_3, x_4, x_5, x_6, x_7, 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); -return x_27; +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(0, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +return x_23; +} } else { -uint8_t x_28; -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_28 = !lean_is_exclusive(x_18); -if (x_28 == 0) +uint8_t x_24; +x_24 = !lean_is_exclusive(x_19); +if (x_24 == 0) { -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_29 = lean_ctor_get(x_18, 0); -x_30 = lean_io_error_to_string(x_29); -x_31 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_31, 0, x_30); -x_32 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set_tag(x_1, 0); -lean_ctor_set(x_1, 1, x_32); -lean_ctor_set(x_1, 0, x_17); -lean_ctor_set(x_18, 0, x_1); -return x_18; +return x_19; } 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_18, 0); -x_34 = lean_ctor_get(x_18, 1); -lean_inc(x_34); -lean_inc(x_33); -lean_dec(x_18); -x_35 = lean_io_error_to_string(x_33); -x_36 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_36, 0, x_35); -x_37 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_37, 0, x_36); -lean_ctor_set_tag(x_1, 0); -lean_ctor_set(x_1, 1, x_37); -lean_ctor_set(x_1, 0, x_17); -x_38 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_38, 0, x_1); -lean_ctor_set(x_38, 1, x_34); -return x_38; +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_19, 0); +x_26 = lean_ctor_get(x_19, 1); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_19); +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_39; lean_object* x_40; -lean_free_object(x_1); +lean_object* x_28; uint8_t x_29; lean_object* x_30; +x_28 = lean_box(0); +x_29 = lean_unbox(x_14); 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_39 = lean_box(0); -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_39); -lean_ctor_set(x_40, 1, x_8); -return x_40; +x_30 = l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__4(x_3, x_2, x_1, x_4, x_12, x_29, x_28, x_5, x_6, x_7, x_8, x_9, x_10, x_16); +return x_30; } } else { -lean_object* x_41; uint8_t x_42; -x_41 = lean_ctor_get(x_1, 0); -lean_inc(x_41); -lean_dec(x_1); -x_42 = l_Lean_Elab_isAbortExceptionId(x_41); -if (x_42 == 0) -{ -lean_object* x_43; lean_object* x_44; -x_43 = lean_ctor_get(x_6, 5); -lean_inc(x_43); -x_44 = l_Lean_InternalExceptionId_getName(x_41, x_8); -lean_dec(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; lean_object* x_49; lean_object* x_50; lean_object* x_51; uint8_t x_52; lean_object* x_53; -lean_dec(x_43); -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_Lean_MessageData_ofName(x_45); -x_48 = l_Lean_Elab_logException___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__7___closed__2; -x_49 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_49, 0, x_48); -lean_ctor_set(x_49, 1, x_47); -x_50 = l_Lean_Elab_Term_MVarErrorInfo_logError___closed__4; -x_51 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_51, 0, x_49); -lean_ctor_set(x_51, 1, x_50); -x_52 = 2; -x_53 = l_Lean_log___at_Lean_Elab_Term_exceptionToSorry___spec__2(x_51, x_52, x_2, x_3, x_4, x_5, x_6, x_7, x_46); -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_53; +lean_object* x_31; lean_object* x_32; uint8_t x_33; lean_object* x_34; +x_31 = lean_ctor_get(x_13, 1); +lean_inc(x_31); +lean_dec(x_13); +x_32 = lean_box(0); +x_33 = lean_unbox(x_14); +lean_dec(x_14); +x_34 = l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__4(x_3, x_2, x_1, x_4, x_12, x_33, x_32, x_5, x_6, x_7, x_8, x_9, x_10, x_31); +return x_34; } -else +} +} +static lean_object* _init_l_Lean_Elab_Term_mkCoe___lambda__1___closed__1() { +_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_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_54 = lean_ctor_get(x_44, 0); -lean_inc(x_54); -x_55 = lean_ctor_get(x_44, 1); -lean_inc(x_55); -if (lean_is_exclusive(x_44)) { - lean_ctor_release(x_44, 0); - lean_ctor_release(x_44, 1); - x_56 = x_44; -} else { - lean_dec_ref(x_44); - x_56 = lean_box(0); +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("adding coercion for ", 20); +return x_1; } -x_57 = lean_io_error_to_string(x_54); -x_58 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_58, 0, x_57); -x_59 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_59, 0, x_58); -x_60 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_60, 0, x_43); -lean_ctor_set(x_60, 1, x_59); -if (lean_is_scalar(x_56)) { - x_61 = lean_alloc_ctor(1, 2, 0); -} else { - x_61 = x_56; } -lean_ctor_set(x_61, 0, x_60); -lean_ctor_set(x_61, 1, x_55); -return x_61; +static lean_object* _init_l_Lean_Elab_Term_mkCoe___lambda__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_Term_mkCoe___lambda__1___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } -else +static lean_object* _init_l_Lean_Elab_Term_mkCoe___lambda__1___closed__3() { +_start: { -lean_object* x_62; lean_object* x_63; -lean_dec(x_41); -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_63 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_63, 0, x_62); -lean_ctor_set(x_63, 1, x_8); -return x_63; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(" : ", 3); +return x_1; } } +static lean_object* _init_l_Lean_Elab_Term_mkCoe___lambda__1___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_Term_mkCoe___lambda__1___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} } +static lean_object* _init_l_Lean_Elab_Term_mkCoe___lambda__1___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(" =?= ", 5); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_exceptionToSorry(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +static lean_object* _init_l_Lean_Elab_Term_mkCoe___lambda__1___closed__6() { _start: { -lean_object* x_10; -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_10 = l___private_Lean_Elab_Term_0__Lean_Elab_Term_mkSyntheticSorryFor(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_10) == 0) +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_Term_mkCoe___lambda__1___closed__5; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkCoe___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; -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_logException___at_Lean_Elab_Term_exceptionToSorry___spec__1(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_12); -if (lean_obj_tag(x_13) == 0) +lean_object* x_11; +lean_inc(x_1); +x_11 = lean_infer_type(x_1, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_11) == 0) { -uint8_t x_14; -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) +uint8_t x_12; +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) { -lean_object* x_15; -x_15 = lean_ctor_get(x_13, 0); -lean_dec(x_15); -lean_ctor_set(x_13, 0, x_11); -return x_13; +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_13 = lean_ctor_get(x_11, 0); +x_14 = l_Lean_MessageData_ofExpr(x_1); +x_15 = l_Lean_Elab_Term_mkCoe___lambda__1___closed__2; +x_16 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_14); +x_17 = l_Lean_Elab_Term_mkCoe___lambda__1___closed__4; +x_18 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_17); +x_19 = l_Lean_MessageData_ofExpr(x_13); +x_20 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_20, 0, x_18); +lean_ctor_set(x_20, 1, x_19); +x_21 = l_Lean_Elab_Term_mkCoe___lambda__1___closed__6; +x_22 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +x_23 = l_Lean_MessageData_ofExpr(x_2); +x_24 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_24, 0, x_22); +lean_ctor_set(x_24, 1, x_23); +x_25 = l_Lean_Elab_Term_MVarErrorInfo_logError___closed__4; +x_26 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +lean_ctor_set(x_11, 0, x_26); +return x_11; } else { -lean_object* x_16; lean_object* x_17; -x_16 = lean_ctor_get(x_13, 1); -lean_inc(x_16); -lean_dec(x_13); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_11); -lean_ctor_set(x_17, 1, x_16); -return x_17; +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_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_27 = lean_ctor_get(x_11, 0); +x_28 = lean_ctor_get(x_11, 1); +lean_inc(x_28); +lean_inc(x_27); +lean_dec(x_11); +x_29 = l_Lean_MessageData_ofExpr(x_1); +x_30 = l_Lean_Elab_Term_mkCoe___lambda__1___closed__2; +x_31 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_29); +x_32 = l_Lean_Elab_Term_mkCoe___lambda__1___closed__4; +x_33 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +x_34 = l_Lean_MessageData_ofExpr(x_27); +x_35 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +x_36 = l_Lean_Elab_Term_mkCoe___lambda__1___closed__6; +x_37 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_37, 0, x_35); +lean_ctor_set(x_37, 1, x_36); +x_38 = l_Lean_MessageData_ofExpr(x_2); +x_39 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +x_40 = l_Lean_Elab_Term_MVarErrorInfo_logError___closed__4; +x_41 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 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_28); +return x_42; } } else { -uint8_t x_18; -lean_dec(x_11); -x_18 = !lean_is_exclusive(x_13); -if (x_18 == 0) +uint8_t x_43; +lean_dec(x_2); +lean_dec(x_1); +x_43 = !lean_is_exclusive(x_11); +if (x_43 == 0) { -return x_13; +return x_11; } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_13, 0); -x_20 = lean_ctor_get(x_13, 1); -lean_inc(x_20); -lean_inc(x_19); -lean_dec(x_13); -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; -} -} -} -else -{ -uint8_t x_22; -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_22 = !lean_is_exclusive(x_10); -if (x_22 == 0) -{ -return x_10; -} -else -{ -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_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_log___at_Lean_Elab_Term_exceptionToSorry___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_10; lean_object* x_11; -x_10 = lean_unbox(x_2); -lean_dec(x_2); -x_11 = l_Lean_log___at_Lean_Elab_Term_exceptionToSorry___spec__2(x_1, x_10, 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_11; +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_11, 0); +x_45 = lean_ctor_get(x_11, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_11); +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; } } -static lean_object* _init_l_Lean_Elab_throwPostpone___at_Lean_Elab_Term_tryPostpone___spec__1___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_Elab_Term_observing___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_throwPostpone___at_Lean_Elab_Term_tryPostpone___spec__1___rarg(lean_object* x_1) { +static lean_object* _init_l_Lean_Elab_Term_mkCoe___lambda__2___closed__1() { _start: { -lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Elab_throwPostpone___at_Lean_Elab_Term_tryPostpone___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_object* x_1; +x_1 = lean_mk_string_from_bytes("failed", 6); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Elab_throwPostpone___at_Lean_Elab_Term_tryPostpone___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) { +static lean_object* _init_l_Lean_Elab_Term_mkCoe___lambda__2___closed__2() { _start: { -lean_object* x_7; -x_7 = lean_alloc_closure((void*)(l_Lean_Elab_throwPostpone___at_Lean_Elab_Term_tryPostpone___spec__1___rarg), 1, 0); -return x_7; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_Term_mkCoe___lambda__2___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_tryPostpone(lean_object* x_1, lean_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_Term_mkCoe___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: { -uint8_t x_8; -x_8 = lean_ctor_get_uint8(x_1, sizeof(void*)*8); -lean_dec(x_1); -if (x_8 == 0) -{ -lean_object* x_9; lean_object* x_10; -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_7); -return x_10; -} -else +lean_object* x_12; +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_2); +lean_inc(x_1); +x_12 = l_Lean_Meta_coerce_x3f(x_1, x_2, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_12) == 0) { -lean_object* x_11; -x_11 = l_Lean_Elab_throwPostpone___at_Lean_Elab_Term_tryPostpone___spec__1___rarg(x_7); -return x_11; -} -} -} -LEAN_EXPORT lean_object* l_Lean_Elab_throwPostpone___at_Lean_Elab_Term_tryPostpone___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) { -_start: +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_7; -x_7 = l_Lean_Elab_throwPostpone___at_Lean_Elab_Term_tryPostpone___spec__1(x_1, x_2, x_3, x_4, x_5, x_6); -lean_dec(x_6); -lean_dec(x_5); +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -return x_7; -} +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = l_Lean_Elab_Term_mkCoe___lambda__2___closed__2; +x_16 = l_Lean_throwError___at_Lean_Expr_abstractRangeM___spec__1(x_15, x_7, x_8, x_9, x_10, x_14); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +return x_16; } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_tryPostpone___boxed(lean_object* x_1, 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: +case 1: { -lean_object* x_8; -x_8 = l_Lean_Elab_Term_tryPostpone(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_6); -lean_dec(x_5); +uint8_t x_17; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Lean_Elab_Term_isMVarApp(lean_object* x_1, lean_object* x_2, 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_Meta_whnfR(x_1, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_9) == 0) -{ -uint8_t x_10; -x_10 = !lean_is_exclusive(x_9); -if (x_10 == 0) +lean_dec(x_1); +x_17 = !lean_is_exclusive(x_12); +if (x_17 == 0) { -lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; -x_11 = lean_ctor_get(x_9, 0); -x_12 = l_Lean_Expr_getAppFn(x_11); -lean_dec(x_11); -x_13 = l_Lean_Expr_isMVar(x_12); -lean_dec(x_12); -x_14 = lean_box(x_13); -lean_ctor_set(x_9, 0, x_14); -return x_9; +lean_object* x_18; lean_object* x_19; +x_18 = lean_ctor_get(x_12, 0); +lean_dec(x_18); +x_19 = lean_ctor_get(x_13, 0); +lean_inc(x_19); +lean_dec(x_13); +lean_ctor_set(x_12, 0, x_19); +return x_12; } else { -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_15 = lean_ctor_get(x_9, 0); -x_16 = lean_ctor_get(x_9, 1); -lean_inc(x_16); -lean_inc(x_15); -lean_dec(x_9); -x_17 = l_Lean_Expr_getAppFn(x_15); -lean_dec(x_15); -x_18 = l_Lean_Expr_isMVar(x_17); -lean_dec(x_17); -x_19 = lean_box(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_16); -return x_20; +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_12, 1); +lean_inc(x_20); +lean_dec(x_12); +x_21 = lean_ctor_get(x_13, 0); +lean_inc(x_21); +lean_dec(x_13); +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 +default: { -uint8_t x_21; -x_21 = !lean_is_exclusive(x_9); -if (x_21 == 0) +lean_object* x_23; lean_object* x_24; uint8_t x_25; lean_object* x_26; 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_23 = lean_ctor_get(x_12, 1); +lean_inc(x_23); +lean_dec(x_12); +lean_inc(x_2); +x_24 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_24, 0, x_2); +x_25 = 2; +x_26 = lean_box(0); +lean_inc(x_7); +x_27 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkFreshExprMVarImpl(x_24, x_25, x_26, x_7, x_8, x_9, x_10, x_23); +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_28); +x_30 = l_Lean_Expr_mvarId_x21(x_28); +x_31 = lean_alloc_ctor(1, 4, 0); +lean_ctor_set(x_31, 0, x_3); +lean_ctor_set(x_31, 1, x_2); +lean_ctor_set(x_31, 2, x_1); +lean_ctor_set(x_31, 3, x_4); +x_32 = l_Lean_Elab_Term_registerSyntheticMVarWithCurrRef(x_30, x_31, x_5, x_6, x_7, x_8, x_9, x_10, x_29); +lean_dec(x_10); +lean_dec(x_8); +lean_dec(x_7); +x_33 = !lean_is_exclusive(x_32); +if (x_33 == 0) { -return x_9; +lean_object* x_34; +x_34 = lean_ctor_get(x_32, 0); +lean_dec(x_34); +lean_ctor_set(x_32, 0, x_28); +return x_32; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_9, 0); -x_23 = lean_ctor_get(x_9, 1); -lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_9); -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_object* x_35; lean_object* x_36; +x_35 = lean_ctor_get(x_32, 1); +lean_inc(x_35); +lean_dec(x_32); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_28); +lean_ctor_set(x_36, 1, x_35); +return x_36; } } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_isMVarApp___boxed(lean_object* x_1, lean_object* x_2, 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_Lean_Elab_Term_isMVarApp(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +uint8_t x_37; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -return x_9; +lean_dec(x_1); +x_37 = !lean_is_exclusive(x_12); +if (x_37 == 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_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; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_tryPostponeIfMVar(lean_object* x_1, lean_object* x_2, lean_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_Term_mkCoe___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: { -lean_object* x_9; +lean_object* x_13; +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_9 = l_Lean_Elab_Term_isMVarApp(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_9) == 0) +x_13 = l_Lean_Elab_Term_withoutMacroStackAtErr___rarg(x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_13) == 0) { -lean_object* x_10; uint8_t x_11; -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -x_11 = lean_unbox(x_10); +lean_dec(x_11); lean_dec(x_10); -if (x_11 == 0) -{ -uint8_t 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); -x_12 = !lean_is_exclusive(x_9); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; -x_13 = lean_ctor_get(x_9, 0); -lean_dec(x_13); -x_14 = lean_box(0); -lean_ctor_set(x_9, 0, x_14); -return x_9; +return x_13; } else { +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_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_9, 1); +x_15 = lean_ctor_get(x_13, 1); lean_inc(x_15); -lean_dec(x_9); -x_16 = lean_box(0); -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; -} +lean_dec(x_13); +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_2); +x_17 = lean_infer_type(x_2, x_8, x_9, x_10, x_11, x_15); +if (lean_obj_tag(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_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(1, 1, 0); +lean_ctor_set(x_20, 0, x_16); +x_21 = l_Lean_Elab_Term_throwTypeMismatchError___rarg(x_3, x_4, x_18, x_2, x_5, x_20, x_6, x_7, x_8, x_9, x_10, x_11, x_19); +return x_21; } else { -lean_object* x_18; lean_object* x_19; -x_18 = lean_ctor_get(x_9, 1); -lean_inc(x_18); +uint8_t x_22; +lean_dec(x_16); +lean_dec(x_11); +lean_dec(x_10); lean_dec(x_9); -x_19 = l_Lean_Elab_Term_tryPostpone(x_2, x_3, x_4, x_5, x_6, x_7, x_18); +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -return x_19; +lean_dec(x_3); +lean_dec(x_2); +x_22 = !lean_is_exclusive(x_17); +if (x_22 == 0) +{ +return x_17; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_17, 0); +x_24 = lean_ctor_get(x_17, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_17); +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 { -uint8_t x_20; +lean_object* x_26; lean_object* x_27; +lean_dec(x_14); +x_26 = lean_ctor_get(x_13, 1); +lean_inc(x_26); +lean_dec(x_13); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_2); +x_27 = lean_infer_type(x_2, x_8, x_9, x_10, x_11, x_26); +if (lean_obj_tag(x_27) == 0) +{ +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 = lean_box(0); +x_31 = l_Lean_Elab_Term_throwTypeMismatchError___rarg(x_3, x_4, x_28, x_2, x_5, x_30, x_6, x_7, x_8, x_9, x_10, x_11, x_29); +return x_31; +} +else +{ +uint8_t x_32; +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); -x_20 = !lean_is_exclusive(x_9); -if (x_20 == 0) +x_32 = !lean_is_exclusive(x_27); +if (x_32 == 0) { -return x_9; +return x_27; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_9, 0); -x_22 = lean_ctor_get(x_9, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_9); -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_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_EXPORT lean_object* l_Lean_Elab_Term_tryPostponeIfMVar___boxed(lean_object* x_1, lean_object* x_2, 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_Elab_Term_mkCoe___closed__1() { _start: { -lean_object* x_9; -x_9 = l_Lean_Elab_Term_tryPostponeIfMVar(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_3); -return x_9; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Term_mkTermElabAttributeUnsafe___closed__9; +x_2 = l_Lean_Elab_Term_instToStringSyntheticMVarKind___closed__2; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_tryPostponeIfNoneOrMVar(lean_object* x_1, lean_object* x_2, lean_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_Term_mkCoe(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_1) == 0) -{ -lean_object* x_9; -x_9 = l_Lean_Elab_Term_tryPostpone(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); -return x_9; +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; lean_object* x_17; +lean_inc(x_1); +lean_inc(x_2); +x_12 = lean_alloc_closure((void*)(l_Lean_Elab_Term_mkCoe___lambda__1___boxed), 10, 2); +lean_closure_set(x_12, 0, x_2); +lean_closure_set(x_12, 1, x_1); +lean_inc(x_3); +lean_inc(x_4); +lean_inc(x_1); +lean_inc(x_2); +x_13 = lean_alloc_closure((void*)(l_Lean_Elab_Term_mkCoe___lambda__2___boxed), 11, 4); +lean_closure_set(x_13, 0, x_2); +lean_closure_set(x_13, 1, x_1); +lean_closure_set(x_13, 2, x_4); +lean_closure_set(x_13, 3, x_3); +x_14 = lean_alloc_closure((void*)(l_Lean_Elab_Term_mkCoe___lambda__3), 12, 5); +lean_closure_set(x_14, 0, x_13); +lean_closure_set(x_14, 1, x_2); +lean_closure_set(x_14, 2, x_4); +lean_closure_set(x_14, 3, x_1); +lean_closure_set(x_14, 4, x_3); +x_15 = l_Lean_Elab_Term_mkCoe___closed__1; +x_16 = 1; +x_17 = l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1(x_15, x_12, x_14, x_16, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_17; } -else +} +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Term_mkCoe___spec__2___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_10; lean_object* x_11; -x_10 = lean_ctor_get(x_1, 0); -lean_inc(x_10); +lean_object* x_3; +x_3 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Term_mkCoe___spec__2___rarg(x_1, x_2); lean_dec(x_1); -x_11 = l_Lean_Elab_Term_tryPostponeIfMVar(x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -return x_11; -} +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_tryPostponeIfNoneOrMVar___boxed(lean_object* x_1, lean_object* x_2, lean_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___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Term_mkCoe___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_9; -x_9 = l_Lean_Elab_Term_tryPostponeIfNoneOrMVar(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_object* x_6; +x_6 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Term_mkCoe___spec__2(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); -return x_9; +lean_dec(x_2); +lean_dec(x_1); +return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_tryPostponeIfHasMVars_x3f___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_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Elab_Term_mkCoe___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) { _start: { -lean_object* x_10; lean_object* x_11; -x_10 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_10, 0, x_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); -return x_11; -} -} -LEAN_EXPORT lean_object* l_Lean_Elab_Term_tryPostponeIfHasMVars_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) { -_start: -{ -lean_object* x_9; -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_2); -lean_inc(x_1); -x_9 = l_Lean_Elab_Term_tryPostponeIfNoneOrMVar(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_9) == 0) -{ -if (lean_obj_tag(x_1) == 0) -{ -uint8_t x_10; -lean_dec(x_7); -lean_dec(x_6); +uint8_t x_13; lean_object* x_14; +x_13 = lean_unbox(x_5); lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_10 = !lean_is_exclusive(x_9); -if (x_10 == 0) -{ -lean_object* x_11; lean_object* x_12; -x_11 = lean_ctor_get(x_9, 0); +x_14 = l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Elab_Term_mkCoe___spec__5(x_1, x_2, x_3, x_4, x_13, x_6, x_7, x_8, x_9, x_10, x_11, x_12); lean_dec(x_11); -x_12 = lean_box(0); -lean_ctor_set(x_9, 0, x_12); -return x_9; -} -else -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_13 = lean_ctor_get(x_9, 1); -lean_inc(x_13); -lean_dec(x_9); -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_13); -return x_15; -} -} -else -{ -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_ctor_get(x_9, 1); -lean_inc(x_16); +lean_dec(x_10); lean_dec(x_9); -x_17 = lean_ctor_get(x_1, 0); -lean_inc(x_17); -lean_dec(x_1); -x_18 = l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(x_17, x_2, x_3, x_4, x_5, x_6, 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_inc(x_20); -lean_dec(x_18); -x_21 = l_Lean_Expr_hasExprMVar(x_19); -if (x_21 == 0) -{ -lean_object* x_22; lean_object* x_23; -x_22 = lean_box(0); -x_23 = l_Lean_Elab_Term_tryPostponeIfHasMVars_x3f___lambda__1(x_19, x_22, x_2, x_3, x_4, x_5, x_6, x_7, x_20); +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_23; +return x_14; } -else +} +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Term_mkCoe___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) { +_start: { -lean_object* x_24; -lean_dec(x_19); -x_24 = l_Lean_Elab_Term_tryPostpone(x_2, x_3, x_4, x_5, x_6, x_7, x_20); +uint8_t x_13; lean_object* x_14; +x_13 = lean_unbox(x_5); +lean_dec(x_5); +x_14 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Term_mkCoe___spec__4(x_1, x_2, x_3, x_4, x_13, 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_5); -lean_dec(x_4); -lean_dec(x_3); -if (lean_obj_tag(x_24) == 0) -{ -uint8_t x_25; -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) -{ -lean_object* x_26; lean_object* x_27; -x_26 = lean_ctor_get(x_24, 0); -lean_dec(x_26); -x_27 = lean_box(0); -lean_ctor_set(x_24, 0, x_27); -return x_24; -} -else -{ -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_24, 1); -lean_inc(x_28); -lean_dec(x_24); -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; -} -} -else -{ -uint8_t x_31; -x_31 = !lean_is_exclusive(x_24); -if (x_31 == 0) -{ -return x_24; -} -else -{ -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_24, 0); -x_33 = lean_ctor_get(x_24, 1); -lean_inc(x_33); -lean_inc(x_32); -lean_dec(x_24); -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; -} -} -} +return x_14; } } -else +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Term_mkCoe___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: { -uint8_t x_35; +lean_object* x_9; +x_9 = l_MonadExcept_ofExcept___at_Lean_Elab_Term_mkCoe___spec__6(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); @@ -23582,220 +23695,183 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_35 = !lean_is_exclusive(x_9); -if (x_35 == 0) -{ return x_9; } -else -{ -lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_36 = lean_ctor_get(x_9, 0); -x_37 = lean_ctor_get(x_9, 1); -lean_inc(x_37); -lean_inc(x_36); -lean_dec(x_9); -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; -} -} } -} -LEAN_EXPORT lean_object* l_Lean_Elab_Term_tryPostponeIfHasMVars_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, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__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, 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_10; -x_10 = l_Lean_Elab_Term_tryPostponeIfHasMVars_x3f___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_8); +uint8_t x_15; lean_object* x_16; +x_15 = lean_unbox(x_4); +lean_dec(x_4); +x_16 = l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___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); +lean_dec(x_13); +lean_dec(x_11); +lean_dec(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_10; +return x_16; } } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_tryPostponeIfHasMVars___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_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___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: { -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, 2); -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); -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; -} +uint8_t x_17; uint8_t x_18; double x_19; lean_object* x_20; +x_17 = lean_unbox(x_5); +lean_dec(x_5); +x_18 = lean_unbox(x_7); +lean_dec(x_7); +x_19 = lean_unbox_float(x_8); +lean_dec(x_8); +x_20 = l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__3(x_1, x_2, x_3, x_4, x_17, x_6, x_18, x_19, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +return x_20; } } -static lean_object* _init_l_Lean_Elab_Term_tryPostponeIfHasMVars___closed__1() { +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___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: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes(", expected type contains metavariables", 38); -return x_1; +uint8_t x_15; uint8_t x_16; lean_object* x_17; +x_15 = lean_unbox(x_4); +lean_dec(x_4); +x_16 = lean_unbox(x_6); +lean_dec(x_6); +x_17 = l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__4(x_1, x_2, x_3, x_15, x_5, x_16, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +return x_17; } } -static lean_object* _init_l_Lean_Elab_Term_tryPostponeIfHasMVars___closed__2() { +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___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) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_tryPostponeIfHasMVars___closed__1; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +uint8_t x_12; lean_object* x_13; +x_12 = lean_unbox(x_4); +lean_dec(x_4); +x_13 = l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1(x_1, x_2, x_3, x_12, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_13; } } -static lean_object* _init_l_Lean_Elab_Term_tryPostponeIfHasMVars___closed__3() { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkCoe___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_1; -x_1 = lean_mk_string_from_bytes("none", 4); -return x_1; +lean_object* x_11; +x_11 = l_Lean_Elab_Term_mkCoe___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_5); +lean_dec(x_4); +lean_dec(x_3); +return x_11; } } -static lean_object* _init_l_Lean_Elab_Term_tryPostponeIfHasMVars___closed__4() { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkCoe___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_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_tryPostponeIfHasMVars___closed__3; -x_2 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; +lean_object* x_12; +x_12 = l_Lean_Elab_Term_mkCoe___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_6); +lean_dec(x_5); +return x_12; } } -static lean_object* _init_l_Lean_Elab_Term_tryPostponeIfHasMVars___closed__5() { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_ensureHasType(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_tryPostponeIfHasMVars___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_Elab_Term_tryPostponeIfHasMVars___closed__6() { -_start: +if (lean_obj_tag(x_1) == 0) { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_tryPostponeIfHasMVars___closed__5; -x_2 = l_Lean_indentD(x_1); -return x_2; -} +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_5); +lean_dec(x_4); +lean_dec(x_3); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_2); +lean_ctor_set(x_12, 1, x_11); +return x_12; } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_tryPostponeIfHasMVars(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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: +else { -lean_object* x_10; +lean_object* x_13; lean_object* x_14; +x_13 = lean_ctor_get(x_1, 0); +lean_inc(x_13); +lean_dec(x_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); -lean_inc(x_4); -lean_inc(x_3); -lean_inc(x_1); -x_10 = l_Lean_Elab_Term_tryPostponeIfHasMVars_x3f(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_10) == 0) +lean_inc(x_2); +x_14 = lean_infer_type(x_2, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_14) == 0) { -lean_object* x_11; -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -if (lean_obj_tag(x_11) == 0) +lean_object* x_15; lean_object* x_16; lean_object* x_17; +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); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_13); +x_17 = l_Lean_Meta_isExprDefEq(x_15, x_13, x_7, x_8, x_9, x_10, x_16); +if (lean_obj_tag(x_17) == 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, 1); -lean_inc(x_12); -lean_dec(x_10); -x_13 = l_Lean_stringToMessageData(x_2); -lean_dec(x_2); -x_14 = l_Lean_Elab_Term_MVarErrorInfo_logError___closed__4; -x_15 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_15, 0, x_14); -lean_ctor_set(x_15, 1, x_13); -x_16 = l_Lean_Elab_Term_tryPostponeIfHasMVars___closed__2; -x_17 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_17, 0, x_15); -lean_ctor_set(x_17, 1, x_16); -if (lean_obj_tag(x_1) == 0) +lean_object* x_18; uint8_t x_19; +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_unbox(x_18); +lean_dec(x_18); +if (x_19 == 0) { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_18 = l_Lean_Elab_Term_tryPostponeIfHasMVars___closed__6; -x_19 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_19, 0, x_17); -lean_ctor_set(x_19, 1, x_18); -x_20 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_14); -x_21 = l_Lean_throwError___at_Lean_Elab_Term_tryPostponeIfHasMVars___spec__1(x_20, x_3, x_4, x_5, x_6, x_7, x_8, x_12); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_17, 1); +lean_inc(x_20); +lean_dec(x_17); +x_21 = l_Lean_Elab_Term_mkCoe(x_13, x_2, x_4, x_3, x_5, x_6, x_7, x_8, x_9, x_10, x_20); return x_21; } 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_1, 0); -lean_inc(x_22); -lean_dec(x_1); -x_23 = l_Lean_MessageData_ofExpr(x_22); -x_24 = l_Lean_indentD(x_23); -x_25 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_25, 0, x_17); -lean_ctor_set(x_25, 1, x_24); -x_26 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_14); -x_27 = l_Lean_throwError___at_Lean_Elab_Term_tryPostponeIfHasMVars___spec__1(x_26, x_3, x_4, x_5, x_6, x_7, x_8, x_12); +uint8_t x_22; +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); -return x_27; +lean_dec(x_3); +x_22 = !lean_is_exclusive(x_17); +if (x_22 == 0) +{ +lean_object* x_23; +x_23 = lean_ctor_get(x_17, 0); +lean_dec(x_23); +lean_ctor_set(x_17, 0, x_2); +return x_17; +} +else +{ +lean_object* x_24; lean_object* x_25; +x_24 = lean_ctor_get(x_17, 1); +lean_inc(x_24); +lean_dec(x_17); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_2); +lean_ctor_set(x_25, 1, x_24); +return x_25; +} } } else { -uint8_t x_28; +uint8_t x_26; +lean_dec(x_13); +lean_dec(x_10); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -23803,38 +23879,32 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_28 = !lean_is_exclusive(x_10); -if (x_28 == 0) +x_26 = !lean_is_exclusive(x_17); +if (x_26 == 0) { -lean_object* x_29; lean_object* x_30; -x_29 = lean_ctor_get(x_10, 0); -lean_dec(x_29); -x_30 = lean_ctor_get(x_11, 0); -lean_inc(x_30); -lean_dec(x_11); -lean_ctor_set(x_10, 0, x_30); -return x_10; +return x_17; } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_10, 1); -lean_inc(x_31); -lean_dec(x_10); -x_32 = lean_ctor_get(x_11, 0); -lean_inc(x_32); -lean_dec(x_11); -x_33 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_33, 0, x_32); -lean_ctor_set(x_33, 1, x_31); -return x_33; +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_17, 0); +x_28 = lean_ctor_get(x_17, 1); +lean_inc(x_28); +lean_inc(x_27); +lean_dec(x_17); +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_34; +uint8_t x_30; +lean_dec(x_13); +lean_dec(x_10); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -23842,50 +23912,1193 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_34 = !lean_is_exclusive(x_10); -if (x_34 == 0) +x_30 = !lean_is_exclusive(x_14); +if (x_30 == 0) { -return x_10; +return x_14; } else { -lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_35 = lean_ctor_get(x_10, 0); -x_36 = lean_ctor_get(x_10, 1); -lean_inc(x_36); -lean_inc(x_35); -lean_dec(x_10); -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; +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_14, 0); +x_32 = lean_ctor_get(x_14, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_14); +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_Lean_throwError___at_Lean_Elab_Term_tryPostponeIfHasMVars___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___private_Lean_Elab_Term_0__Lean_Elab_Term_mkSyntheticSorryFor___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; -x_9 = l_Lean_throwError___at_Lean_Elab_Term_tryPostponeIfHasMVars___spec__1(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); +x_9 = l_Lean_Meta_mkSyntheticSorry(x_1, x_4, x_5, x_6, x_7, x_8); return x_9; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_saveContext(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +static lean_object* _init_l___private_Lean_Elab_Term_0__Lean_Elab_Term_mkSyntheticSorryFor___closed__1() { _start: { -lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_8 = lean_ctor_get(x_5, 2); -x_9 = lean_ctor_get(x_5, 7); -x_10 = lean_st_ref_get(x_2, x_7); -x_11 = !lean_is_exclusive(x_10); -if (x_11 == 0) +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Elab_Term_0__Lean_Elab_Term_mkSyntheticSorryFor___lambda__1___boxed), 8, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_mkSyntheticSorryFor(lean_object* x_1, lean_object* x_2, 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_Term_0__Lean_Elab_Term_mkSyntheticSorryFor___closed__1; +if (lean_obj_tag(x_1) == 0) +{ +uint8_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_10 = 0; +x_11 = lean_box(0); +lean_inc(x_4); +x_12 = l_Lean_Meta_mkFreshTypeMVar(x_10, x_11, 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 = lean_apply_8(x_9, x_13, x_2, x_3, x_4, x_5, x_6, x_7, x_14); +return x_15; +} +else +{ +lean_object* x_16; lean_object* x_17; +x_16 = lean_ctor_get(x_1, 0); +lean_inc(x_16); +lean_dec(x_1); +x_17 = lean_apply_8(x_9, x_16, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +return x_17; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_mkSyntheticSorryFor___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___private_Lean_Elab_Term_0__Lean_Elab_Term_mkSyntheticSorryFor___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_3); +lean_dec(x_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_log___at_Lean_Elab_Term_exceptionToSorry___spec__2(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: +{ +lean_object* x_10; lean_object* x_11; +x_10 = lean_ctor_get(x_7, 5); +x_11 = l_Lean_logAt___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__2(x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_logException___at_Lean_Elab_Term_exceptionToSorry___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: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_9; lean_object* x_10; uint8_t x_11; lean_object* x_12; +x_9 = lean_ctor_get(x_1, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_1, 1); +lean_inc(x_10); +lean_dec(x_1); +x_11 = 2; +x_12 = l_Lean_logAt___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__2(x_9, x_10, x_11, 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); +lean_dec(x_9); +return x_12; +} +else +{ +uint8_t x_13; +x_13 = !lean_is_exclusive(x_1); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_14 = lean_ctor_get(x_1, 0); +x_15 = lean_ctor_get(x_1, 1); +lean_dec(x_15); +x_16 = l_Lean_Elab_isAbortExceptionId(x_14); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; +x_17 = lean_ctor_get(x_6, 5); +lean_inc(x_17); +x_18 = l_Lean_InternalExceptionId_getName(x_14, x_8); +lean_dec(x_14); +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; lean_object* x_27; +lean_dec(x_17); +lean_free_object(x_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); +lean_dec(x_18); +x_21 = l_Lean_MessageData_ofName(x_19); +x_22 = l_Lean_Elab_logException___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__7___closed__2; +x_23 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_21); +x_24 = l_Lean_Elab_Term_MVarErrorInfo_logError___closed__4; +x_25 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +x_26 = 2; +x_27 = l_Lean_log___at_Lean_Elab_Term_exceptionToSorry___spec__2(x_25, x_26, x_2, x_3, x_4, x_5, x_6, x_7, 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); +return x_27; +} +else +{ +uint8_t x_28; +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_28 = !lean_is_exclusive(x_18); +if (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_18, 0); +x_30 = lean_io_error_to_string(x_29); +x_31 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_31, 0, x_30); +x_32 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set_tag(x_1, 0); +lean_ctor_set(x_1, 1, x_32); +lean_ctor_set(x_1, 0, x_17); +lean_ctor_set(x_18, 0, x_1); +return x_18; +} +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_18, 0); +x_34 = lean_ctor_get(x_18, 1); +lean_inc(x_34); +lean_inc(x_33); +lean_dec(x_18); +x_35 = lean_io_error_to_string(x_33); +x_36 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_36, 0, x_35); +x_37 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set_tag(x_1, 0); +lean_ctor_set(x_1, 1, x_37); +lean_ctor_set(x_1, 0, x_17); +x_38 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_38, 0, x_1); +lean_ctor_set(x_38, 1, x_34); +return x_38; +} +} +} +else +{ +lean_object* x_39; lean_object* x_40; +lean_free_object(x_1); +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_39 = lean_box(0); +x_40 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_8); +return x_40; +} +} +else +{ +lean_object* x_41; uint8_t x_42; +x_41 = lean_ctor_get(x_1, 0); +lean_inc(x_41); +lean_dec(x_1); +x_42 = l_Lean_Elab_isAbortExceptionId(x_41); +if (x_42 == 0) +{ +lean_object* x_43; lean_object* x_44; +x_43 = lean_ctor_get(x_6, 5); +lean_inc(x_43); +x_44 = l_Lean_InternalExceptionId_getName(x_41, x_8); +lean_dec(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; lean_object* x_49; lean_object* x_50; lean_object* x_51; uint8_t x_52; lean_object* x_53; +lean_dec(x_43); +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_Lean_MessageData_ofName(x_45); +x_48 = l_Lean_Elab_logException___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__7___closed__2; +x_49 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_49, 0, x_48); +lean_ctor_set(x_49, 1, x_47); +x_50 = l_Lean_Elab_Term_MVarErrorInfo_logError___closed__4; +x_51 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_51, 0, x_49); +lean_ctor_set(x_51, 1, x_50); +x_52 = 2; +x_53 = l_Lean_log___at_Lean_Elab_Term_exceptionToSorry___spec__2(x_51, x_52, x_2, x_3, x_4, x_5, x_6, x_7, x_46); +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_53; +} +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_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_54 = lean_ctor_get(x_44, 0); +lean_inc(x_54); +x_55 = lean_ctor_get(x_44, 1); +lean_inc(x_55); +if (lean_is_exclusive(x_44)) { + lean_ctor_release(x_44, 0); + lean_ctor_release(x_44, 1); + x_56 = x_44; +} else { + lean_dec_ref(x_44); + x_56 = lean_box(0); +} +x_57 = lean_io_error_to_string(x_54); +x_58 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_58, 0, x_57); +x_59 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_59, 0, x_58); +x_60 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_60, 0, x_43); +lean_ctor_set(x_60, 1, x_59); +if (lean_is_scalar(x_56)) { + x_61 = lean_alloc_ctor(1, 2, 0); +} else { + x_61 = x_56; +} +lean_ctor_set(x_61, 0, x_60); +lean_ctor_set(x_61, 1, x_55); +return x_61; +} +} +else +{ +lean_object* x_62; lean_object* x_63; +lean_dec(x_41); +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_63 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_63, 0, x_62); +lean_ctor_set(x_63, 1, x_8); +return x_63; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_exceptionToSorry(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_10 = l___private_Lean_Elab_Term_0__Lean_Elab_Term_mkSyntheticSorryFor(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_10) == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; +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_logException___at_Lean_Elab_Term_exceptionToSorry___spec__1(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_12); +if (lean_obj_tag(x_13) == 0) +{ +uint8_t x_14; +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +lean_object* x_15; +x_15 = lean_ctor_get(x_13, 0); +lean_dec(x_15); +lean_ctor_set(x_13, 0, x_11); +return x_13; +} +else +{ +lean_object* x_16; lean_object* x_17; +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); +lean_dec(x_13); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_11); +lean_ctor_set(x_17, 1, x_16); +return x_17; +} +} +else +{ +uint8_t x_18; +lean_dec(x_11); +x_18 = !lean_is_exclusive(x_13); +if (x_18 == 0) +{ +return x_13; +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_13, 0); +x_20 = lean_ctor_get(x_13, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_13); +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; +} +} +} +else +{ +uint8_t x_22; +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_22 = !lean_is_exclusive(x_10); +if (x_22 == 0) +{ +return x_10; +} +else +{ +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_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_log___at_Lean_Elab_Term_exceptionToSorry___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_10; lean_object* x_11; +x_10 = lean_unbox(x_2); +lean_dec(x_2); +x_11 = l_Lean_log___at_Lean_Elab_Term_exceptionToSorry___spec__2(x_1, x_10, 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_11; +} +} +static lean_object* _init_l_Lean_Elab_throwPostpone___at_Lean_Elab_Term_tryPostpone___spec__1___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_Elab_Term_observing___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_throwPostpone___at_Lean_Elab_Term_tryPostpone___spec__1___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Lean_Elab_throwPostpone___at_Lean_Elab_Term_tryPostpone___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_throwPostpone___at_Lean_Elab_Term_tryPostpone___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; +x_7 = lean_alloc_closure((void*)(l_Lean_Elab_throwPostpone___at_Lean_Elab_Term_tryPostpone___spec__1___rarg), 1, 0); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_tryPostpone(lean_object* x_1, 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; +x_8 = lean_ctor_get_uint8(x_1, sizeof(void*)*8); +lean_dec(x_1); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; +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_7); +return x_10; +} +else +{ +lean_object* x_11; +x_11 = l_Lean_Elab_throwPostpone___at_Lean_Elab_Term_tryPostpone___spec__1___rarg(x_7); +return x_11; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_throwPostpone___at_Lean_Elab_Term_tryPostpone___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) { +_start: +{ +lean_object* x_7; +x_7 = l_Lean_Elab_throwPostpone___at_Lean_Elab_Term_tryPostpone___spec__1(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_Elab_Term_tryPostpone___boxed(lean_object* x_1, 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_Term_tryPostpone(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_isMVarApp(lean_object* x_1, lean_object* x_2, 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_Meta_whnfR(x_1, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_9) == 0) +{ +uint8_t x_10; +x_10 = !lean_is_exclusive(x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; +x_11 = lean_ctor_get(x_9, 0); +x_12 = l_Lean_Expr_getAppFn(x_11); +lean_dec(x_11); +x_13 = l_Lean_Expr_isMVar(x_12); +lean_dec(x_12); +x_14 = lean_box(x_13); +lean_ctor_set(x_9, 0, x_14); +return x_9; +} +else +{ +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_15 = lean_ctor_get(x_9, 0); +x_16 = lean_ctor_get(x_9, 1); +lean_inc(x_16); +lean_inc(x_15); +lean_dec(x_9); +x_17 = l_Lean_Expr_getAppFn(x_15); +lean_dec(x_15); +x_18 = l_Lean_Expr_isMVar(x_17); +lean_dec(x_17); +x_19 = lean_box(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_16); +return x_20; +} +} +else +{ +uint8_t x_21; +x_21 = !lean_is_exclusive(x_9); +if (x_21 == 0) +{ +return x_9; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_9, 0); +x_23 = lean_ctor_get(x_9, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_9); +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 lean_object* l_Lean_Elab_Term_isMVarApp___boxed(lean_object* x_1, lean_object* x_2, 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_Term_isMVarApp(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_3); +lean_dec(x_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_tryPostponeIfMVar(lean_object* x_1, lean_object* x_2, 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_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_9 = l_Lean_Elab_Term_isMVarApp(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_9) == 0) +{ +lean_object* x_10; uint8_t x_11; +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +x_11 = lean_unbox(x_10); +lean_dec(x_10); +if (x_11 == 0) +{ +uint8_t x_12; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +x_12 = !lean_is_exclusive(x_9); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; +x_13 = lean_ctor_get(x_9, 0); +lean_dec(x_13); +x_14 = lean_box(0); +lean_ctor_set(x_9, 0, x_14); +return x_9; +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_9, 1); +lean_inc(x_15); +lean_dec(x_9); +x_16 = lean_box(0); +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; +} +} +else +{ +lean_object* x_18; lean_object* x_19; +x_18 = lean_ctor_get(x_9, 1); +lean_inc(x_18); +lean_dec(x_9); +x_19 = l_Lean_Elab_Term_tryPostpone(x_2, x_3, x_4, x_5, x_6, x_7, x_18); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +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_2); +x_20 = !lean_is_exclusive(x_9); +if (x_20 == 0) +{ +return x_9; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_9, 0); +x_22 = lean_ctor_get(x_9, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_9); +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_Elab_Term_tryPostponeIfMVar___boxed(lean_object* x_1, lean_object* x_2, 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_Term_tryPostponeIfMVar(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_3); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_tryPostponeIfNoneOrMVar(lean_object* x_1, lean_object* x_2, 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: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_9; +x_9 = l_Lean_Elab_Term_tryPostpone(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); +return x_9; +} +else +{ +lean_object* x_10; lean_object* x_11; +x_10 = lean_ctor_get(x_1, 0); +lean_inc(x_10); +lean_dec(x_1); +x_11 = l_Lean_Elab_Term_tryPostponeIfMVar(x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +return x_11; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_tryPostponeIfNoneOrMVar___boxed(lean_object* x_1, lean_object* x_2, 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_Term_tryPostponeIfNoneOrMVar(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_3); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_tryPostponeIfHasMVars_x3f___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; lean_object* x_11; +x_10 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_10, 0, x_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); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_tryPostponeIfHasMVars_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) { +_start: +{ +lean_object* x_9; +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_2); +lean_inc(x_1); +x_9 = l_Lean_Elab_Term_tryPostponeIfNoneOrMVar(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_9) == 0) +{ +if (lean_obj_tag(x_1) == 0) +{ +uint8_t 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_10 = !lean_is_exclusive(x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_ctor_get(x_9, 0); +lean_dec(x_11); +x_12 = lean_box(0); +lean_ctor_set(x_9, 0, x_12); +return x_9; +} +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_ctor_get(x_9, 1); +lean_inc(x_13); +lean_dec(x_9); +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_13); +return x_15; +} +} +else +{ +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_ctor_get(x_9, 1); +lean_inc(x_16); +lean_dec(x_9); +x_17 = lean_ctor_get(x_1, 0); +lean_inc(x_17); +lean_dec(x_1); +x_18 = l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(x_17, x_2, x_3, x_4, x_5, x_6, 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_inc(x_20); +lean_dec(x_18); +x_21 = l_Lean_Expr_hasExprMVar(x_19); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; +x_22 = lean_box(0); +x_23 = l_Lean_Elab_Term_tryPostponeIfHasMVars_x3f___lambda__1(x_19, x_22, x_2, x_3, x_4, x_5, x_6, x_7, 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); +return x_23; +} +else +{ +lean_object* x_24; +lean_dec(x_19); +x_24 = l_Lean_Elab_Term_tryPostpone(x_2, x_3, x_4, x_5, x_6, x_7, x_20); +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_24) == 0) +{ +uint8_t x_25; +x_25 = !lean_is_exclusive(x_24); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; +x_26 = lean_ctor_get(x_24, 0); +lean_dec(x_26); +x_27 = lean_box(0); +lean_ctor_set(x_24, 0, x_27); +return x_24; +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_24, 1); +lean_inc(x_28); +lean_dec(x_24); +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; +} +} +else +{ +uint8_t x_31; +x_31 = !lean_is_exclusive(x_24); +if (x_31 == 0) +{ +return x_24; +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_24, 0); +x_33 = lean_ctor_get(x_24, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_24); +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_35; +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_35 = !lean_is_exclusive(x_9); +if (x_35 == 0) +{ +return x_9; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_9, 0); +x_37 = lean_ctor_get(x_9, 1); +lean_inc(x_37); +lean_inc(x_36); +lean_dec(x_9); +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; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_tryPostponeIfHasMVars_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, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_Elab_Term_tryPostponeIfHasMVars_x3f___lambda__1(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_throwError___at_Lean_Elab_Term_tryPostponeIfHasMVars___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_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, 2); +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); +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; +} +} +} +static lean_object* _init_l_Lean_Elab_Term_tryPostponeIfHasMVars___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(", expected type contains metavariables", 38); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_tryPostponeIfHasMVars___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_Term_tryPostponeIfHasMVars___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Elab_Term_tryPostponeIfHasMVars___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("none", 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_tryPostponeIfHasMVars___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_Term_tryPostponeIfHasMVars___closed__3; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Elab_Term_tryPostponeIfHasMVars___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_Term_tryPostponeIfHasMVars___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_Elab_Term_tryPostponeIfHasMVars___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_Term_tryPostponeIfHasMVars___closed__5; +x_2 = l_Lean_indentD(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_tryPostponeIfHasMVars(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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_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_10 = l_Lean_Elab_Term_tryPostponeIfHasMVars_x3f(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_10) == 0) +{ +lean_object* x_11; +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +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; +x_12 = lean_ctor_get(x_10, 1); +lean_inc(x_12); +lean_dec(x_10); +x_13 = l_Lean_stringToMessageData(x_2); +lean_dec(x_2); +x_14 = l_Lean_Elab_Term_MVarErrorInfo_logError___closed__4; +x_15 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_13); +x_16 = l_Lean_Elab_Term_tryPostponeIfHasMVars___closed__2; +x_17 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_17, 0, x_15); +lean_ctor_set(x_17, 1, x_16); +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_18 = l_Lean_Elab_Term_tryPostponeIfHasMVars___closed__6; +x_19 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_19, 1, x_18); +x_20 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_14); +x_21 = l_Lean_throwError___at_Lean_Elab_Term_tryPostponeIfHasMVars___spec__1(x_20, x_3, x_4, x_5, x_6, x_7, x_8, x_12); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_21; +} +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_1, 0); +lean_inc(x_22); +lean_dec(x_1); +x_23 = l_Lean_MessageData_ofExpr(x_22); +x_24 = l_Lean_indentD(x_23); +x_25 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_25, 0, x_17); +lean_ctor_set(x_25, 1, x_24); +x_26 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_14); +x_27 = l_Lean_throwError___at_Lean_Elab_Term_tryPostponeIfHasMVars___spec__1(x_26, x_3, x_4, x_5, x_6, x_7, x_8, x_12); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_27; +} +} +else +{ +uint8_t 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); +lean_dec(x_1); +x_28 = !lean_is_exclusive(x_10); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; +x_29 = lean_ctor_get(x_10, 0); +lean_dec(x_29); +x_30 = lean_ctor_get(x_11, 0); +lean_inc(x_30); +lean_dec(x_11); +lean_ctor_set(x_10, 0, x_30); +return x_10; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_10, 1); +lean_inc(x_31); +lean_dec(x_10); +x_32 = lean_ctor_get(x_11, 0); +lean_inc(x_32); +lean_dec(x_11); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_31); +return x_33; +} +} +} +else +{ +uint8_t x_34; +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_34 = !lean_is_exclusive(x_10); +if (x_34 == 0) +{ +return x_10; +} +else +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_10, 0); +x_36 = lean_ctor_get(x_10, 1); +lean_inc(x_36); +lean_inc(x_35); +lean_dec(x_10); +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; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_tryPostponeIfHasMVars___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: +{ +lean_object* x_9; +x_9 = l_Lean_throwError___at_Lean_Elab_Term_tryPostponeIfHasMVars___spec__1(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_saveContext(lean_object* x_1, 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; lean_object* x_10; uint8_t x_11; +x_8 = lean_ctor_get(x_5, 2); +x_9 = lean_ctor_get(x_5, 7); +x_10 = lean_st_ref_get(x_2, x_7); +x_11 = !lean_is_exclusive(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; x_12 = lean_ctor_get(x_10, 0); @@ -24269,7 +25482,7 @@ x_19 = l_Lean_Elab_Term_MVarErrorInfo_logError___closed__4; x_20 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_20, 0, x_19); lean_ctor_set(x_20, 1, x_18); -x_21 = l_Lean_Elab_Term_mkCoe___closed__5; +x_21 = l_Lean_Elab_Term_mkCoe___lambda__1___closed__4; x_22 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_22, 0, x_20); lean_ctor_set(x_22, 1, x_21); @@ -32351,7 +33564,7 @@ lean_inc(x_1); x_20 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_20, 0, x_1); lean_ctor_set(x_20, 1, x_19); -x_21 = l_Lean_Elab_Term_mkCoe___closed__5; +x_21 = l_Lean_Elab_Term_mkCoe___lambda__1___closed__4; x_22 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_22, 0, x_20); lean_ctor_set(x_22, 1, x_21); @@ -33633,502 +34846,59 @@ if (lean_is_exclusive(x_30)) { x_38 = x_30; } else { lean_dec_ref(x_30); - x_38 = lean_box(0); -} -x_39 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_39, 0, x_37); -if (lean_is_scalar(x_38)) { - x_40 = lean_alloc_ctor(1, 1, 0); -} else { - x_40 = x_38; -} -lean_ctor_set(x_40, 0, x_39); -x_41 = l_liftExcept___at_Lean_Elab_liftMacroM___spec__1(x_40, x_3, x_36); -lean_dec(x_40); -return x_41; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__1___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -uint8_t x_5; lean_object* x_6; lean_object* x_7; -x_5 = l_Lean_Environment_contains(x_1, x_2); -x_6 = lean_box(x_5); -x_7 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_7, 0, x_6); -lean_ctor_set(x_7, 1, x_4); -return x_7; -} -} -LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__1___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) { -_start: -{ -lean_object* x_7; lean_object* x_8; -x_7 = l_Lean_ResolveName_resolveNamespace(x_1, x_2, x_3, x_4); -x_8 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_8, 0, x_7); -lean_ctor_set(x_8, 1, x_6); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__1___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) { -_start: -{ -lean_object* x_7; lean_object* x_8; -x_7 = l_Lean_ResolveName_resolveGlobalName(x_1, x_2, x_3, x_4); -x_8 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_8, 0, x_7); -lean_ctor_set(x_8, 1, x_6); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___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_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); -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___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___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___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___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___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___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___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___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) -{ -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__2(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); -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_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; lean_object* x_65; -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); -x_56 = lean_ctor_get(x_38, 6); -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_38); -x_57 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_57, 0, x_51); -lean_ctor_set(x_57, 1, x_36); -lean_ctor_set(x_57, 2, x_52); -lean_ctor_set(x_57, 3, x_53); -lean_ctor_set(x_57, 4, x_54); -lean_ctor_set(x_57, 5, x_55); -lean_ctor_set(x_57, 6, x_56); -x_58 = lean_st_ref_set(x_7, x_57, x_39); -x_59 = lean_ctor_get(x_58, 1); -lean_inc(x_59); -lean_dec(x_58); -x_60 = lean_ctor_get(x_35, 1); -lean_inc(x_60); -lean_dec(x_35); -x_61 = l_List_reverse___rarg(x_60); -x_62 = l_List_forM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__2(x_61, x_2, x_3, x_4, x_5, x_6, x_7, 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_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_62); - x_64 = lean_box(0); -} -if (lean_is_scalar(x_64)) { - x_65 = lean_alloc_ctor(0, 2, 0); -} else { - x_65 = x_64; -} -lean_ctor_set(x_65, 0, x_34); -lean_ctor_set(x_65, 1, x_63); -return x_65; -} -} -else -{ -lean_object* x_66; -x_66 = lean_ctor_get(x_33, 0); -lean_inc(x_66); -lean_dec(x_33); -if (lean_obj_tag(x_66) == 0) -{ -lean_object* x_67; lean_object* x_68; lean_object* x_69; uint8_t x_70; -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 = l_Lean_maxRecDepthErrorMessage; -x_70 = lean_string_dec_eq(x_68, x_69); -if (x_70 == 0) -{ -lean_object* x_71; lean_object* x_72; lean_object* x_73; -x_71 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_71, 0, x_68); -x_72 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_72, 0, x_71); -x_73 = l_Lean_throwErrorAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__3(x_67, x_72, x_2, x_3, x_4, x_5, x_6, x_7, x_27); -return x_73; -} -else -{ -lean_object* x_74; -lean_dec(x_68); -x_74 = l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__5(x_67, 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_74; -} -} -else -{ -lean_object* x_75; -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_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__6___rarg(x_27); -return x_75; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Elab_throwPostpone___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__7___rarg(lean_object* x_1) { -_start: -{ -lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Elab_throwPostpone___at_Lean_Elab_Term_tryPostpone___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_throwPostpone___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___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_7; -x_7 = lean_alloc_closure((void*)(l_Lean_Elab_throwPostpone___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__7___rarg), 1, 0); -return x_7; -} -} -LEAN_EXPORT lean_object* l_liftExcept___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -if (lean_obj_tag(x_1) == 0) -{ -lean_object* x_4; lean_object* x_5; -x_4 = lean_ctor_get(x_1, 0); -lean_inc(x_4); -x_5 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_5, 0, x_4); -lean_ctor_set(x_5, 1, x_3); -return x_5; -} -else -{ -lean_object* x_6; lean_object* x_7; -x_6 = lean_ctor_get(x_1, 0); -lean_inc(x_6); -x_7 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_7, 0, x_6); -lean_ctor_set(x_7, 1, x_3); -return x_7; -} -} -} -LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___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) { -_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, 2); -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); -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_Term_0__Lean_Elab_Term_elabTermAux___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) { -_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_Term_0__Lean_Elab_Term_elabTermAux___spec__11(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_Term_0__Lean_Elab_Term_elabTermAux___spec__11(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; + x_38 = lean_box(0); +} +x_39 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_39, 0, x_37); +if (lean_is_scalar(x_38)) { + x_40 = lean_alloc_ctor(1, 1, 0); +} else { + x_40 = x_38; } +lean_ctor_set(x_40, 0, x_39); +x_41 = l_liftExcept___at_Lean_Elab_liftMacroM___spec__1(x_40, x_3, x_36); +lean_dec(x_40); +return x_41; } } -LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___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_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__1___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_9 = l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__5___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; +uint8_t x_5; lean_object* x_6; lean_object* x_7; +x_5 = l_Lean_Environment_contains(x_1, x_2); +x_6 = lean_box(x_5); +x_7 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_4); +return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__13___rarg(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__1___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) { _start: { -lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___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_object* x_7; lean_object* x_8; +x_7 = l_Lean_ResolveName_resolveNamespace(x_1, x_2, x_3, x_4); +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_7); +lean_ctor_set(x_8, 1, x_6); +return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___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_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__1___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) { _start: { -lean_object* x_7; -x_7 = lean_alloc_closure((void*)(l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__13___rarg), 1, 0); -return x_7; +lean_object* x_7; lean_object* x_8; +x_7 = l_Lean_ResolveName_resolveGlobalName(x_1, x_2, x_3, x_4); +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_7); +lean_ctor_set(x_8, 1, x_6); +return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___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_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___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_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; @@ -34221,1105 +34991,604 @@ 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__2(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); -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_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; lean_object* x_65; -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); -x_56 = lean_ctor_get(x_38, 6); -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_38); -x_57 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_57, 0, x_51); -lean_ctor_set(x_57, 1, x_36); -lean_ctor_set(x_57, 2, x_52); -lean_ctor_set(x_57, 3, x_53); -lean_ctor_set(x_57, 4, x_54); -lean_ctor_set(x_57, 5, x_55); -lean_ctor_set(x_57, 6, x_56); -x_58 = lean_st_ref_set(x_7, x_57, x_39); -x_59 = lean_ctor_get(x_58, 1); -lean_inc(x_59); -lean_dec(x_58); -x_60 = lean_ctor_get(x_35, 1); -lean_inc(x_60); -lean_dec(x_35); -x_61 = l_List_reverse___rarg(x_60); -x_62 = l_List_forM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__2(x_61, x_2, x_3, x_4, x_5, x_6, x_7, 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_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_62); - x_64 = lean_box(0); -} -if (lean_is_scalar(x_64)) { - x_65 = lean_alloc_ctor(0, 2, 0); -} else { - x_65 = x_64; -} -lean_ctor_set(x_65, 0, x_34); -lean_ctor_set(x_65, 1, x_63); -return x_65; -} -} -else -{ -lean_object* x_66; -x_66 = lean_ctor_get(x_33, 0); -lean_inc(x_66); -lean_dec(x_33); -if (lean_obj_tag(x_66) == 0) -{ -lean_object* x_67; lean_object* x_68; lean_object* x_69; uint8_t x_70; -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 = l_Lean_maxRecDepthErrorMessage; -x_70 = lean_string_dec_eq(x_68, x_69); -if (x_70 == 0) -{ -lean_object* x_71; lean_object* x_72; lean_object* x_73; -x_71 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_71, 0, x_68); -x_72 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_72, 0, x_71); -x_73 = l_Lean_throwErrorAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__10(x_67, x_72, x_2, x_3, x_4, x_5, x_6, x_7, x_27); -return x_73; -} -else -{ -lean_object* x_74; -lean_dec(x_68); -x_74 = l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__12(x_67, 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_74; -} -} -else -{ -lean_object* x_75; -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_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__13___rarg(x_27); -return x_75; -} -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__15___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; lean_object* x_9; uint8_t x_10; -x_3 = lean_st_ref_get(x_1, x_2); -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 = lean_ctor_get(x_4, 3); -lean_inc(x_6); -lean_dec(x_4); -x_7 = lean_st_ref_take(x_1, x_5); -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 = !lean_is_exclusive(x_8); -if (x_10 == 0) -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_11 = lean_ctor_get(x_8, 3); -lean_dec(x_11); -x_12 = l_Lean_Elab_Term_Context_autoBoundImplicits___default___closed__3; -lean_ctor_set(x_8, 3, x_12); -x_13 = lean_st_ref_set(x_1, x_8, x_9); -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) -{ -lean_object* x_15; -x_15 = lean_ctor_get(x_13, 0); -lean_dec(x_15); -lean_ctor_set(x_13, 0, x_6); -return x_13; -} -else -{ -lean_object* x_16; lean_object* x_17; -x_16 = lean_ctor_get(x_13, 1); -lean_inc(x_16); -lean_dec(x_13); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_6); -lean_ctor_set(x_17, 1, x_16); -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; -x_18 = lean_ctor_get(x_8, 0); -x_19 = lean_ctor_get(x_8, 1); -x_20 = lean_ctor_get(x_8, 2); -x_21 = lean_ctor_get(x_8, 4); -x_22 = lean_ctor_get(x_8, 5); -x_23 = lean_ctor_get(x_8, 6); -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_dec(x_8); -x_24 = l_Lean_Elab_Term_Context_autoBoundImplicits___default___closed__3; -x_25 = lean_alloc_ctor(0, 7, 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_24); -lean_ctor_set(x_25, 4, x_21); -lean_ctor_set(x_25, 5, x_22); -lean_ctor_set(x_25, 6, x_23); -x_26 = lean_st_ref_set(x_1, x_25, x_9); -x_27 = lean_ctor_get(x_26, 1); -lean_inc(x_27); -if (lean_is_exclusive(x_26)) { - lean_ctor_release(x_26, 0); - lean_ctor_release(x_26, 1); - x_28 = x_26; -} else { - lean_dec_ref(x_26); - x_28 = lean_box(0); -} -if (lean_is_scalar(x_28)) { - x_29 = lean_alloc_ctor(0, 2, 0); -} else { - x_29 = x_28; -} -lean_ctor_set(x_29, 0, x_6); -lean_ctor_set(x_29, 1, x_27); -return x_29; -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__15(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 = lean_alloc_closure((void*)(l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__15___rarg___boxed), 2, 0); -return x_6; -} -} -LEAN_EXPORT lean_object* l_Lean_getBoolOption___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__17(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: -{ -lean_object* x_10; uint8_t x_11; lean_object* x_12; lean_object* x_13; -x_10 = lean_ctor_get(x_7, 2); -x_11 = l_Lean_KVMap_getBool(x_10, x_1, x_2); -x_12 = lean_box(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_9); -return x_13; -} -} -static double _init_l_Lean_withSeconds___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__18___closed__1() { -_start: -{ -lean_object* x_1; uint8_t x_2; lean_object* x_3; double x_4; -x_1 = lean_unsigned_to_nat(1000000000u); -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_Lean_withSeconds___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__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) { -_start: -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_9 = lean_io_mono_nanos_now(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_apply_7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_11); -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_io_mono_nanos_now(x_14); -x_16 = !lean_is_exclusive(x_15); -if (x_16 == 0) -{ -lean_object* x_17; lean_object* x_18; uint8_t x_19; lean_object* x_20; double x_21; double x_22; double x_23; lean_object* x_24; lean_object* x_25; -x_17 = lean_ctor_get(x_15, 0); -x_18 = lean_nat_sub(x_17, x_10); -lean_dec(x_10); -lean_dec(x_17); -x_19 = 0; -x_20 = lean_unsigned_to_nat(0u); -x_21 = l_Float_ofScientific(x_18, x_19, x_20); -lean_dec(x_18); -x_22 = l_Lean_withSeconds___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__18___closed__1; -x_23 = lean_float_div(x_21, x_22); -x_24 = lean_box_float(x_23); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_13); -lean_ctor_set(x_25, 1, x_24); -lean_ctor_set(x_15, 0, x_25); -return x_15; -} -else -{ -lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; lean_object* x_30; double x_31; double x_32; double x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_26 = lean_ctor_get(x_15, 0); -x_27 = lean_ctor_get(x_15, 1); -lean_inc(x_27); -lean_inc(x_26); -lean_dec(x_15); -x_28 = lean_nat_sub(x_26, x_10); -lean_dec(x_10); -lean_dec(x_26); -x_29 = 0; -x_30 = lean_unsigned_to_nat(0u); -x_31 = l_Float_ofScientific(x_28, x_29, x_30); -lean_dec(x_28); -x_32 = l_Lean_withSeconds___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__18___closed__1; -x_33 = lean_float_div(x_31, x_32); -x_34 = lean_box_float(x_33); -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_13); -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_27); -return x_36; -} -} -else -{ -uint8_t x_37; -lean_dec(x_10); -x_37 = !lean_is_exclusive(x_12); -if (x_37 == 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_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; -} -} -} -} -static lean_object* _init_l_Lean_withOptProfile___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__16___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("profiler", 8); -return x_1; -} -} -static lean_object* _init_l_Lean_withOptProfile___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__16___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_withOptProfile___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__16___closed__1; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Lean_withOptProfile___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___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) { -_start: -{ -lean_object* x_9; uint8_t x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_9 = l_Lean_withOptProfile___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__16___closed__2; -x_10 = 0; -x_11 = l_Lean_getBoolOption___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__17(x_9, x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_unbox(x_12); -lean_dec(x_12); -if (x_13 == 0) -{ -lean_object* x_14; lean_object* x_15; -x_14 = lean_ctor_get(x_11, 1); -lean_inc(x_14); -lean_dec(x_11); -x_15 = lean_apply_7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_14); -if (lean_obj_tag(x_15) == 0) -{ -uint8_t x_16; -x_16 = !lean_is_exclusive(x_15); -if (x_16 == 0) -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_17 = lean_ctor_get(x_15, 0); -x_18 = lean_box(0); -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_17); -lean_ctor_set(x_19, 1, x_18); -lean_ctor_set(x_15, 0, x_19); -return x_15; -} -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_box(0); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_20); -lean_ctor_set(x_23, 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_21); -return x_24; -} -} -else -{ -uint8_t x_25; -x_25 = !lean_is_exclusive(x_15); -if (x_25 == 0) -{ -return x_15; -} -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_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; -} -} -} -else -{ -lean_object* x_29; lean_object* x_30; -x_29 = lean_ctor_get(x_11, 1); -lean_inc(x_29); -lean_dec(x_11); -x_30 = l_Lean_withSeconds___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__18(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_29); -if (lean_obj_tag(x_30) == 0) -{ -uint8_t x_31; -x_31 = !lean_is_exclusive(x_30); -if (x_31 == 0) -{ -lean_object* x_32; uint8_t x_33; -x_32 = lean_ctor_get(x_30, 0); -x_33 = !lean_is_exclusive(x_32); -if (x_33 == 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__2(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); +x_47 = !lean_is_exclusive(x_46); +if (x_47 == 0) { -lean_object* x_34; lean_object* x_35; -x_34 = lean_ctor_get(x_32, 1); -x_35 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_35, 0, x_34); -lean_ctor_set(x_32, 1, x_35); -return x_30; +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_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 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_38, 0, x_37); -x_39 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_39, 0, x_36); -lean_ctor_set(x_39, 1, x_38); -lean_ctor_set(x_30, 0, x_39); -return x_30; +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_40; lean_object* x_41; 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_40 = lean_ctor_get(x_30, 0); -x_41 = lean_ctor_get(x_30, 1); -lean_inc(x_41); -lean_inc(x_40); -lean_dec(x_30); -x_42 = lean_ctor_get(x_40, 0); -lean_inc(x_42); -x_43 = lean_ctor_get(x_40, 1); -lean_inc(x_43); -if (lean_is_exclusive(x_40)) { - lean_ctor_release(x_40, 0); - lean_ctor_release(x_40, 1); - x_44 = x_40; +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_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); +x_56 = lean_ctor_get(x_38, 6); +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_38); +x_57 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_57, 0, x_51); +lean_ctor_set(x_57, 1, x_36); +lean_ctor_set(x_57, 2, x_52); +lean_ctor_set(x_57, 3, x_53); +lean_ctor_set(x_57, 4, x_54); +lean_ctor_set(x_57, 5, x_55); +lean_ctor_set(x_57, 6, x_56); +x_58 = lean_st_ref_set(x_7, x_57, x_39); +x_59 = lean_ctor_get(x_58, 1); +lean_inc(x_59); +lean_dec(x_58); +x_60 = lean_ctor_get(x_35, 1); +lean_inc(x_60); +lean_dec(x_35); +x_61 = l_List_reverse___rarg(x_60); +x_62 = l_List_forM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__2(x_61, x_2, x_3, x_4, x_5, x_6, x_7, 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_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_40); - x_44 = lean_box(0); + lean_dec_ref(x_62); + x_64 = lean_box(0); } -x_45 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_45, 0, x_43); -if (lean_is_scalar(x_44)) { - x_46 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_64)) { + x_65 = lean_alloc_ctor(0, 2, 0); } else { - x_46 = x_44; -} -lean_ctor_set(x_46, 0, x_42); -lean_ctor_set(x_46, 1, 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_41); -return x_47; + x_65 = x_64; } +lean_ctor_set(x_65, 0, x_34); +lean_ctor_set(x_65, 1, x_63); +return x_65; } -else -{ -uint8_t x_48; -x_48 = !lean_is_exclusive(x_30); -if (x_48 == 0) -{ -return x_30; } else { -lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_49 = lean_ctor_get(x_30, 0); -x_50 = lean_ctor_get(x_30, 1); -lean_inc(x_50); -lean_inc(x_49); -lean_dec(x_30); -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_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__20(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; uint8_t x_16; -x_13 = lean_st_ref_take(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_is_exclusive(x_14); -if (x_16 == 0) +lean_object* x_66; +x_66 = lean_ctor_get(x_33, 0); +lean_inc(x_66); +lean_dec(x_33); +if (lean_obj_tag(x_66) == 0) { -lean_object* x_17; lean_object* x_18; lean_object* x_19; size_t x_20; size_t 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_17 = lean_ctor_get(x_14, 3); -x_18 = l_Lean_PersistentArray_toArray___rarg(x_17); -x_19 = lean_array_get_size(x_18); -x_20 = lean_usize_of_nat(x_19); -lean_dec(x_19); -x_21 = 0; -x_22 = l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___spec__1(x_20, x_21, x_18); -x_23 = lean_alloc_ctor(9, 3, 1); -lean_ctor_set(x_23, 0, x_2); -lean_ctor_set(x_23, 1, x_4); -lean_ctor_set(x_23, 2, x_22); -lean_ctor_set_uint8(x_23, sizeof(void*)*3, x_5); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_3); -lean_ctor_set(x_24, 1, x_23); -x_25 = l_Lean_PersistentArray_push___rarg(x_1, x_24); -lean_ctor_set(x_14, 3, x_25); -x_26 = lean_st_ref_set(x_11, x_14, x_15); -x_27 = !lean_is_exclusive(x_26); -if (x_27 == 0) +lean_object* x_67; lean_object* x_68; lean_object* x_69; uint8_t x_70; +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 = l_Lean_maxRecDepthErrorMessage; +x_70 = lean_string_dec_eq(x_68, x_69); +if (x_70 == 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; +lean_object* x_71; lean_object* x_72; lean_object* x_73; +x_71 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_71, 0, x_68); +x_72 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_72, 0, x_71); +x_73 = l_Lean_throwErrorAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__3(x_67, x_72, x_2, x_3, x_4, x_5, x_6, x_7, x_27); +return x_73; } 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_74; +lean_dec(x_68); +x_74 = l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__5(x_67, 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_74; } } 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; size_t 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; lean_object* x_53; -x_33 = lean_ctor_get(x_14, 0); -x_34 = lean_ctor_get(x_14, 1); -x_35 = lean_ctor_get(x_14, 2); -x_36 = lean_ctor_get(x_14, 3); -x_37 = lean_ctor_get(x_14, 4); -x_38 = lean_ctor_get(x_14, 5); -x_39 = lean_ctor_get(x_14, 6); -lean_inc(x_39); -lean_inc(x_38); -lean_inc(x_37); -lean_inc(x_36); -lean_inc(x_35); -lean_inc(x_34); -lean_inc(x_33); -lean_dec(x_14); -x_40 = l_Lean_PersistentArray_toArray___rarg(x_36); -x_41 = lean_array_get_size(x_40); -x_42 = lean_usize_of_nat(x_41); -lean_dec(x_41); -x_43 = 0; -x_44 = l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___spec__1(x_42, x_43, x_40); -x_45 = lean_alloc_ctor(9, 3, 1); -lean_ctor_set(x_45, 0, x_2); -lean_ctor_set(x_45, 1, x_4); -lean_ctor_set(x_45, 2, x_44); -lean_ctor_set_uint8(x_45, sizeof(void*)*3, x_5); -x_46 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_46, 0, x_3); -lean_ctor_set(x_46, 1, x_45); -x_47 = l_Lean_PersistentArray_push___rarg(x_1, x_46); -x_48 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_48, 0, x_33); -lean_ctor_set(x_48, 1, x_34); -lean_ctor_set(x_48, 2, x_35); -lean_ctor_set(x_48, 3, x_47); -lean_ctor_set(x_48, 4, x_37); -lean_ctor_set(x_48, 5, x_38); -lean_ctor_set(x_48, 6, x_39); -x_49 = lean_st_ref_set(x_11, x_48, x_15); -x_50 = lean_ctor_get(x_49, 1); -lean_inc(x_50); -if (lean_is_exclusive(x_49)) { - lean_ctor_release(x_49, 0); - lean_ctor_release(x_49, 1); - x_51 = x_49; -} else { - lean_dec_ref(x_49); - x_51 = lean_box(0); -} -x_52 = lean_box(0); -if (lean_is_scalar(x_51)) { - x_53 = lean_alloc_ctor(0, 2, 0); -} else { - x_53 = x_51; +lean_object* x_75; +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_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__6___rarg(x_27); +return x_75; } -lean_ctor_set(x_53, 0, x_52); -lean_ctor_set(x_53, 1, x_50); -return x_53; } } } -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__19(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) { +LEAN_EXPORT lean_object* l_Lean_Elab_throwPostpone___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__7___rarg(lean_object* x_1) { _start: { -uint8_t x_13; -x_13 = !lean_is_exclusive(x_10); -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_10, 5); -x_15 = l_Lean_replaceRef(x_3, x_14); -lean_dec(x_14); -lean_ctor_set(x_10, 5, x_15); -x_16 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_4, x_8, x_9, x_10, x_11, x_12); -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___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__20(x_1, x_2, x_3, x_17, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_18); -lean_dec(x_10); -return x_19; +lean_object* x_2; lean_object* x_3; +x_2 = l_Lean_Elab_throwPostpone___at_Lean_Elab_Term_tryPostpone___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; } -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; -x_20 = lean_ctor_get(x_10, 0); -x_21 = lean_ctor_get(x_10, 1); -x_22 = lean_ctor_get(x_10, 2); -x_23 = lean_ctor_get(x_10, 3); -x_24 = lean_ctor_get(x_10, 4); -x_25 = lean_ctor_get(x_10, 5); -x_26 = lean_ctor_get(x_10, 6); -x_27 = lean_ctor_get(x_10, 7); -x_28 = lean_ctor_get(x_10, 8); -x_29 = lean_ctor_get(x_10, 9); -x_30 = lean_ctor_get(x_10, 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_10); -x_31 = l_Lean_replaceRef(x_3, x_25); -lean_dec(x_25); -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_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_4, x_8, x_9, x_32, x_11, x_12); -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 = l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__20(x_1, x_2, x_3, x_34, x_5, x_6, x_7, x_8, x_9, x_32, x_11, x_35); -lean_dec(x_32); -return x_36; } +LEAN_EXPORT lean_object* l_Lean_Elab_throwPostpone___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___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_7; +x_7 = lean_alloc_closure((void*)(l_Lean_Elab_throwPostpone___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__7___rarg), 1, 0); +return x_7; } } -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__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_EXPORT lean_object* l_liftExcept___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) { -lean_object* x_9; lean_object* x_10; -x_9 = lean_ctor_get(x_1, 0); -lean_inc(x_9); -x_10 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_10, 0, x_9); -lean_ctor_set(x_10, 1, x_8); -return x_10; +lean_object* x_4; lean_object* x_5; +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_4); +x_5 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_5, 0, x_4); +lean_ctor_set(x_5, 1, x_3); +return x_5; } else { -lean_object* x_11; lean_object* x_12; -x_11 = lean_ctor_get(x_1, 0); -lean_inc(x_11); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_11); -lean_ctor_set(x_12, 1, x_8); -return x_12; +lean_object* x_6; lean_object* x_7; +x_6 = lean_ctor_get(x_1, 0); +lean_inc(x_6); +x_7 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_3); +return x_7; } } } -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__14___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_throwError___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___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) { _start: { -lean_object* x_9; -x_9 = lean_apply_7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_9) == 0) -{ -uint8_t x_10; -x_10 = !lean_is_exclusive(x_9); -if (x_10 == 0) -{ -lean_object* x_11; lean_object* x_12; -x_11 = lean_ctor_get(x_9, 0); -x_12 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_12, 0, x_11); -lean_ctor_set(x_9, 0, x_12); -return x_9; -} -else -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_13 = lean_ctor_get(x_9, 0); -x_14 = lean_ctor_get(x_9, 1); -lean_inc(x_14); +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, 2); +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); -lean_dec(x_9); -x_15 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_15, 0, 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 -{ -uint8_t x_17; -x_17 = !lean_is_exclusive(x_9); -if (x_17 == 0) +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); +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 0) { -lean_object* x_18; lean_object* x_19; -x_18 = lean_ctor_get(x_9, 0); -x_19 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_19, 0, x_18); -lean_ctor_set_tag(x_9, 0); -lean_ctor_set(x_9, 0, x_19); -return x_9; +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_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_20 = lean_ctor_get(x_9, 0); -x_21 = lean_ctor_get(x_9, 1); -lean_inc(x_21); +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_dec(x_9); -x_22 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_22, 0, 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; -} +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_withTraceNode___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__14___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) { +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___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) { _start: { -lean_object* x_15; lean_object* x_16; lean_object* x_17; -lean_inc(x_12); -x_15 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__19(x_1, x_2, x_3, x_6, x_4, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -x_16 = lean_ctor_get(x_15, 1); +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_Term_0__Lean_Elab_Term_elabTermAux___spec__11(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_dec(x_15); -x_17 = l_MonadExcept_ofExcept___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__21(x_5, x_8, x_9, x_10, x_11, x_12, x_13, x_16); -lean_dec(x_12); +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_Term_0__Lean_Elab_Term_elabTermAux___spec__11(x_2, x_3, x_4, x_5, x_6, x_26, x_8, x_9); lean_dec(x_8); -return x_17; -} +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_withTraceNode___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__14___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("[", 1); -return x_1; } } -static lean_object* _init_l_Lean_withTraceNode___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__14___closed__2() { +LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___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) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_withTraceNode___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__14___closed__1; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_9 = l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__5___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; } } -static lean_object* _init_l_Lean_withTraceNode___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__14___closed__3() { +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__13___rarg(lean_object* x_1) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("s] ", 3); -return x_1; +lean_object* x_2; lean_object* x_3; +x_2 = l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___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; } } -static lean_object* _init_l_Lean_withTraceNode___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__14___closed__4() { +LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___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) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_withTraceNode___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__14___closed__3; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +lean_object* x_7; +x_7 = lean_alloc_closure((void*)(l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__13___rarg), 1, 0); +return x_7; } } -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__14(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_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___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) { _start: { -lean_object* x_12; lean_object* x_13; uint8_t x_14; -lean_inc(x_1); -x_12 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Term_traceAtCmdPos___spec__1(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -x_13 = lean_ctor_get(x_12, 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); +lean_dec(x_10); +x_13 = lean_ctor_get(x_6, 3); lean_inc(x_13); -x_14 = lean_unbox(x_13); -lean_dec(x_13); -if (x_14 == 0) -{ -lean_object* x_15; lean_object* x_16; -lean_dec(x_2); -lean_dec(x_1); -x_15 = lean_ctor_get(x_12, 1); +x_14 = lean_ctor_get(x_6, 4); +lean_inc(x_14); +x_15 = lean_ctor_get(x_6, 5); lean_inc(x_15); -lean_dec(x_12); -x_16 = lean_apply_7(x_3, x_5, x_6, x_7, x_8, x_9, x_10, 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; lean_object* x_22; lean_object* x_23; -x_17 = lean_ctor_get(x_12, 1); +x_16 = lean_ctor_get(x_6, 6); +lean_inc(x_16); +x_17 = lean_ctor_get(x_6, 7); lean_inc(x_17); -lean_dec(x_12); -x_18 = lean_ctor_get(x_9, 5); +x_18 = lean_ctor_get(x_6, 10); lean_inc(x_18); -x_19 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__15___rarg(x_10, x_17); -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_alloc_closure((void*)(l_Lean_withTraceNode___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__14___lambda__1), 8, 1); -lean_closure_set(x_22, 0, x_3); -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_23 = l_Lean_withOptProfile___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__16(x_22, x_5, x_6, x_7, x_8, x_9, x_10, x_21); -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; -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_24, 0); +lean_inc(x_12); +x_19 = lean_alloc_closure((void*)(l_Lean_Elab_liftMacroM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___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___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___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___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___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___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___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_24, 1); +x_27 = lean_ctor_get(x_25, 1); lean_inc(x_27); -lean_dec(x_24); -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_26); -x_28 = lean_apply_8(x_2, x_26, x_5, x_6, x_7, x_8, x_9, x_10, x_25); -if (lean_obj_tag(x_28) == 0) -{ -if (lean_obj_tag(x_27) == 0) -{ -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; -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_box(0); -x_32 = l_Lean_withTraceNode___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__14___lambda__2(x_20, x_1, x_18, x_4, x_26, x_29, x_31, x_5, x_6, x_7, x_8, x_9, x_10, x_30); -lean_dec(x_10); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); +lean_dec(x_25); +x_28 = lean_ctor_get(x_26, 1); +lean_inc(x_28); lean_dec(x_26); -return x_32; -} -else +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_33; lean_object* x_34; lean_object* x_35; double x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; 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_28, 0); -lean_inc(x_33); -x_34 = lean_ctor_get(x_28, 1); +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); -lean_dec(x_28); -x_35 = lean_ctor_get(x_27, 0); +x_35 = lean_ctor_get(x_33, 1); lean_inc(x_35); -lean_dec(x_27); -x_36 = lean_unbox_float(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_37 = lean_float_to_string(x_36); -x_38 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_38, 0, x_37); -x_39 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_39, 0, x_38); -x_40 = l_Lean_withTraceNode___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__14___closed__2; -x_41 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_41, 0, x_40); -lean_ctor_set(x_41, 1, x_39); -x_42 = l_Lean_withTraceNode___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__14___closed__4; -x_43 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_43, 0, x_41); -lean_ctor_set(x_43, 1, x_42); -x_44 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_44, 0, x_43); -lean_ctor_set(x_44, 1, x_33); -x_45 = l_Lean_Elab_Term_MVarErrorInfo_logError___closed__4; -x_46 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -x_47 = lean_box(0); -x_48 = l_Lean_withTraceNode___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__14___lambda__2(x_20, x_1, x_18, x_4, x_26, x_46, x_47, x_5, x_6, x_7, x_8, x_9, x_10, x_34); -lean_dec(x_10); -lean_dec(x_8); +x_45 = l_List_reverse___rarg(x_44); +x_46 = l_List_forM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__2(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_26); -return x_48; +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_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_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 { -uint8_t x_49; -lean_dec(x_27); -lean_dec(x_26); -lean_dec(x_20); -lean_dec(x_18); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +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_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); +x_56 = lean_ctor_get(x_38, 6); +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_38); +x_57 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_57, 0, x_51); +lean_ctor_set(x_57, 1, x_36); +lean_ctor_set(x_57, 2, x_52); +lean_ctor_set(x_57, 3, x_53); +lean_ctor_set(x_57, 4, x_54); +lean_ctor_set(x_57, 5, x_55); +lean_ctor_set(x_57, 6, x_56); +x_58 = lean_st_ref_set(x_7, x_57, x_39); +x_59 = lean_ctor_get(x_58, 1); +lean_inc(x_59); +lean_dec(x_58); +x_60 = lean_ctor_get(x_35, 1); +lean_inc(x_60); +lean_dec(x_35); +x_61 = l_List_reverse___rarg(x_60); +x_62 = l_List_forM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__2(x_61, x_2, x_3, x_4, x_5, x_6, x_7, x_59); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_1); -x_49 = !lean_is_exclusive(x_28); -if (x_49 == 0) -{ -return x_28; +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +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_62); + x_64 = lean_box(0); } -else -{ -lean_object* x_50; lean_object* x_51; lean_object* x_52; -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, 2, 0); -lean_ctor_set(x_52, 0, x_50); -lean_ctor_set(x_52, 1, x_51); -return x_52; +if (lean_is_scalar(x_64)) { + x_65 = lean_alloc_ctor(0, 2, 0); +} else { + x_65 = x_64; } +lean_ctor_set(x_65, 0, x_34); +lean_ctor_set(x_65, 1, x_63); +return x_65; } } else { -uint8_t x_53; -lean_dec(x_20); -lean_dec(x_18); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +lean_object* x_66; +x_66 = lean_ctor_get(x_33, 0); +lean_inc(x_66); +lean_dec(x_33); +if (lean_obj_tag(x_66) == 0) +{ +lean_object* x_67; lean_object* x_68; lean_object* x_69; uint8_t x_70; +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 = l_Lean_maxRecDepthErrorMessage; +x_70 = lean_string_dec_eq(x_68, x_69); +if (x_70 == 0) +{ +lean_object* x_71; lean_object* x_72; lean_object* x_73; +x_71 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_71, 0, x_68); +x_72 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_72, 0, x_71); +x_73 = l_Lean_throwErrorAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__10(x_67, x_72, x_2, x_3, x_4, x_5, x_6, x_7, x_27); +return x_73; +} +else +{ +lean_object* x_74; +lean_dec(x_68); +x_74 = l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__12(x_67, 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_53 = !lean_is_exclusive(x_23); -if (x_53 == 0) -{ -return x_23; +return x_74; +} } else { -lean_object* x_54; lean_object* x_55; lean_object* x_56; -x_54 = lean_ctor_get(x_23, 0); -x_55 = lean_ctor_get(x_23, 1); -lean_inc(x_55); -lean_inc(x_54); -lean_dec(x_23); -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_object* x_75; +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_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__13___rarg(x_27); +return x_75; } } } } -LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___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_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___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) { _start: { lean_object* x_9; lean_object* x_10; lean_object* x_11; @@ -36253,7 +36522,7 @@ x_36 = lean_nat_add(x_14, x_35); lean_dec(x_14); lean_ctor_set(x_8, 3, x_36); x_37 = 1; -x_38 = l_Lean_withTraceNode___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__14(x_1, x_2, x_3, x_37, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_38 = l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1(x_1, x_2, x_3, x_37, x_4, x_5, x_6, x_7, x_8, x_9, x_10); return x_38; } else @@ -36276,7 +36545,7 @@ lean_ctor_set(x_41, 8, x_19); lean_ctor_set(x_41, 9, x_20); lean_ctor_set(x_41, 10, x_21); x_42 = 1; -x_43 = l_Lean_withTraceNode___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__14(x_1, x_2, x_3, x_42, x_4, x_5, x_6, x_7, x_41, x_9, x_10); +x_43 = l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1(x_1, x_2, x_3, x_42, x_4, x_5, x_6, x_7, x_41, x_9, x_10); return x_43; } } @@ -36296,7 +36565,7 @@ lean_dec(x_11); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_44 = l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__22(x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_44 = l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__14(x_16, 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); @@ -36511,122 +36780,11 @@ lean_dec(x_1); return x_7; } } -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__15___rarg___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__15___rarg(x_1, x_2); -lean_dec(x_1); -return x_3; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__15___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___private_Lean_Util_Trace_0__Lean_getResetTraces___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__15(x_1, x_2, x_3, x_4, x_5); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_6; -} -} -LEAN_EXPORT lean_object* l_Lean_getBoolOption___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___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) { -_start: -{ -uint8_t x_10; lean_object* x_11; -x_10 = lean_unbox(x_2); -lean_dec(x_2); -x_11 = l_Lean_getBoolOption___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__17(x_1, x_10, 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_1); -return x_11; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___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) { -_start: -{ -uint8_t x_13; lean_object* x_14; -x_13 = lean_unbox(x_5); -lean_dec(x_5); -x_14 = l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__20(x_1, x_2, x_3, x_4, x_13, 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); -lean_dec(x_6); -return x_14; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___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: -{ -uint8_t x_13; lean_object* x_14; -x_13 = lean_unbox(x_5); -lean_dec(x_5); -x_14 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__19(x_1, x_2, x_3, x_4, x_13, 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); -return x_14; -} -} -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__21___boxed(lean_object* x_1, lean_object* x_2, 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_MonadExcept_ofExcept___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__21(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); -lean_dec(x_1); -return x_9; -} -} -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__14___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: -{ -uint8_t x_15; lean_object* x_16; -x_15 = lean_unbox(x_4); -lean_dec(x_4); -x_16 = l_Lean_withTraceNode___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__14___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); -lean_dec(x_13); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_7); -lean_dec(x_5); -return x_16; -} -} -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___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) { -_start: -{ -uint8_t x_12; lean_object* x_13; -x_12 = lean_unbox(x_4); -lean_dec(x_4); -x_13 = l_Lean_withTraceNode___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__14(x_1, x_2, x_3, x_12, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_13; -} -} -LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__22___boxed(lean_object* x_1, lean_object* x_2, lean_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_throwMaxRecDepthAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___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) { _start: { lean_object* x_9; -x_9 = l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__22(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_9 = l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__14(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); @@ -37707,7 +37865,7 @@ x_18 = l_Lean_Elab_Term_ensureType___lambda__1___closed__2; x_19 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_19, 0, x_18); lean_ctor_set(x_19, 1, x_17); -x_20 = l_Lean_Elab_Term_mkCoe___closed__5; +x_20 = l_Lean_Elab_Term_mkCoe___lambda__1___closed__4; x_21 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_21, 0, x_19); lean_ctor_set(x_21, 1, x_20); @@ -38473,7 +38631,7 @@ lean_object* x_9; lean_object* x_10; uint8_t x_11; x_9 = lean_ctor_get(x_6, 2); lean_inc(x_9); x_10 = l_Lean_Elab_Term_withAutoBoundImplicit___rarg___closed__1; -x_11 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_9, x_10); +x_11 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_9, x_10); lean_dec(x_9); if (x_11 == 0) { @@ -38886,162 +39044,171 @@ x_2 = lean_array_to_list(lean_box(0), x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_collectUnassignedMVars_go(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_collectUnassignedMVars_go(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, 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_2) == 0) { -lean_object* x_11; +lean_object* x_12; +lean_dec(x_4); lean_dec(x_1); -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; +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_3); +lean_ctor_set(x_12, 1, x_11); +return x_12; } else { -uint8_t x_12; -x_12 = !lean_is_exclusive(x_2); -if (x_12 == 0) +uint8_t x_13; +x_13 = !lean_is_exclusive(x_2); +if (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_2, 0); -x_14 = lean_ctor_get(x_2, 1); -lean_inc(x_13); -x_15 = l_Lean_MVarId_isAssigned___at_Lean_Elab_Term_synthesizeInstMVarCore___spec__2(x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -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; uint8_t x_20; +x_14 = lean_ctor_get(x_2, 0); +x_15 = lean_ctor_get(x_2, 1); +lean_inc(x_14); +x_16 = l_Lean_Expr_mvar___override(x_14); lean_inc(x_16); -x_17 = lean_unbox(x_16); -lean_dec(x_16); -if (x_17 == 0) -{ -lean_object* x_18; lean_object* x_19; uint8_t x_20; -x_18 = lean_ctor_get(x_15, 1); -lean_inc(x_18); -lean_dec(x_15); -lean_inc(x_13); -x_19 = l_Lean_Expr_mvar___override(x_13); -x_20 = l_Array_contains___at_Lean_Elab_Term_collectUnassignedMVars_go___spec__1(x_3, x_19); +x_17 = lean_array_push(x_4, x_16); +lean_inc(x_14); +x_18 = l_Lean_MVarId_isAssigned___at_Lean_Elab_Term_synthesizeInstMVarCore___spec__2(x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_unbox(x_19); +lean_dec(x_19); if (x_20 == 0) { lean_object* x_21; uint8_t x_22; -lean_inc(x_1); -lean_inc(x_13); -x_21 = lean_apply_1(x_1, x_13); -x_22 = lean_unbox(x_21); -lean_dec(x_21); +x_21 = lean_ctor_get(x_18, 1); +lean_inc(x_21); +lean_dec(x_18); +x_22 = l_Array_contains___at_Lean_Elab_Term_collectUnassignedMVars_go___spec__1(x_3, x_16); 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; uint8_t x_32; -lean_inc(x_13); -x_23 = l_Lean_Elab_Term_getMVarDecl(x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_18); -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_object* x_23; uint8_t x_24; +lean_inc(x_1); +lean_inc(x_14); +x_23 = lean_apply_1(x_1, x_14); +x_24 = lean_unbox(x_23); lean_dec(x_23); -x_26 = lean_ctor_get(x_24, 2); +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; uint8_t x_34; +lean_inc(x_14); +x_25 = l_Lean_Elab_Term_getMVarDecl(x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_21); +x_26 = lean_ctor_get(x_25, 0); lean_inc(x_26); -lean_dec(x_24); -x_27 = l_Lean_Meta_getMVars(x_26, x_6, x_7, x_8, x_9, x_25); -x_28 = lean_ctor_get(x_27, 0); +x_27 = lean_ctor_get(x_25, 1); +lean_inc(x_27); +lean_dec(x_25); +x_28 = lean_ctor_get(x_26, 2); lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); -lean_inc(x_29); -lean_dec(x_27); -x_30 = lean_array_get_size(x_28); -x_31 = lean_unsigned_to_nat(0u); -x_32 = lean_nat_dec_lt(x_31, x_30); -if (x_32 == 0) +lean_dec(x_26); +x_29 = l_Lean_Meta_getMVars(x_28, x_7, x_8, x_9, x_10, 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_array_get_size(x_30); +x_33 = lean_unsigned_to_nat(0u); +x_34 = lean_nat_dec_lt(x_33, x_32); +if (x_34 == 0) { -uint8_t x_33; +uint8_t x_35; +lean_dec(x_32); lean_dec(x_30); -lean_dec(x_28); -x_33 = l_Lean_Elab_Term_collectUnassignedMVars_go___closed__1; -if (x_33 == 0) +x_35 = l_Lean_Elab_Term_collectUnassignedMVars_go___closed__1; +if (x_35 == 0) { -lean_object* x_34; lean_object* x_35; -lean_dec(x_19); -x_34 = l_Lean_Elab_Term_collectUnassignedMVars_go___closed__2; -x_35 = l_List_appendTR___rarg(x_34, x_2); -x_2 = x_35; -x_10 = x_29; +lean_object* x_36; lean_object* x_37; +lean_dec(x_16); +x_36 = l_Lean_Elab_Term_collectUnassignedMVars_go___closed__2; +x_37 = l_List_appendTR___rarg(x_36, x_2); +x_2 = x_37; +x_4 = x_17; +x_11 = x_31; goto _start; } else { -lean_object* x_37; +lean_object* x_39; lean_free_object(x_2); -lean_dec(x_13); -x_37 = lean_array_push(x_3, x_19); -x_2 = x_14; -x_3 = x_37; -x_10 = x_29; +lean_dec(x_14); +x_39 = lean_array_push(x_3, x_16); +x_2 = x_15; +x_3 = x_39; +x_4 = x_17; +x_11 = x_31; goto _start; } } else { -uint8_t x_39; -x_39 = lean_nat_dec_le(x_30, x_30); -if (x_39 == 0) +uint8_t x_41; +x_41 = lean_nat_dec_le(x_32, x_32); +if (x_41 == 0) { -uint8_t x_40; +uint8_t x_42; +lean_dec(x_32); lean_dec(x_30); -lean_dec(x_28); -x_40 = l_Lean_Elab_Term_collectUnassignedMVars_go___closed__1; -if (x_40 == 0) +x_42 = l_Lean_Elab_Term_collectUnassignedMVars_go___closed__1; +if (x_42 == 0) { -lean_object* x_41; lean_object* x_42; -lean_dec(x_19); -x_41 = l_Lean_Elab_Term_collectUnassignedMVars_go___closed__2; -x_42 = l_List_appendTR___rarg(x_41, x_2); -x_2 = x_42; -x_10 = x_29; +lean_object* x_43; lean_object* x_44; +lean_dec(x_16); +x_43 = l_Lean_Elab_Term_collectUnassignedMVars_go___closed__2; +x_44 = l_List_appendTR___rarg(x_43, x_2); +x_2 = x_44; +x_4 = x_17; +x_11 = x_31; goto _start; } else { -lean_object* x_44; +lean_object* x_46; lean_free_object(x_2); -lean_dec(x_13); -x_44 = lean_array_push(x_3, x_19); -x_2 = x_14; -x_3 = x_44; -x_10 = x_29; +lean_dec(x_14); +x_46 = lean_array_push(x_3, x_16); +x_2 = x_15; +x_3 = x_46; +x_4 = x_17; +x_11 = x_31; goto _start; } } else { -size_t x_46; size_t x_47; lean_object* x_48; lean_object* x_49; uint8_t x_50; -x_46 = 0; -x_47 = lean_usize_of_nat(x_30); +size_t x_48; size_t x_49; lean_object* x_50; lean_object* x_51; uint8_t x_52; +x_48 = 0; +x_49 = lean_usize_of_nat(x_32); +lean_dec(x_32); +x_50 = l_Lean_Elab_Term_instInhabitedLetRecToLift___closed__1; +x_51 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_collectUnassignedMVars_go___spec__3(x_17, x_30, x_48, x_49, x_50); lean_dec(x_30); -x_48 = l_Lean_Elab_Term_instInhabitedLetRecToLift___closed__1; -x_49 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_collectUnassignedMVars_go___spec__3(x_3, x_28, x_46, x_47, x_48); -lean_dec(x_28); -x_50 = l_Array_isEmpty___rarg(x_49); -if (x_50 == 0) +x_52 = l_Array_isEmpty___rarg(x_51); +if (x_52 == 0) { -lean_object* x_51; lean_object* x_52; -lean_dec(x_19); -x_51 = lean_array_to_list(lean_box(0), x_49); -x_52 = l_List_appendTR___rarg(x_51, x_2); -x_2 = x_52; -x_10 = x_29; +lean_object* x_53; lean_object* x_54; +lean_dec(x_16); +x_53 = lean_array_to_list(lean_box(0), x_51); +x_54 = l_List_appendTR___rarg(x_53, x_2); +x_2 = x_54; +x_4 = x_17; +x_11 = x_31; goto _start; } else { -lean_object* x_54; -lean_dec(x_49); +lean_object* x_56; +lean_dec(x_51); lean_free_object(x_2); -lean_dec(x_13); -x_54 = lean_array_push(x_3, x_19); -x_2 = x_14; -x_3 = x_54; -x_10 = x_29; +lean_dec(x_14); +x_56 = lean_array_push(x_3, x_16); +x_2 = x_15; +x_3 = x_56; +x_4 = x_17; +x_11 = x_31; goto _start; } } @@ -39049,186 +39216,198 @@ goto _start; } else { -lean_dec(x_19); +lean_dec(x_16); lean_free_object(x_2); -lean_dec(x_13); -x_2 = x_14; -x_10 = x_18; +lean_dec(x_14); +x_2 = x_15; +x_4 = x_17; +x_11 = x_21; goto _start; } } else { -lean_dec(x_19); +lean_dec(x_16); lean_free_object(x_2); -lean_dec(x_13); -x_2 = x_14; -x_10 = x_18; +lean_dec(x_14); +x_2 = x_15; +x_4 = x_17; +x_11 = x_21; goto _start; } } else { -lean_object* x_58; +lean_object* x_60; +lean_dec(x_16); lean_free_object(x_2); -lean_dec(x_13); -x_58 = lean_ctor_get(x_15, 1); -lean_inc(x_58); -lean_dec(x_15); -x_2 = x_14; -x_10 = x_58; +lean_dec(x_14); +x_60 = lean_ctor_get(x_18, 1); +lean_inc(x_60); +lean_dec(x_18); +x_2 = x_15; +x_4 = x_17; +x_11 = x_60; goto _start; } } else { -lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; uint8_t x_64; -x_60 = lean_ctor_get(x_2, 0); -x_61 = lean_ctor_get(x_2, 1); -lean_inc(x_61); -lean_inc(x_60); -lean_dec(x_2); -lean_inc(x_60); -x_62 = l_Lean_MVarId_isAssigned___at_Lean_Elab_Term_synthesizeInstMVarCore___spec__2(x_60, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -x_63 = lean_ctor_get(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; uint8_t x_68; +x_62 = lean_ctor_get(x_2, 0); +x_63 = lean_ctor_get(x_2, 1); lean_inc(x_63); -x_64 = lean_unbox(x_63); -lean_dec(x_63); -if (x_64 == 0) +lean_inc(x_62); +lean_dec(x_2); +lean_inc(x_62); +x_64 = l_Lean_Expr_mvar___override(x_62); +lean_inc(x_64); +x_65 = lean_array_push(x_4, x_64); +lean_inc(x_62); +x_66 = l_Lean_MVarId_isAssigned___at_Lean_Elab_Term_synthesizeInstMVarCore___spec__2(x_62, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_67 = lean_ctor_get(x_66, 0); +lean_inc(x_67); +x_68 = lean_unbox(x_67); +lean_dec(x_67); +if (x_68 == 0) { -lean_object* x_65; lean_object* x_66; uint8_t x_67; -x_65 = lean_ctor_get(x_62, 1); -lean_inc(x_65); -lean_dec(x_62); -lean_inc(x_60); -x_66 = l_Lean_Expr_mvar___override(x_60); -x_67 = l_Array_contains___at_Lean_Elab_Term_collectUnassignedMVars_go___spec__1(x_3, x_66); -if (x_67 == 0) +lean_object* x_69; uint8_t x_70; +x_69 = lean_ctor_get(x_66, 1); +lean_inc(x_69); +lean_dec(x_66); +x_70 = l_Array_contains___at_Lean_Elab_Term_collectUnassignedMVars_go___spec__1(x_3, x_64); +if (x_70 == 0) { -lean_object* x_68; uint8_t x_69; +lean_object* x_71; uint8_t x_72; lean_inc(x_1); -lean_inc(x_60); -x_68 = lean_apply_1(x_1, x_60); -x_69 = lean_unbox(x_68); -lean_dec(x_68); -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; uint8_t x_79; -lean_inc(x_60); -x_70 = l_Lean_Elab_Term_getMVarDecl(x_60, x_4, x_5, x_6, x_7, x_8, x_9, x_65); -x_71 = lean_ctor_get(x_70, 0); -lean_inc(x_71); -x_72 = lean_ctor_get(x_70, 1); -lean_inc(x_72); -lean_dec(x_70); -x_73 = lean_ctor_get(x_71, 2); -lean_inc(x_73); +lean_inc(x_62); +x_71 = lean_apply_1(x_1, x_62); +x_72 = lean_unbox(x_71); lean_dec(x_71); -x_74 = l_Lean_Meta_getMVars(x_73, x_6, x_7, x_8, x_9, x_72); -x_75 = lean_ctor_get(x_74, 0); +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; uint8_t x_82; +lean_inc(x_62); +x_73 = l_Lean_Elab_Term_getMVarDecl(x_62, x_5, x_6, x_7, x_8, x_9, x_10, x_69); +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_74, 2); lean_inc(x_76); lean_dec(x_74); -x_77 = lean_array_get_size(x_75); -x_78 = lean_unsigned_to_nat(0u); -x_79 = lean_nat_dec_lt(x_78, x_77); -if (x_79 == 0) -{ -uint8_t x_80; +x_77 = l_Lean_Meta_getMVars(x_76, x_7, x_8, x_9, x_10, 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); lean_dec(x_77); -lean_dec(x_75); -x_80 = l_Lean_Elab_Term_collectUnassignedMVars_go___closed__1; -if (x_80 == 0) +x_80 = lean_array_get_size(x_78); +x_81 = lean_unsigned_to_nat(0u); +x_82 = lean_nat_dec_lt(x_81, x_80); +if (x_82 == 0) { -lean_object* x_81; lean_object* x_82; lean_object* x_83; -lean_dec(x_66); -x_81 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_81, 0, x_60); -lean_ctor_set(x_81, 1, x_61); -x_82 = l_Lean_Elab_Term_collectUnassignedMVars_go___closed__2; -x_83 = l_List_appendTR___rarg(x_82, x_81); -x_2 = x_83; -x_10 = x_76; +uint8_t x_83; +lean_dec(x_80); +lean_dec(x_78); +x_83 = l_Lean_Elab_Term_collectUnassignedMVars_go___closed__1; +if (x_83 == 0) +{ +lean_object* x_84; lean_object* x_85; lean_object* x_86; +lean_dec(x_64); +x_84 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_84, 0, x_62); +lean_ctor_set(x_84, 1, x_63); +x_85 = l_Lean_Elab_Term_collectUnassignedMVars_go___closed__2; +x_86 = l_List_appendTR___rarg(x_85, x_84); +x_2 = x_86; +x_4 = x_65; +x_11 = x_79; goto _start; } else { -lean_object* x_85; -lean_dec(x_60); -x_85 = lean_array_push(x_3, x_66); -x_2 = x_61; -x_3 = x_85; -x_10 = x_76; +lean_object* x_88; +lean_dec(x_62); +x_88 = lean_array_push(x_3, x_64); +x_2 = x_63; +x_3 = x_88; +x_4 = x_65; +x_11 = x_79; goto _start; } } else { -uint8_t x_87; -x_87 = lean_nat_dec_le(x_77, x_77); -if (x_87 == 0) +uint8_t x_90; +x_90 = lean_nat_dec_le(x_80, x_80); +if (x_90 == 0) { -uint8_t x_88; -lean_dec(x_77); -lean_dec(x_75); -x_88 = l_Lean_Elab_Term_collectUnassignedMVars_go___closed__1; -if (x_88 == 0) +uint8_t x_91; +lean_dec(x_80); +lean_dec(x_78); +x_91 = l_Lean_Elab_Term_collectUnassignedMVars_go___closed__1; +if (x_91 == 0) { -lean_object* x_89; lean_object* x_90; lean_object* x_91; -lean_dec(x_66); -x_89 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_89, 0, x_60); -lean_ctor_set(x_89, 1, x_61); -x_90 = l_Lean_Elab_Term_collectUnassignedMVars_go___closed__2; -x_91 = l_List_appendTR___rarg(x_90, x_89); -x_2 = x_91; -x_10 = x_76; +lean_object* x_92; lean_object* x_93; lean_object* x_94; +lean_dec(x_64); +x_92 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_92, 0, x_62); +lean_ctor_set(x_92, 1, x_63); +x_93 = l_Lean_Elab_Term_collectUnassignedMVars_go___closed__2; +x_94 = l_List_appendTR___rarg(x_93, x_92); +x_2 = x_94; +x_4 = x_65; +x_11 = x_79; goto _start; } else { -lean_object* x_93; -lean_dec(x_60); -x_93 = lean_array_push(x_3, x_66); -x_2 = x_61; -x_3 = x_93; -x_10 = x_76; +lean_object* x_96; +lean_dec(x_62); +x_96 = lean_array_push(x_3, x_64); +x_2 = x_63; +x_3 = x_96; +x_4 = x_65; +x_11 = x_79; goto _start; } } else { -size_t x_95; size_t x_96; lean_object* x_97; lean_object* x_98; uint8_t x_99; -x_95 = 0; -x_96 = lean_usize_of_nat(x_77); -lean_dec(x_77); -x_97 = l_Lean_Elab_Term_instInhabitedLetRecToLift___closed__1; -x_98 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_collectUnassignedMVars_go___spec__3(x_3, x_75, x_95, x_96, x_97); -lean_dec(x_75); -x_99 = l_Array_isEmpty___rarg(x_98); -if (x_99 == 0) +size_t x_98; size_t x_99; lean_object* x_100; lean_object* x_101; uint8_t x_102; +x_98 = 0; +x_99 = lean_usize_of_nat(x_80); +lean_dec(x_80); +x_100 = l_Lean_Elab_Term_instInhabitedLetRecToLift___closed__1; +x_101 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_collectUnassignedMVars_go___spec__3(x_65, x_78, x_98, x_99, x_100); +lean_dec(x_78); +x_102 = l_Array_isEmpty___rarg(x_101); +if (x_102 == 0) { -lean_object* x_100; lean_object* x_101; lean_object* x_102; -lean_dec(x_66); -x_100 = lean_array_to_list(lean_box(0), x_98); -x_101 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_101, 0, x_60); -lean_ctor_set(x_101, 1, x_61); -x_102 = l_List_appendTR___rarg(x_100, x_101); -x_2 = x_102; -x_10 = x_76; +lean_object* x_103; lean_object* x_104; lean_object* x_105; +lean_dec(x_64); +x_103 = lean_array_to_list(lean_box(0), x_101); +x_104 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_104, 0, x_62); +lean_ctor_set(x_104, 1, x_63); +x_105 = l_List_appendTR___rarg(x_103, x_104); +x_2 = x_105; +x_4 = x_65; +x_11 = x_79; goto _start; } else { -lean_object* x_104; -lean_dec(x_98); -lean_dec(x_60); -x_104 = lean_array_push(x_3, x_66); -x_2 = x_61; -x_3 = x_104; -x_10 = x_76; +lean_object* x_107; +lean_dec(x_101); +lean_dec(x_62); +x_107 = lean_array_push(x_3, x_64); +x_2 = x_63; +x_3 = x_107; +x_4 = x_65; +x_11 = x_79; goto _start; } } @@ -39236,31 +39415,35 @@ goto _start; } else { -lean_dec(x_66); -lean_dec(x_60); -x_2 = x_61; -x_10 = x_65; +lean_dec(x_64); +lean_dec(x_62); +x_2 = x_63; +x_4 = x_65; +x_11 = x_69; goto _start; } } else { -lean_dec(x_66); -lean_dec(x_60); -x_2 = x_61; -x_10 = x_65; +lean_dec(x_64); +lean_dec(x_62); +x_2 = x_63; +x_4 = x_65; +x_11 = x_69; goto _start; } } else { -lean_object* x_108; -lean_dec(x_60); -x_108 = lean_ctor_get(x_62, 1); -lean_inc(x_108); +lean_object* x_111; +lean_dec(x_64); lean_dec(x_62); -x_2 = x_61; -x_10 = x_108; +x_111 = lean_ctor_get(x_66, 1); +lean_inc(x_111); +lean_dec(x_66); +x_2 = x_63; +x_4 = x_65; +x_11 = x_111; goto _start; } } @@ -39307,18 +39490,18 @@ lean_dec(x_1); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_collectUnassignedMVars_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, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_collectUnassignedMVars_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, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_11; -x_11 = l_Lean_Elab_Term_collectUnassignedMVars_go(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_Lean_Elab_Term_collectUnassignedMVars_go(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); -return x_11; +return x_12; } } LEAN_EXPORT lean_object* l_Lean_Elab_Term_collectUnassignedMVars(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { @@ -39338,7 +39521,8 @@ if (x_15 == 0) lean_object* x_16; lean_object* x_17; lean_free_object(x_11); x_16 = lean_array_to_list(lean_box(0), x_13); -x_17 = l_Lean_Elab_Term_collectUnassignedMVars_go(x_3, x_16, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_14); +lean_inc(x_2); +x_17 = l_Lean_Elab_Term_collectUnassignedMVars_go(x_3, x_16, x_2, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_14); return x_17; } else @@ -39362,7 +39546,8 @@ if (x_20 == 0) { lean_object* x_21; lean_object* x_22; x_21 = lean_array_to_list(lean_box(0), x_18); -x_22 = l_Lean_Elab_Term_collectUnassignedMVars_go(x_3, x_21, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_19); +lean_inc(x_2); +x_22 = l_Lean_Elab_Term_collectUnassignedMVars_go(x_3, x_21, x_2, x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_19); return x_22; } else @@ -45788,7 +45973,7 @@ lean_dec(x_3); return x_10; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__1() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__1() { _start: { lean_object* x_1; @@ -45796,17 +45981,17 @@ 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_18409____closed__2() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_Term_mkTermElabAttributeUnsafe___closed__9; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__1; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__1; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__3() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -45816,27 +46001,27 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__4() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__3; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__3; x_2 = l_Lean_Elab_Term_mkTermElabAttributeUnsafe___closed__9; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__5() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__4; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__4; x_2 = l_Lean_Elab_Term_mkTermElabAttributeUnsafe___closed__7; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__6() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__6() { _start: { lean_object* x_1; @@ -45844,17 +46029,17 @@ x_1 = lean_mk_string_from_bytes("initFn", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__7() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__5; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__6; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__5; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__8() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__8() { _start: { lean_object* x_1; @@ -45862,47 +46047,47 @@ x_1 = lean_mk_string_from_bytes("_@", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__9() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__7; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__8; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__7; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__8; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__10() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__9; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__9; x_2 = l_Lean_Elab_Term_mkTermElabAttributeUnsafe___closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__11() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__10; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__10; x_2 = l_Lean_Elab_Term_mkTermElabAttributeUnsafe___closed__9; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__12() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__11; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__11; x_2 = l_Lean_Elab_Term_mkTermElabAttributeUnsafe___closed__7; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__13() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__13() { _start: { lean_object* x_1; @@ -45910,33 +46095,33 @@ x_1 = lean_mk_string_from_bytes("_hyg", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__14() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__12; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__13; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__12; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__13; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__15() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__14; -x_2 = lean_unsigned_to_nat(18409u); +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__14; +x_2 = lean_unsigned_to_nat(18351u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__2; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__2; x_3 = 0; -x_4 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__15; +x_4 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__15; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } @@ -46284,7 +46469,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_isLetRecAuxMVar(lean_object* x_1, lean _start: { lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_9 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__2; +x_9 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__2; x_10 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Term_traceAtCmdPos___spec__1(x_9, x_2, x_3, x_4, x_5, x_6, x_7, x_8); x_11 = lean_ctor_get(x_10, 0); lean_inc(x_11); @@ -46730,7 +46915,7 @@ LEAN_EXPORT lean_object* l_Lean_logWarning___at___private_Lean_Elab_Term_0__Lean lean_object* x_9; lean_object* x_10; uint8_t x_11; x_9 = lean_ctor_get(x_6, 2); x_10 = l_Lean_logAt___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__2___closed__1; -x_11 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_9, x_10); +x_11 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_9, x_10); if (x_11 == 0) { uint8_t x_12; lean_object* x_13; @@ -47469,7 +47654,7 @@ if (x_32 == 0) { lean_object* x_33; uint8_t x_34; uint8_t x_35; x_33 = l_Lean_Elab_Term_resolveName_process___closed__3; -x_34 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_29, x_33); +x_34 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_29, x_33); lean_dec(x_29); lean_inc(x_1); x_35 = l_Lean_Elab_isValidAutoBoundImplicitName(x_1, x_34); @@ -50846,7 +51031,7 @@ x_7 = lean_alloc_closure((void*)(l_Lean_Elab_Term_instMetaEvalTermElabM___rarg__ lean_closure_set(x_7, 0, x_4); x_8 = lean_box(0); x_31 = l_Lean_Elab_Term_instMetaEvalTermElabM___rarg___closed__25; -x_32 = l_Lean_Option_get___at_Std_Format_pretty_x27___spec__1(x_3, x_31); +x_32 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_3, x_31); x_33 = l_Lean_Core_getMaxHeartbeats(x_3); x_34 = l_Lean_Elab_Term_instMetaEvalTermElabM___rarg___closed__26; x_35 = l_Lean_Elab_Term_instMetaEvalTermElabM___rarg___closed__29; @@ -54565,77 +54750,77 @@ 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_21681____closed__1() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__4; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__6; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__4; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681____closed__2() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681____closed__1; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__8; +x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623____closed__1; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__8; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681____closed__3() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681____closed__2; +x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623____closed__2; x_2 = l_Lean_Elab_Term_mkTermElabAttributeUnsafe___closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681____closed__4() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681____closed__3; +x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623____closed__3; x_2 = l_Lean_Elab_Term_mkTermElabAttributeUnsafe___closed__9; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681____closed__5() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681____closed__4; +x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623____closed__4; x_2 = l_Lean_Elab_Term_mkTermElabAttributeUnsafe___closed__7; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681____closed__6() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681____closed__5; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__13; +x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623____closed__5; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__13; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681____closed__7() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681____closed__6; -x_2 = lean_unsigned_to_nat(21681u); +x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623____closed__6; +x_2 = lean_unsigned_to_nat(21623u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681____closed__8() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623____closed__8() { _start: { lean_object* x_1; @@ -54643,23 +54828,23 @@ 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_21681____closed__9() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_Term_mkTermElabAttributeUnsafe___closed__9; -x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681____closed__8; +x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623____closed__8; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; x_2 = l___private_Lean_Elab_Term_0__Lean_Elab_Term_postponeElabTermCore___closed__2; x_3 = 0; -x_4 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681____closed__7; +x_4 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623____closed__7; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); if (lean_obj_tag(x_5) == 0) { @@ -54675,7 +54860,7 @@ lean_object* x_9; lean_object* x_10; lean_object* x_11; x_9 = lean_ctor_get(x_8, 1); lean_inc(x_9); lean_dec(x_8); -x_10 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681____closed__9; +x_10 = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623____closed__9; x_11 = l_Lean_registerTraceClass(x_10, x_3, x_4, x_9); return x_11; } @@ -55120,24 +55305,37 @@ l_Lean_Elab_Term_synthesizeInstMVarCore___closed__4 = _init_l_Lean_Elab_Term_syn lean_mark_persistent(l_Lean_Elab_Term_synthesizeInstMVarCore___closed__4); l_Lean_Elab_Term_synthesizeInstMVarCore___closed__5 = _init_l_Lean_Elab_Term_synthesizeInstMVarCore___closed__5(); lean_mark_persistent(l_Lean_Elab_Term_synthesizeInstMVarCore___closed__5); +l_Lean_withSeconds___at_Lean_Elab_Term_mkCoe___spec__3___closed__1 = _init_l_Lean_withSeconds___at_Lean_Elab_Term_mkCoe___spec__3___closed__1(); +l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__3___closed__1 = _init_l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__3___closed__1); +l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__3___closed__2 = _init_l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__3___closed__2); +l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__3___closed__3 = _init_l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__3___closed__3(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__3___closed__3); +l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__3___closed__4 = _init_l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__3___closed__4(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__3___closed__4); +l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__3___closed__5 = _init_l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__3___closed__5(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__3___closed__5); +l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__4___closed__1 = _init_l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__4___closed__1); l_Lean_Elab_Term_mkCoe___lambda__1___closed__1 = _init_l_Lean_Elab_Term_mkCoe___lambda__1___closed__1(); lean_mark_persistent(l_Lean_Elab_Term_mkCoe___lambda__1___closed__1); l_Lean_Elab_Term_mkCoe___lambda__1___closed__2 = _init_l_Lean_Elab_Term_mkCoe___lambda__1___closed__2(); lean_mark_persistent(l_Lean_Elab_Term_mkCoe___lambda__1___closed__2); +l_Lean_Elab_Term_mkCoe___lambda__1___closed__3 = _init_l_Lean_Elab_Term_mkCoe___lambda__1___closed__3(); +lean_mark_persistent(l_Lean_Elab_Term_mkCoe___lambda__1___closed__3); +l_Lean_Elab_Term_mkCoe___lambda__1___closed__4 = _init_l_Lean_Elab_Term_mkCoe___lambda__1___closed__4(); +lean_mark_persistent(l_Lean_Elab_Term_mkCoe___lambda__1___closed__4); +l_Lean_Elab_Term_mkCoe___lambda__1___closed__5 = _init_l_Lean_Elab_Term_mkCoe___lambda__1___closed__5(); +lean_mark_persistent(l_Lean_Elab_Term_mkCoe___lambda__1___closed__5); +l_Lean_Elab_Term_mkCoe___lambda__1___closed__6 = _init_l_Lean_Elab_Term_mkCoe___lambda__1___closed__6(); +lean_mark_persistent(l_Lean_Elab_Term_mkCoe___lambda__1___closed__6); +l_Lean_Elab_Term_mkCoe___lambda__2___closed__1 = _init_l_Lean_Elab_Term_mkCoe___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Elab_Term_mkCoe___lambda__2___closed__1); +l_Lean_Elab_Term_mkCoe___lambda__2___closed__2 = _init_l_Lean_Elab_Term_mkCoe___lambda__2___closed__2(); +lean_mark_persistent(l_Lean_Elab_Term_mkCoe___lambda__2___closed__2); l_Lean_Elab_Term_mkCoe___closed__1 = _init_l_Lean_Elab_Term_mkCoe___closed__1(); lean_mark_persistent(l_Lean_Elab_Term_mkCoe___closed__1); -l_Lean_Elab_Term_mkCoe___closed__2 = _init_l_Lean_Elab_Term_mkCoe___closed__2(); -lean_mark_persistent(l_Lean_Elab_Term_mkCoe___closed__2); -l_Lean_Elab_Term_mkCoe___closed__3 = _init_l_Lean_Elab_Term_mkCoe___closed__3(); -lean_mark_persistent(l_Lean_Elab_Term_mkCoe___closed__3); -l_Lean_Elab_Term_mkCoe___closed__4 = _init_l_Lean_Elab_Term_mkCoe___closed__4(); -lean_mark_persistent(l_Lean_Elab_Term_mkCoe___closed__4); -l_Lean_Elab_Term_mkCoe___closed__5 = _init_l_Lean_Elab_Term_mkCoe___closed__5(); -lean_mark_persistent(l_Lean_Elab_Term_mkCoe___closed__5); -l_Lean_Elab_Term_mkCoe___closed__6 = _init_l_Lean_Elab_Term_mkCoe___closed__6(); -lean_mark_persistent(l_Lean_Elab_Term_mkCoe___closed__6); -l_Lean_Elab_Term_mkCoe___closed__7 = _init_l_Lean_Elab_Term_mkCoe___closed__7(); -lean_mark_persistent(l_Lean_Elab_Term_mkCoe___closed__7); l___private_Lean_Elab_Term_0__Lean_Elab_Term_mkSyntheticSorryFor___closed__1 = _init_l___private_Lean_Elab_Term_0__Lean_Elab_Term_mkSyntheticSorryFor___closed__1(); lean_mark_persistent(l___private_Lean_Elab_Term_0__Lean_Elab_Term_mkSyntheticSorryFor___closed__1); l_Lean_Elab_throwPostpone___at_Lean_Elab_Term_tryPostpone___spec__1___rarg___closed__1 = _init_l_Lean_Elab_throwPostpone___at_Lean_Elab_Term_tryPostpone___spec__1___rarg___closed__1(); @@ -55270,19 +55468,6 @@ l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTe lean_mark_persistent(l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__5___closed__2); l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__6___rarg___closed__1 = _init_l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__6___rarg___closed__1(); lean_mark_persistent(l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__6___rarg___closed__1); -l_Lean_withSeconds___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__18___closed__1 = _init_l_Lean_withSeconds___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__18___closed__1(); -l_Lean_withOptProfile___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__16___closed__1 = _init_l_Lean_withOptProfile___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__16___closed__1(); -lean_mark_persistent(l_Lean_withOptProfile___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__16___closed__1); -l_Lean_withOptProfile___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__16___closed__2 = _init_l_Lean_withOptProfile___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__16___closed__2(); -lean_mark_persistent(l_Lean_withOptProfile___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__16___closed__2); -l_Lean_withTraceNode___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__14___closed__1 = _init_l_Lean_withTraceNode___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__14___closed__1(); -lean_mark_persistent(l_Lean_withTraceNode___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__14___closed__1); -l_Lean_withTraceNode___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__14___closed__2 = _init_l_Lean_withTraceNode___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__14___closed__2(); -lean_mark_persistent(l_Lean_withTraceNode___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__14___closed__2); -l_Lean_withTraceNode___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__14___closed__3 = _init_l_Lean_withTraceNode___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__14___closed__3(); -lean_mark_persistent(l_Lean_withTraceNode___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__14___closed__3); -l_Lean_withTraceNode___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__14___closed__4 = _init_l_Lean_withTraceNode___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__14___closed__4(); -lean_mark_persistent(l_Lean_withTraceNode___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__14___closed__4); l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__1___closed__1 = _init_l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__1___closed__1(); lean_mark_persistent(l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__1___closed__1); l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__1___closed__2 = _init_l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__1___closed__2(); @@ -55334,37 +55519,37 @@ 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_18409____closed__1 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__1(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__1); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__2 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__2(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__2); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__3 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__3(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__3); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__4 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__4(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__4); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__5 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__5(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__5); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__6 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__6(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__6); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__7 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__7(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__7); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__8 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__8(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__8); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__9 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__9(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__9); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__10 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__10(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__10); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__11 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__11(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__11); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__12 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__12(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__12); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__13 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__13(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__13); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__14 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__14(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__14); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__15 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__15(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409____closed__15); -res = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18409_(lean_io_mk_world()); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__1 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__1(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__1); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__2 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__2(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__2); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__3 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__3(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__3); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__4 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__4(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__4); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__5 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__5(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__5); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__6 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__6(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__6); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__7 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__7(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__7); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__8 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__8(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__8); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__9 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__9(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__9); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__10 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__10(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__10); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__11 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__11(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__11); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__12 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__12(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__12); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__13 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__13(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__13); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__14 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__14(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__14); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__15 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__15(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351____closed__15); +res = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_18351_(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(); @@ -55579,25 +55764,25 @@ l_Lean_Elab_Term_exprToSyntax___lambda__1___closed__3 = _init_l_Lean_Elab_Term_e lean_mark_persistent(l_Lean_Elab_Term_exprToSyntax___lambda__1___closed__3); l_Lean_Elab_Term_exprToSyntax___lambda__1___closed__4 = _init_l_Lean_Elab_Term_exprToSyntax___lambda__1___closed__4(); lean_mark_persistent(l_Lean_Elab_Term_exprToSyntax___lambda__1___closed__4); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681____closed__1 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681____closed__1(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681____closed__1); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681____closed__2 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681____closed__2(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681____closed__2); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681____closed__3 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681____closed__3(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681____closed__3); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681____closed__4 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681____closed__4(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681____closed__4); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681____closed__5 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681____closed__5(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681____closed__5); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681____closed__6 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681____closed__6(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681____closed__6); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681____closed__7 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681____closed__7(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681____closed__7); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681____closed__8 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681____closed__8(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681____closed__8); -l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681____closed__9 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681____closed__9(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681____closed__9); -res = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21681_(lean_io_mk_world()); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623____closed__1 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623____closed__1(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623____closed__1); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623____closed__2 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623____closed__2(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623____closed__2); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623____closed__3 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623____closed__3(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623____closed__3); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623____closed__4 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623____closed__4(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623____closed__4); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623____closed__5 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623____closed__5(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623____closed__5); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623____closed__6 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623____closed__6(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623____closed__6); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623____closed__7 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623____closed__7(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623____closed__7); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623____closed__8 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623____closed__8(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623____closed__8); +l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623____closed__9 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623____closed__9(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623____closed__9); +res = l_Lean_Elab_initFn____x40_Lean_Elab_Term___hyg_21623_(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 a5f9df4cbced..8e263589c2ca 100644 --- a/stage0/stdlib/Lean/Elab/Util.c +++ b/stage0/stdlib/Lean/Elab/Util.c @@ -15,7 +15,6 @@ extern "C" { #endif lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); lean_object* l_Lean_evalPrio(lean_object*, lean_object*, lean_object*); -uint8_t l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2981____closed__3; lean_object* l_Lean_addTrace___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___rarg___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -183,6 +182,7 @@ static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_370____closed LEAN_EXPORT lean_object* l_Lean_Elab_logException(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_macroAttribute; LEAN_EXPORT lean_object* l_Lean_Elab_pp_macroStack; +uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_checkSyntaxNodeKindAtNamespaces___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Elab_Util___hyg_843____closed__4; @@ -202,7 +202,6 @@ static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_1375____close static lean_object* l_Lean_Elab_nestedExceptionToMessageData___rarg___lambda__1___closed__3; uint8_t l_Lean_Environment_contains(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_Elab_liftMacroM___spec__2(lean_object*); -lean_object* l_Lean_Option_register___at_Std_Format_initFn____x40_Lean_Data_Format___hyg_95____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_expandOptNamedPrio___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Util_0__Lean_Elab_evalSyntaxConstantUnsafe___closed__2; static lean_object* l___auto____x40_Lean_Elab_Util___hyg_843____closed__2; @@ -258,6 +257,7 @@ static lean_object* l___auto____x40_Lean_Elab_Util___hyg_843____closed__25; uint8_t l_Lean_Syntax_isNone(lean_object*); LEAN_EXPORT lean_object* l_List_forM___at_Lean_Elab_liftMacroM___spec__3(lean_object*); lean_object* l_Lean_ResolveName_resolveGlobalName(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_logException___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_addMacroStack(lean_object*); static lean_object* l_Lean_Elab_mkElabAttribute___rarg___lambda__3___closed__8; @@ -689,7 +689,7 @@ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_370____closed__3; x_3 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_370____closed__5; x_4 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_370____closed__7; -x_5 = l_Lean_Option_register___at_Std_Format_initFn____x40_Lean_Data_Format___hyg_95____spec__1(x_2, x_3, x_4, x_1); +x_5 = l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__1(x_2, x_3, x_4, x_1); return x_5; } } @@ -810,7 +810,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_addMacroStack___rarg___lambda__1(lean_objec { lean_object* x_5; uint8_t x_6; x_5 = l_Lean_Elab_addMacroStack___rarg___lambda__1___closed__1; -x_6 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_4, x_5); +x_6 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_4, x_5); if (x_6 == 0) { lean_object* x_7; lean_object* x_8; lean_object* x_9; diff --git a/stage0/stdlib/Lean/Exception.c b/stage0/stdlib/Lean/Exception.c index d7e6941b8eba..a202829dbc26 100644 --- a/stage0/stdlib/Lean/Exception.c +++ b/stage0/stdlib/Lean/Exception.c @@ -13,23 +13,22 @@ #ifdef __cplusplus extern "C" { #endif -static lean_object* l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__6; lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); static lean_object* l_Lean_throwUnknownConstant___rarg___closed__1; -static lean_object* l_Lean_termThrowErrorAt_________closed__1; LEAN_EXPORT lean_object* l_Lean_Exception_getRef___boxed(lean_object*); +static lean_object* l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt__________1___closed__8; LEAN_EXPORT lean_object* l_Lean_throwError___rarg___lambda__1(lean_object*, lean_object*); static lean_object* l_Lean_throwMaxRecDepthAt___rarg___closed__2; +static lean_object* l_Lean_termThrowErrorAt___________closed__1; LEAN_EXPORT uint8_t l_Lean_Exception_isMaxRecDepth(lean_object*); extern lean_object* l_Lean_maxRecDepthErrorMessage; static lean_object* l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__21; +static lean_object* l_Lean_termThrowErrorAt___________closed__7; static lean_object* l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__22; static lean_object* l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__8; -static lean_object* l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__5; static lean_object* l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__10; -static lean_object* l_Lean_termThrowErrorAt_________closed__7; +static lean_object* l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt__________1___closed__4; static lean_object* l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__4; -LEAN_EXPORT lean_object* l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_replaceRef(lean_object*, lean_object*); static lean_object* l_Lean_throwMaxRecDepthAt___rarg___closed__1; static lean_object* l_Lean_termThrowError_______closed__4; @@ -49,27 +48,30 @@ LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___rarg(lean_object*, lean_obj static lean_object* l_Lean_instInhabitedException___closed__1; LEAN_EXPORT lean_object* l_Lean_throwError___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_string_dec_eq(lean_object*, lean_object*); +static lean_object* l_Lean_termThrowErrorAt___________closed__6; LEAN_EXPORT lean_object* l_Lean_Exception_toMessageData(lean_object*); 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__17; LEAN_EXPORT lean_object* l_Lean_throwKernelException___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_termThrowErrorAt________; LEAN_EXPORT lean_object* l_Lean_throwError(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt(lean_object*, lean_object*); static lean_object* l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__11; +static lean_object* l_Lean_termThrowErrorAt___________closed__11; static lean_object* l_Lean_instInhabitedException___closed__2; static lean_object* l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__1; LEAN_EXPORT lean_object* l_Lean_ofExcept___rarg(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instAddErrorMessageContext___rarg___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instAddErrorMessageContext(lean_object*); +static lean_object* l_Lean_termThrowErrorAt___________closed__10; LEAN_EXPORT lean_object* l_Lean_instMonadRecDepthMonadCacheT___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*, uint8_t); LEAN_EXPORT lean_object* l_Lean_instMonadRecDepthReaderT___rarg(lean_object*); static lean_object* l_Lean_termThrowError_______closed__17; LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt(lean_object*, lean_object*); static lean_object* l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__23; -static lean_object* l_Lean_termThrowErrorAt_________closed__2; -static lean_object* l_Lean_termThrowErrorAt_________closed__6; +static lean_object* l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt__________1___closed__6; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withIncRecDepth___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Exception_getRef(lean_object*); @@ -84,9 +86,8 @@ static lean_object* l_Lean___aux__Lean__Exception______macroRules__Lean__termThr LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_termThrowError_______closed__13; lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__1; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__8; +static lean_object* l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt__________1___closed__2; lean_object* l_Lean_Syntax_node2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instMonadRecDepthMonadCacheT___rarg(lean_object*); lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); @@ -94,18 +95,20 @@ static lean_object* l_Lean_termThrowError_______closed__3; static lean_object* l_Lean_termThrowError_______closed__10; LEAN_EXPORT lean_object* l_Lean_ofExceptKernelException(lean_object*, lean_object*); static lean_object* l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__6; +static lean_object* l_Lean_termThrowErrorAt___________closed__4; LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant(lean_object*, lean_object*); static lean_object* l_Lean_termThrowError_______closed__16; -static lean_object* l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__7; +static lean_object* l_Lean_termThrowErrorAt___________closed__9; static lean_object* l_Lean_termThrowError_______closed__18; +LEAN_EXPORT lean_object* l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt__________1(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofExpr(lean_object*); static lean_object* l_Lean_termThrowError_______closed__12; LEAN_EXPORT lean_object* l_Lean_withIncRecDepth___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_termThrowError_______closed__14; +static lean_object* l_Lean_termThrowErrorAt___________closed__5; LEAN_EXPORT lean_object* l_Lean_instMonadRecDepthReaderT___rarg___lambda__3___boxed(lean_object*, lean_object*); static lean_object* l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__18; -static lean_object* l_Lean_termThrowErrorAt_________closed__3; LEAN_EXPORT lean_object* l_Lean_instAddErrorMessageContext___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instMonadRecDepthReaderT___rarg___lambda__3(lean_object*, lean_object*); @@ -113,18 +116,18 @@ LEAN_EXPORT lean_object* l_Lean_throwKernelException___rarg___lambda__1(lean_obj static lean_object* l_Lean_termThrowError_______closed__1; LEAN_EXPORT lean_object* l_Lean_Exception_isMaxRecDepth___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_ofExcept(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_termThrowErrorAt_________closed__8; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withIncRecDepth___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_termThrowErrorAt___________closed__8; lean_object* l_Lean_InternalExceptionId_toString(lean_object*); lean_object* l_Lean_Syntax_node1(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt__________1___closed__3; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_termThrowError_______closed__7; +static lean_object* l_Lean_termThrowErrorAt___________closed__12; static lean_object* l_Lean_termThrowError_______closed__6; LEAN_EXPORT lean_object* l_Lean_instMonadRecDepthReaderT(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__4; -static lean_object* l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__3; -static lean_object* l_Lean_termThrowErrorAt_________closed__4; +static lean_object* l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt__________1___closed__7; static lean_object* l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__9; static lean_object* l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__16; static lean_object* l_Lean_termThrowError_______closed__5; @@ -139,19 +142,20 @@ lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_o static lean_object* l_Lean_termThrowError_______closed__2; static lean_object* l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__14; static lean_object* l_Lean_termThrowError_______closed__8; +static lean_object* l_Lean_termThrowErrorAt___________closed__3; static lean_object* l_Lean_throwUnknownConstant___rarg___closed__2; +static lean_object* l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt__________1___closed__5; lean_object* lean_nat_add(lean_object*, lean_object*); -static lean_object* l_Lean_termThrowErrorAt_________closed__5; +static lean_object* l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt__________1___closed__1; LEAN_EXPORT lean_object* l_Lean_withIncRecDepth(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instMonadRecDepthStateRefT_x27___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_KernelException_toMessageData(lean_object*, lean_object*); lean_object* l_String_toSubstring_x27(lean_object*); LEAN_EXPORT lean_object* l_Lean_instMonadRecDepthReaderT___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__15; -static lean_object* l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__2; -LEAN_EXPORT lean_object* l_Lean_termThrowErrorAt______; static lean_object* l_Lean_termThrowError_______closed__9; LEAN_EXPORT lean_object* l_Lean_throwError___rarg___lambda__1___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_termThrowErrorAt___________closed__2; LEAN_EXPORT lean_object* l_Lean_instMonadRecDepthStateRefT_x27___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Exception_toMessageData(lean_object* x_1) { _start: @@ -1110,25 +1114,25 @@ x_1 = l_Lean_termThrowError_______closed__18; return x_1; } } -static lean_object* _init_l_Lean_termThrowErrorAt_________closed__1() { +static lean_object* _init_l_Lean_termThrowErrorAt___________closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("termThrowErrorAt___", 19); +x_1 = lean_mk_string_from_bytes("termThrowErrorAt____", 20); return x_1; } } -static lean_object* _init_l_Lean_termThrowErrorAt_________closed__2() { +static lean_object* _init_l_Lean_termThrowErrorAt___________closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_termThrowError_______closed__1; -x_2 = l_Lean_termThrowErrorAt_________closed__1; +x_2 = l_Lean_termThrowErrorAt___________closed__1; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_termThrowErrorAt_________closed__3() { +static lean_object* _init_l_Lean_termThrowErrorAt___________closed__3() { _start: { lean_object* x_1; @@ -1136,17 +1140,17 @@ x_1 = lean_mk_string_from_bytes("throwErrorAt ", 13); return x_1; } } -static lean_object* _init_l_Lean_termThrowErrorAt_________closed__4() { +static lean_object* _init_l_Lean_termThrowErrorAt___________closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_termThrowErrorAt_________closed__3; +x_1 = l_Lean_termThrowErrorAt___________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_termThrowErrorAt_________closed__5() { +static lean_object* _init_l_Lean_termThrowErrorAt___________closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -1158,13 +1162,13 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_termThrowErrorAt_________closed__6() { +static lean_object* _init_l_Lean_termThrowErrorAt___________closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_termThrowError_______closed__5; -x_2 = l_Lean_termThrowErrorAt_________closed__4; -x_3 = l_Lean_termThrowErrorAt_________closed__5; +x_2 = l_Lean_termThrowErrorAt___________closed__4; +x_3 = l_Lean_termThrowErrorAt___________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); @@ -1172,12 +1176,54 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_termThrowErrorAt_________closed__7() { +static lean_object* _init_l_Lean_termThrowErrorAt___________closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("ppSpace", 7); +return x_1; +} +} +static lean_object* _init_l_Lean_termThrowErrorAt___________closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_termThrowErrorAt___________closed__7; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_termThrowErrorAt___________closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_termThrowErrorAt___________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_termThrowErrorAt___________closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_termThrowError_______closed__5; +x_2 = l_Lean_termThrowErrorAt___________closed__6; +x_3 = l_Lean_termThrowErrorAt___________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); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_termThrowErrorAt___________closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_termThrowError_______closed__5; -x_2 = l_Lean_termThrowErrorAt_________closed__6; +x_2 = l_Lean_termThrowErrorAt___________closed__10; x_3 = l_Lean_termThrowError_______closed__16; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); @@ -1186,13 +1232,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_termThrowErrorAt_________closed__8() { +static lean_object* _init_l_Lean_termThrowErrorAt___________closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_termThrowErrorAt_________closed__2; +x_1 = l_Lean_termThrowErrorAt___________closed__2; x_2 = lean_unsigned_to_nat(1022u); -x_3 = l_Lean_termThrowErrorAt_________closed__7; +x_3 = l_Lean_termThrowErrorAt___________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); @@ -1200,11 +1246,11 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_termThrowErrorAt______() { +static lean_object* _init_l_Lean_termThrowErrorAt________() { _start: { lean_object* x_1; -x_1 = l_Lean_termThrowErrorAt_________closed__8; +x_1 = l_Lean_termThrowErrorAt___________closed__12; return x_1; } } @@ -1538,7 +1584,7 @@ return x_51; } } } -static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__1() { +static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt__________1___closed__1() { _start: { lean_object* x_1; @@ -1546,16 +1592,16 @@ x_1 = lean_mk_string_from_bytes("Lean.throwErrorAt", 17); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__2() { +static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt__________1___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__1; +x_1 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt__________1___closed__1; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__3() { +static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt__________1___closed__3() { _start: { lean_object* x_1; @@ -1563,67 +1609,67 @@ x_1 = lean_mk_string_from_bytes("throwErrorAt", 12); return x_1; } } -static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__4() { +static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt__________1___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_termThrowError_______closed__1; -x_2 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__3; +x_2 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt__________1___closed__3; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__5() { +static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt__________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___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__4; +x_2 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt__________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_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__6() { +static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt__________1___closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__4; +x_1 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt__________1___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___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__7() { +static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt__________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__Exception______macroRules__Lean__termThrowErrorAt________1___closed__6; +x_2 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt__________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___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__8() { +static lean_object* _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt__________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__termThrowErrorAt________1___closed__5; -x_2 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__7; +x_1 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt__________1___closed__5; +x_2 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt__________1___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; } } -LEAN_EXPORT lean_object* l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt__________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_termThrowErrorAt_________closed__2; +x_4 = l_Lean_termThrowErrorAt___________closed__2; lean_inc(x_1); x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); if (x_5 == 0) @@ -1660,10 +1706,10 @@ lean_inc(x_17); x_18 = lean_ctor_get(x_2, 1); lean_inc(x_18); lean_dec(x_2); -x_19 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__4; +x_19 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt__________1___closed__4; x_20 = l_Lean_addMacroScope(x_18, x_19, x_17); -x_21 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__2; -x_22 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__8; +x_21 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt__________1___closed__2; +x_22 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt__________1___closed__8; lean_inc(x_16); x_23 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_23, 0, x_16); @@ -1692,10 +1738,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___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__4; +x_34 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt__________1___closed__4; x_35 = l_Lean_addMacroScope(x_33, x_34, x_32); -x_36 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__2; -x_37 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__8; +x_36 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt__________1___closed__2; +x_37 = l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt__________1___closed__8; lean_inc(x_31); x_38 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_38, 0, x_31); @@ -1817,24 +1863,32 @@ l_Lean_termThrowError_______closed__18 = _init_l_Lean_termThrowError_______close lean_mark_persistent(l_Lean_termThrowError_______closed__18); l_Lean_termThrowError____ = _init_l_Lean_termThrowError____(); lean_mark_persistent(l_Lean_termThrowError____); -l_Lean_termThrowErrorAt_________closed__1 = _init_l_Lean_termThrowErrorAt_________closed__1(); -lean_mark_persistent(l_Lean_termThrowErrorAt_________closed__1); -l_Lean_termThrowErrorAt_________closed__2 = _init_l_Lean_termThrowErrorAt_________closed__2(); -lean_mark_persistent(l_Lean_termThrowErrorAt_________closed__2); -l_Lean_termThrowErrorAt_________closed__3 = _init_l_Lean_termThrowErrorAt_________closed__3(); -lean_mark_persistent(l_Lean_termThrowErrorAt_________closed__3); -l_Lean_termThrowErrorAt_________closed__4 = _init_l_Lean_termThrowErrorAt_________closed__4(); -lean_mark_persistent(l_Lean_termThrowErrorAt_________closed__4); -l_Lean_termThrowErrorAt_________closed__5 = _init_l_Lean_termThrowErrorAt_________closed__5(); -lean_mark_persistent(l_Lean_termThrowErrorAt_________closed__5); -l_Lean_termThrowErrorAt_________closed__6 = _init_l_Lean_termThrowErrorAt_________closed__6(); -lean_mark_persistent(l_Lean_termThrowErrorAt_________closed__6); -l_Lean_termThrowErrorAt_________closed__7 = _init_l_Lean_termThrowErrorAt_________closed__7(); -lean_mark_persistent(l_Lean_termThrowErrorAt_________closed__7); -l_Lean_termThrowErrorAt_________closed__8 = _init_l_Lean_termThrowErrorAt_________closed__8(); -lean_mark_persistent(l_Lean_termThrowErrorAt_________closed__8); -l_Lean_termThrowErrorAt______ = _init_l_Lean_termThrowErrorAt______(); -lean_mark_persistent(l_Lean_termThrowErrorAt______); +l_Lean_termThrowErrorAt___________closed__1 = _init_l_Lean_termThrowErrorAt___________closed__1(); +lean_mark_persistent(l_Lean_termThrowErrorAt___________closed__1); +l_Lean_termThrowErrorAt___________closed__2 = _init_l_Lean_termThrowErrorAt___________closed__2(); +lean_mark_persistent(l_Lean_termThrowErrorAt___________closed__2); +l_Lean_termThrowErrorAt___________closed__3 = _init_l_Lean_termThrowErrorAt___________closed__3(); +lean_mark_persistent(l_Lean_termThrowErrorAt___________closed__3); +l_Lean_termThrowErrorAt___________closed__4 = _init_l_Lean_termThrowErrorAt___________closed__4(); +lean_mark_persistent(l_Lean_termThrowErrorAt___________closed__4); +l_Lean_termThrowErrorAt___________closed__5 = _init_l_Lean_termThrowErrorAt___________closed__5(); +lean_mark_persistent(l_Lean_termThrowErrorAt___________closed__5); +l_Lean_termThrowErrorAt___________closed__6 = _init_l_Lean_termThrowErrorAt___________closed__6(); +lean_mark_persistent(l_Lean_termThrowErrorAt___________closed__6); +l_Lean_termThrowErrorAt___________closed__7 = _init_l_Lean_termThrowErrorAt___________closed__7(); +lean_mark_persistent(l_Lean_termThrowErrorAt___________closed__7); +l_Lean_termThrowErrorAt___________closed__8 = _init_l_Lean_termThrowErrorAt___________closed__8(); +lean_mark_persistent(l_Lean_termThrowErrorAt___________closed__8); +l_Lean_termThrowErrorAt___________closed__9 = _init_l_Lean_termThrowErrorAt___________closed__9(); +lean_mark_persistent(l_Lean_termThrowErrorAt___________closed__9); +l_Lean_termThrowErrorAt___________closed__10 = _init_l_Lean_termThrowErrorAt___________closed__10(); +lean_mark_persistent(l_Lean_termThrowErrorAt___________closed__10); +l_Lean_termThrowErrorAt___________closed__11 = _init_l_Lean_termThrowErrorAt___________closed__11(); +lean_mark_persistent(l_Lean_termThrowErrorAt___________closed__11); +l_Lean_termThrowErrorAt___________closed__12 = _init_l_Lean_termThrowErrorAt___________closed__12(); +lean_mark_persistent(l_Lean_termThrowErrorAt___________closed__12); +l_Lean_termThrowErrorAt________ = _init_l_Lean_termThrowErrorAt________(); +lean_mark_persistent(l_Lean_termThrowErrorAt________); l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__1 = _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__1(); lean_mark_persistent(l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__1); l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__2 = _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__2(); @@ -1881,22 +1935,22 @@ l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___clo lean_mark_persistent(l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__22); l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__23 = _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__23(); lean_mark_persistent(l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowError______1___closed__23); -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(); -lean_mark_persistent(l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__2); -l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__3 = _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__3(); -lean_mark_persistent(l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__3); -l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__4 = _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__4(); -lean_mark_persistent(l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__4); -l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__5 = _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__5(); -lean_mark_persistent(l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__5); -l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__6 = _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__6(); -lean_mark_persistent(l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__6); -l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__7 = _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__7(); -lean_mark_persistent(l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__7); -l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__8 = _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__8(); -lean_mark_persistent(l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt________1___closed__8); +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(); +lean_mark_persistent(l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt__________1___closed__2); +l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt__________1___closed__3 = _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt__________1___closed__3(); +lean_mark_persistent(l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt__________1___closed__3); +l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt__________1___closed__4 = _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt__________1___closed__4(); +lean_mark_persistent(l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt__________1___closed__4); +l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt__________1___closed__5 = _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt__________1___closed__5(); +lean_mark_persistent(l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt__________1___closed__5); +l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt__________1___closed__6 = _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt__________1___closed__6(); +lean_mark_persistent(l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt__________1___closed__6); +l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt__________1___closed__7 = _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt__________1___closed__7(); +lean_mark_persistent(l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt__________1___closed__7); +l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt__________1___closed__8 = _init_l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt__________1___closed__8(); +lean_mark_persistent(l_Lean___aux__Lean__Exception______macroRules__Lean__termThrowErrorAt__________1___closed__8); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lean/Expr.c b/stage0/stdlib/Lean/Expr.c index 64f2914ae27a..757ed758b7b5 100644 --- a/stage0/stdlib/Lean/Expr.c +++ b/stage0/stdlib/Lean/Expr.c @@ -284,6 +284,7 @@ LEAN_EXPORT lean_object* l_Lean_Expr_replaceFVarId(lean_object*, lean_object*, l LEAN_EXPORT lean_object* l_Lean_Expr_betaRev_go(lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Expr_containsFVar(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Expr_hasMVar(lean_object*); +LEAN_EXPORT uint8_t l_Lean_Expr_isSemiOutParam(lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_Expr_constLevels_x21___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_isLHSGoal_x3f(lean_object*); LEAN_EXPORT lean_object* l_Lean_Expr_getAutoParamTactic_x3f___boxed(lean_object*); @@ -454,6 +455,7 @@ LEAN_EXPORT lean_object* l_Lean_Expr_instantiateRange___boxed(lean_object*, lean static lean_object* l___private_Lean_Expr_0__Lean_reprExpr____x40_Lean_Expr___hyg_3663____closed__35; static lean_object* l___private_Lean_Expr_0__Lean_reprExpr____x40_Lean_Expr___hyg_3663____closed__30; static lean_object* l___private_Lean_Expr_0__Lean_reprExpr____x40_Lean_Expr___hyg_3663____closed__31; +static lean_object* l_Lean_Expr_isSemiOutParam___closed__1; uint16_t lean_uint8_to_uint16(uint8_t); static lean_object* l___private_Lean_Expr_0__Lean_reprBinderInfo____x40_Lean_Expr___hyg_497____closed__2; LEAN_EXPORT lean_object* l_Lean_Expr_casesOn___override(lean_object*); @@ -794,6 +796,7 @@ lean_object* l_panic___at_Lean_Level_normalize___spec__1(lean_object*); static lean_object* l___private_Lean_Expr_0__Lean_reprExpr____x40_Lean_Expr___hyg_3663____closed__6; static lean_object* l_Lean_instForInFVarIdSetFVarId___closed__2; LEAN_EXPORT lean_object* l_Lean_mkEM(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_isSemiOutParam___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Expr_hasMVar___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_mkApp(lean_object*, lean_object*); static lean_object* l___private_Lean_Expr_0__Lean_reprBinderInfo____x40_Lean_Expr___hyg_497____closed__16; @@ -969,6 +972,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Expr_0__Lean_Expr_updateMData_x21Impl( LEAN_EXPORT lean_object* l_Lean_Expr_letBody_x21(lean_object*); LEAN_EXPORT lean_object* l_Lean_mkFreshFVarId(lean_object*); static lean_object* l_Lean_Expr_instBEqExpr___closed__1; +static lean_object* l_Lean_Expr_isSemiOutParam___closed__2; static lean_object* l_Lean_Expr_ctorName___closed__8; LEAN_EXPORT lean_object* l___private_Lean_Expr_0__Lean_Expr_updateForall_x21Impl___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Expr_etaExpandedStrict_x3f(lean_object*); @@ -26895,6 +26899,44 @@ x_3 = lean_box(x_2); return x_3; } } +static lean_object* _init_l_Lean_Expr_isSemiOutParam___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("semiOutParam", 12); +return x_1; +} +} +static lean_object* _init_l_Lean_Expr_isSemiOutParam___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Expr_isSemiOutParam___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT uint8_t l_Lean_Expr_isSemiOutParam(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; uint8_t x_4; +x_2 = l_Lean_Expr_isSemiOutParam___closed__2; +x_3 = lean_unsigned_to_nat(1u); +x_4 = l_Lean_Expr_isAppOfArity(x_1, x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Expr_isSemiOutParam___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l_Lean_Expr_isSemiOutParam(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} LEAN_EXPORT uint8_t l_Lean_Expr_isOptParam(lean_object* x_1) { _start: { @@ -26951,24 +26993,26 @@ lean_inc(x_1); x_4 = lean_is_out_param(x_1); if (x_4 == 0) { +uint8_t x_5; +x_5 = l_Lean_Expr_isSemiOutParam(x_1); +if (x_5 == 0) +{ return x_1; } else { -lean_object* x_5; -x_5 = l_Lean_Expr_appArg_x21(x_1); +lean_object* x_6; +x_6 = l_Lean_Expr_appArg_x21(x_1); lean_dec(x_1); -x_1 = x_5; +x_1 = x_6; goto _start; } } else { -lean_object* x_7; lean_object* x_8; -x_7 = l_Lean_Expr_appFn_x21(x_1); +lean_object* x_8; +x_8 = l_Lean_Expr_appArg_x21(x_1); lean_dec(x_1); -x_8 = l_Lean_Expr_appArg_x21(x_7); -lean_dec(x_7); x_1 = x_8; goto _start; } @@ -26984,6 +27028,17 @@ x_1 = x_11; goto _start; } } +else +{ +lean_object* x_13; lean_object* x_14; +x_13 = l_Lean_Expr_appFn_x21(x_1); +lean_dec(x_1); +x_14 = l_Lean_Expr_appArg_x21(x_13); +lean_dec(x_13); +x_1 = x_14; +goto _start; +} +} } LEAN_EXPORT lean_object* l_Lean_Expr_cleanupAnnotations(lean_object* x_1) { _start: @@ -27371,7 +27426,7 @@ static lean_object* _init_l___private_Lean_Expr_0__Lean_Expr_updateApp_x21Impl__ 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_Expr_mkData___closed__5; x_2 = l___private_Lean_Expr_0__Lean_Expr_updateApp_x21Impl___closed__1; -x_3 = lean_unsigned_to_nat(1465u); +x_3 = lean_unsigned_to_nat(1469u); x_4 = lean_unsigned_to_nat(18u); x_5 = l_Lean_Expr_appFn_x21___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -27450,7 +27505,7 @@ static lean_object* _init_l_Lean_Expr_updateFVar_x21___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_Expr_mkData___closed__5; x_2 = l_Lean_Expr_updateFVar_x21___closed__1; -x_3 = lean_unsigned_to_nat(1476u); +x_3 = lean_unsigned_to_nat(1480u); x_4 = lean_unsigned_to_nat(20u); x_5 = l_Lean_Expr_fvarId_x21___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -27511,7 +27566,7 @@ static lean_object* _init_l___private_Lean_Expr_0__Lean_Expr_updateConst_x21Impl 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_Expr_mkData___closed__5; x_2 = l___private_Lean_Expr_0__Lean_Expr_updateConst_x21Impl___closed__1; -x_3 = lean_unsigned_to_nat(1481u); +x_3 = lean_unsigned_to_nat(1485u); x_4 = lean_unsigned_to_nat(18u); x_5 = l_Lean_Expr_constName_x21___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -27577,7 +27632,7 @@ static lean_object* _init_l___private_Lean_Expr_0__Lean_Expr_updateSort_x21Impl_ 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_Expr_mkData___closed__5; x_2 = l___private_Lean_Expr_0__Lean_Expr_updateSort_x21Impl___closed__1; -x_3 = lean_unsigned_to_nat(1492u); +x_3 = lean_unsigned_to_nat(1496u); x_4 = lean_unsigned_to_nat(14u); x_5 = l___private_Lean_Expr_0__Lean_Expr_updateSort_x21Impl___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -27648,7 +27703,7 @@ static lean_object* _init_l___private_Lean_Expr_0__Lean_Expr_updateMData_x21Impl 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_Expr_mkData___closed__5; x_2 = l___private_Lean_Expr_0__Lean_Expr_updateMData_x21Impl___closed__1; -x_3 = lean_unsigned_to_nat(1503u); +x_3 = lean_unsigned_to_nat(1507u); x_4 = lean_unsigned_to_nat(17u); x_5 = l___private_Lean_Expr_0__Lean_Expr_updateMData_x21Impl___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -27716,7 +27771,7 @@ static lean_object* _init_l___private_Lean_Expr_0__Lean_Expr_updateProj_x21Impl_ 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_Expr_mkData___closed__5; x_2 = l___private_Lean_Expr_0__Lean_Expr_updateProj_x21Impl___closed__1; -x_3 = lean_unsigned_to_nat(1514u); +x_3 = lean_unsigned_to_nat(1518u); x_4 = lean_unsigned_to_nat(18u); x_5 = l___private_Lean_Expr_0__Lean_Expr_updateProj_x21Impl___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -27787,7 +27842,7 @@ static lean_object* _init_l___private_Lean_Expr_0__Lean_Expr_updateForall_x21Imp 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_Expr_mkData___closed__5; x_2 = l___private_Lean_Expr_0__Lean_Expr_updateForall_x21Impl___closed__1; -x_3 = lean_unsigned_to_nat(1529u); +x_3 = lean_unsigned_to_nat(1533u); x_4 = lean_unsigned_to_nat(23u); x_5 = l___private_Lean_Expr_0__Lean_Expr_updateForall_x21Impl___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -27890,7 +27945,7 @@ static lean_object* _init_l_Lean_Expr_updateForallE_x21___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_Expr_mkData___closed__5; x_2 = l_Lean_Expr_updateForallE_x21___closed__1; -x_3 = lean_unsigned_to_nat(1540u); +x_3 = lean_unsigned_to_nat(1544u); x_4 = lean_unsigned_to_nat(24u); x_5 = l___private_Lean_Expr_0__Lean_Expr_updateForall_x21Impl___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -27996,7 +28051,7 @@ static lean_object* _init_l___private_Lean_Expr_0__Lean_Expr_updateLambda_x21Imp 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_Expr_mkData___closed__5; x_2 = l___private_Lean_Expr_0__Lean_Expr_updateLambda_x21Impl___closed__1; -x_3 = lean_unsigned_to_nat(1549u); +x_3 = lean_unsigned_to_nat(1553u); x_4 = lean_unsigned_to_nat(19u); x_5 = l___private_Lean_Expr_0__Lean_Expr_updateLambda_x21Impl___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -28099,7 +28154,7 @@ static lean_object* _init_l_Lean_Expr_updateLambdaE_x21___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_Expr_mkData___closed__5; x_2 = l_Lean_Expr_updateLambdaE_x21___closed__1; -x_3 = lean_unsigned_to_nat(1560u); +x_3 = lean_unsigned_to_nat(1564u); x_4 = lean_unsigned_to_nat(20u); x_5 = l___private_Lean_Expr_0__Lean_Expr_updateLambda_x21Impl___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -28197,7 +28252,7 @@ static lean_object* _init_l___private_Lean_Expr_0__Lean_Expr_updateLet_x21Impl__ 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_Expr_mkData___closed__5; x_2 = l___private_Lean_Expr_0__Lean_Expr_updateLet_x21Impl___closed__1; -x_3 = lean_unsigned_to_nat(1569u); +x_3 = lean_unsigned_to_nat(1573u); x_4 = lean_unsigned_to_nat(22u); x_5 = l_Lean_Expr_letName_x21___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -30297,6 +30352,10 @@ l_Lean_Expr_isOutParam___closed__1 = _init_l_Lean_Expr_isOutParam___closed__1(); lean_mark_persistent(l_Lean_Expr_isOutParam___closed__1); l_Lean_Expr_isOutParam___closed__2 = _init_l_Lean_Expr_isOutParam___closed__2(); lean_mark_persistent(l_Lean_Expr_isOutParam___closed__2); +l_Lean_Expr_isSemiOutParam___closed__1 = _init_l_Lean_Expr_isSemiOutParam___closed__1(); +lean_mark_persistent(l_Lean_Expr_isSemiOutParam___closed__1); +l_Lean_Expr_isSemiOutParam___closed__2 = _init_l_Lean_Expr_isSemiOutParam___closed__2(); +lean_mark_persistent(l_Lean_Expr_isSemiOutParam___closed__2); l___private_Lean_Expr_0__Lean_Expr_updateApp_x21Impl___closed__1 = _init_l___private_Lean_Expr_0__Lean_Expr_updateApp_x21Impl___closed__1(); lean_mark_persistent(l___private_Lean_Expr_0__Lean_Expr_updateApp_x21Impl___closed__1); l___private_Lean_Expr_0__Lean_Expr_updateApp_x21Impl___closed__2 = _init_l___private_Lean_Expr_0__Lean_Expr_updateApp_x21Impl___closed__2(); diff --git a/stage0/stdlib/Lean/Linter/Builtin.c b/stage0/stdlib/Lean/Linter/Builtin.c index 5f3a69f13873..42e478191de8 100644 --- a/stage0/stdlib/Lean/Linter/Builtin.c +++ b/stage0/stdlib/Lean/Linter/Builtin.c @@ -13,129 +13,133 @@ #ifdef __cplusplus extern "C" { #endif +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__7; +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___boxed(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_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__7; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___lambda__1___closed__2; +static lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__5___closed__5; +static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___closed__2; uint8_t l_Lean_Syntax_matchesIdent(lean_object*, lean_object*); -static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__14; lean_object* l_Lean_logAt___at_Lean_Elab_Command_elabCommand___spec__4(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__1___closed__3; -static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__2___closed__2; +static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__7; 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_sequenceMap_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__5___closed__1; +static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__18; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_Builtin___hyg_6____closed__3; -static lean_object* l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__3___closed__5; +LEAN_EXPORT lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_Builtin___hyg_2485_(lean_object*); +static lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__5___closed__1; lean_object* lean_array_push(lean_object*, lean_object*); -static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_Builtin___hyg_2470____closed__1; -static lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__5___closed__4; -static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__3; +static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__1___closed__1; uint8_t lean_usize_dec_eq(size_t, size_t); -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__2; +static lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__5___closed__4; +LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getArgs(lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___lambda__1___closed__2; -static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__5___closed__3; +LEAN_EXPORT lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__2(lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); -static lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__5___closed__2; -LEAN_EXPORT lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__3___closed__2; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__5; uint8_t l_Lean_Linter_getLinterValue(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_linter_suspiciousUnexpanderPatterns; uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); -static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__17; lean_object* l_Lean_stringToMessageData(lean_object*); -static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__12; -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__8; -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___lambda__1___closed__1; lean_object* l_List_head_x21___rarg(lean_object*, lean_object*); -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__5(lean_object*, size_t, size_t); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__3; +static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__17; lean_object* l_Lean_Option_register___at_Lean_Elab_initFn____x40_Lean_Elab_AutoBound___hyg_7____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__2___closed__1; -LEAN_EXPORT lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__1___closed__2; +static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__4; +LEAN_EXPORT lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns; size_t lean_usize_of_nat(lean_object*); -LEAN_EXPORT lean_object* l_Array_sequenceMap___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__1___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__3; -LEAN_EXPORT lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__1(lean_object*); -static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__10; +static lean_object* l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__3___closed__4; +LEAN_EXPORT lean_object* l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__4; +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__10; +static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__5___closed__1; +LEAN_EXPORT lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__1(lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_Builtin___hyg_6____closed__1; -static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__7; -LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__6; +LEAN_EXPORT lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__5___closed__3; +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__5___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_sequenceMap___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__2(lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__4; -static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__6; -static lean_object* l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__3___closed__1; +static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__1___closed__3; +static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__5___closed__6; +static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__8; +static lean_object* l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__3___closed__1; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__1; +static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__1___closed__4; +static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__1___closed__2; LEAN_EXPORT uint8_t l_Lean_Linter_getLinterSuspiciousUnexpanderPatterns(lean_object*); -static lean_object* l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__3___closed__4; -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__6; -static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__15; +static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__15; +static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__11; +static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__2___closed__2; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__5(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_Linter_suspiciousUnexpanderPatterns___closed__3; uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__5; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__8; +static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__1; +static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__9; +static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__5; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_Builtin___hyg_6____closed__4; -static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__18; -static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__5___closed__5; +static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__3; LEAN_EXPORT lean_object* l_Lean_Linter_getLinterSuspiciousUnexpanderPatterns___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_Builtin___hyg_6_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__3___closed__3; lean_object* l_Lean_Elab_Command_addLinter(lean_object*, lean_object*); static lean_object* l_Lean_Linter_getLinterSuspiciousUnexpanderPatterns___closed__1; -LEAN_EXPORT lean_object* l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__3___closed__2; -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___lambda__1___closed__3; -static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__8; -static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__5___closed__6; +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___lambda__1___closed__1; +LEAN_EXPORT lean_object* l_Array_sequenceMap___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___closed__1; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_Builtin___hyg_6____closed__8; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_Builtin___hyg_6____closed__5; -static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__5; +static lean_object* l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__3___closed__5; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -static lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__5___closed__5; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); -static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__11; +LEAN_EXPORT lean_object* l_Array_sequenceMap___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__1___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__5___closed__2; +static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__13; +static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__5___closed__4; uint8_t l_Lean_Syntax_isNone(lean_object*); -static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__9; +LEAN_EXPORT lean_object* l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); -static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__5___closed__1; -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__5___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__3___closed__3; -static lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__5___closed__3; +static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__5___closed__3; +static lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__5___closed__2; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_Builtin___hyg_6____closed__2; +static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__19; extern lean_object* l_Lean_Elab_Command_instInhabitedScope; -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___lambda__1___closed__4; size_t lean_usize_add(size_t, size_t); -static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__4; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_Builtin___hyg_6____closed__9; +static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__2___closed__1; lean_object* lean_array_uget(lean_object*, size_t); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___lambda__1___closed__4; +static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__2; +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__5(lean_object*, size_t, size_t); lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__1; +static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__5___closed__5; +static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__14; +static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__6; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__2; +static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__12; 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_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__2; lean_object* lean_array_get_size(lean_object*); -static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__13; -static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__16; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_Builtin___hyg_6____closed__6; +LEAN_EXPORT lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__9; lean_object* lean_nat_add(lean_object*, lean_object*); -static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__5___closed__2; +LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_Builtin___hyg_6____closed__7; -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__9; -static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__1___closed__1; lean_object* l_Lean_MessageData_ofName(lean_object*); -static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__1___closed__4; -LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_Builtin___hyg_2470_(lean_object*); -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__1; -static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__19; -static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__5___closed__4; +static lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__16; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___lambda__1___closed__3; static lean_object* _init_l_Lean_Linter_initFn____x40_Lean_Linter_Builtin___hyg_6____closed__1() { _start: { @@ -259,7 +263,7 @@ x_3 = lean_box(x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___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_Array_sequenceMap_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__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; @@ -328,18 +332,18 @@ return x_20; } } } -LEAN_EXPORT lean_object* l_Array_sequenceMap___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Array_sequenceMap___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__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_Linter_suspiciousUnexpanderPatterns___spec__2(x_1, x_2, x_3, x_5, x_4); +x_6 = l_Array_sequenceMap_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__2(x_1, x_2, x_3, x_5, x_4); return x_6; } } -static lean_object* _init_l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__3___closed__1() { +static lean_object* _init_l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__3___closed__1() { _start: { lean_object* x_1; lean_object* x_2; @@ -348,7 +352,7 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__3___closed__2() { +static lean_object* _init_l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__3___closed__2() { _start: { lean_object* x_1; @@ -356,16 +360,16 @@ x_1 = lean_mk_string_from_bytes(" [", 2); return x_1; } } -static lean_object* _init_l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__3___closed__3() { +static lean_object* _init_l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__3___closed__3() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__3___closed__2; +x_1 = l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__3___closed__2; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__3___closed__4() { +static lean_object* _init_l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__3___closed__4() { _start: { lean_object* x_1; @@ -373,27 +377,27 @@ x_1 = lean_mk_string_from_bytes("]", 1); return x_1; } } -static lean_object* _init_l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__3___closed__5() { +static lean_object* _init_l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__3___closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__3___closed__4; +x_1 = l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__3___closed__4; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___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_EXPORT lean_object* l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__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) { _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; uint8_t x_17; lean_object* x_18; x_7 = lean_ctor_get(x_1, 0); lean_inc(x_7); lean_dec(x_1); -x_8 = l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__3___closed__1; +x_8 = l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__3___closed__1; x_9 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_9, 0, x_8); lean_ctor_set(x_9, 1, x_3); -x_10 = l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__3___closed__3; +x_10 = l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__3___closed__3; x_11 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_11, 0, x_9); lean_ctor_set(x_11, 1, x_10); @@ -402,7 +406,7 @@ x_12 = l_Lean_MessageData_ofName(x_7); x_13 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_13, 0, x_11); lean_ctor_set(x_13, 1, x_12); -x_14 = l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__3___closed__5; +x_14 = l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__3___closed__5; x_15 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_15, 0, x_13); lean_ctor_set(x_15, 1, x_14); @@ -414,7 +418,7 @@ x_18 = l_Lean_logAt___at_Lean_Elab_Command_elabCommand___spec__4(x_2, x_16, x_17 return x_18; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___lambda__1___closed__1() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___lambda__1___closed__1() { _start: { lean_object* x_1; @@ -422,27 +426,27 @@ x_1 = lean_mk_string_from_bytes("Unexpanders should match the function name agai return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___lambda__1___closed__2() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___lambda__1___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___lambda__1___closed__1; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___lambda__1___closed__1; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___lambda__1___closed__3() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___lambda__1___closed__3() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___lambda__1___closed__2; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___lambda__1___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_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___lambda__1___closed__4() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___lambda__1___closed__4() { _start: { lean_object* x_1; lean_object* x_2; @@ -452,20 +456,20 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___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; uint8_t x_8; x_5 = l_Lean_Linter_getLinterSuspiciousUnexpanderPatterns___closed__1; -x_6 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___lambda__1___closed__3; -x_7 = l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__3(x_5, x_1, x_6, x_2, x_3, x_4); +x_6 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___lambda__1___closed__3; +x_7 = l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__3(x_5, x_1, x_6, x_2, x_3, x_4); x_8 = !lean_is_exclusive(x_7); if (x_8 == 0) { lean_object* x_9; lean_object* x_10; x_9 = lean_ctor_get(x_7, 0); lean_dec(x_9); -x_10 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___lambda__1___closed__4; +x_10 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___lambda__1___closed__4; lean_ctor_set(x_7, 0, x_10); return x_7; } @@ -475,7 +479,7 @@ lean_object* x_11; lean_object* x_12; lean_object* x_13; x_11 = lean_ctor_get(x_7, 1); lean_inc(x_11); lean_dec(x_7); -x_12 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___lambda__1___closed__4; +x_12 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___lambda__1___closed__4; x_13 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_13, 0, x_12); lean_ctor_set(x_13, 1, x_11); @@ -483,15 +487,15 @@ return x_13; } } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__1() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___lambda__1___boxed), 4, 0); +x_1 = lean_alloc_closure((void*)(l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___lambda__1___boxed), 4, 0); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__2() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__2() { _start: { lean_object* x_1; @@ -499,7 +503,7 @@ x_1 = lean_mk_string_from_bytes("Parser", 6); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__3() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__3() { _start: { lean_object* x_1; @@ -507,7 +511,7 @@ x_1 = lean_mk_string_from_bytes("Term", 4); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__4() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__4() { _start: { lean_object* x_1; @@ -515,19 +519,19 @@ x_1 = lean_mk_string_from_bytes("quot", 4); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__5() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__5() { _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_Linter_initFn____x40_Lean_Linter_Builtin___hyg_6____closed__7; -x_2 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__2; -x_3 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__3; -x_4 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__4; +x_2 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__2; +x_3 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__3; +x_4 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__4; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__6() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__6() { _start: { lean_object* x_1; @@ -535,19 +539,19 @@ x_1 = lean_mk_string_from_bytes("app", 3); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__7() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__7() { _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_Linter_initFn____x40_Lean_Linter_Builtin___hyg_6____closed__7; -x_2 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__2; -x_3 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__3; -x_4 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__6; +x_2 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__2; +x_3 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__3; +x_4 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__6; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__8() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__8() { _start: { lean_object* x_1; @@ -555,17 +559,17 @@ x_1 = lean_mk_string_from_bytes("ident", 5); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__9() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___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_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__8; +x_2 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__8; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___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_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__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_object* x_7) { _start: { uint8_t x_8; @@ -585,8 +589,8 @@ else lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_dec(x_4); x_10 = lean_array_uget(x_1, x_3); -x_11 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__1; -x_12 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__5; +x_11 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__1; +x_12 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__5; lean_inc(x_10); x_13 = l_Lean_Syntax_isOfKind(x_10, x_12); if (x_13 == 0) @@ -606,13 +610,13 @@ 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_10, x_18); lean_dec(x_10); -x_20 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__7; +x_20 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__7; lean_inc(x_19); x_21 = l_Lean_Syntax_isOfKind(x_19, x_20); if (x_21 == 0) { lean_object* x_22; uint8_t x_23; -x_22 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__9; +x_22 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__9; lean_inc(x_19); x_23 = l_Lean_Syntax_isOfKind(x_19, x_22); if (x_23 == 0) @@ -718,7 +722,7 @@ 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_19, x_45); lean_dec(x_19); -x_47 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__9; +x_47 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__9; lean_inc(x_46); x_48 = l_Lean_Syntax_isOfKind(x_46, x_47); if (x_48 == 0) @@ -822,7 +826,7 @@ return x_69; } } } -static lean_object* _init_l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__5___closed__1() { +static lean_object* _init_l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__5___closed__1() { _start: { lean_object* x_1; @@ -830,7 +834,7 @@ x_1 = lean_mk_string_from_bytes("Attr", 4); return x_1; } } -static lean_object* _init_l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__5___closed__2() { +static lean_object* _init_l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__5___closed__2() { _start: { lean_object* x_1; @@ -838,19 +842,19 @@ x_1 = lean_mk_string_from_bytes("simple", 6); return x_1; } } -static lean_object* _init_l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__5___closed__3() { +static lean_object* _init_l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__5___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_Linter_initFn____x40_Lean_Linter_Builtin___hyg_6____closed__7; -x_2 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__2; -x_3 = l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__5___closed__1; -x_4 = l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__5___closed__2; +x_2 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__2; +x_3 = l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__5___closed__1; +x_4 = l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__5___closed__2; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__5___closed__4() { +static lean_object* _init_l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__5___closed__4() { _start: { lean_object* x_1; @@ -858,17 +862,17 @@ x_1 = lean_mk_string_from_bytes("app_unexpander", 14); return x_1; } } -static lean_object* _init_l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__5___closed__5() { +static lean_object* _init_l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__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_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__5___closed__4; +x_2 = l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__5___closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__5(lean_object* x_1, size_t x_2, size_t x_3) { +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__5(lean_object* x_1, size_t x_2, size_t x_3) { _start: { uint8_t x_4; @@ -877,7 +881,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; uint8_t x_7; x_5 = lean_array_uget(x_1, x_2); -x_6 = l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__5___closed__3; +x_6 = l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__5___closed__3; lean_inc(x_5); x_7 = l_Lean_Syntax_isOfKind(x_5, x_6); if (x_7 == 0) @@ -894,7 +898,7 @@ else lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; x_11 = lean_unsigned_to_nat(0u); x_12 = l_Lean_Syntax_getArg(x_5, x_11); -x_13 = l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__5___closed__5; +x_13 = l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__5___closed__5; x_14 = l_Lean_Syntax_matchesIdent(x_12, x_13); lean_dec(x_12); if (x_14 == 0) @@ -938,7 +942,7 @@ return x_25; } } } -static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__1___closed__1() { +static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__1___closed__1() { _start: { lean_object* x_1; @@ -946,19 +950,19 @@ x_1 = lean_mk_string_from_bytes("attrInstance", 12); return x_1; } } -static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__1___closed__2() { +static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__1___closed__2() { _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_Linter_initFn____x40_Lean_Linter_Builtin___hyg_6____closed__7; -x_2 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__2; -x_3 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__3; -x_4 = l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__1___closed__1; +x_2 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__2; +x_3 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__3; +x_4 = l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__1___closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__1___closed__3() { +static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__1___closed__3() { _start: { lean_object* x_1; @@ -966,23 +970,23 @@ x_1 = lean_mk_string_from_bytes("attrKind", 8); return x_1; } } -static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__1___closed__4() { +static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__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_Linter_initFn____x40_Lean_Linter_Builtin___hyg_6____closed__7; -x_2 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__2; -x_3 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__3; -x_4 = l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__1___closed__3; +x_2 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__2; +x_3 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__3; +x_4 = l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__1___closed__3; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__1(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; -x_2 = l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__1___closed__2; +x_2 = l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__1___closed__2; lean_inc(x_1); x_3 = l_Lean_Syntax_isOfKind(x_1, x_2); if (x_3 == 0) @@ -997,7 +1001,7 @@ else 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); -x_7 = l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__1___closed__4; +x_7 = l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__1___closed__4; lean_inc(x_6); x_8 = l_Lean_Syntax_isOfKind(x_6, x_7); if (x_8 == 0) @@ -1035,7 +1039,7 @@ return x_15; } } } -static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__2___closed__1() { +static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__2___closed__1() { _start: { lean_object* x_1; @@ -1043,23 +1047,23 @@ x_1 = lean_mk_string_from_bytes("matchAlt", 8); return x_1; } } -static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__2___closed__2() { +static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__2___closed__2() { _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_Linter_initFn____x40_Lean_Linter_Builtin___hyg_6____closed__7; -x_2 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__2; -x_3 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__3; -x_4 = l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__2___closed__1; +x_2 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__2; +x_3 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__3; +x_4 = l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__2___closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__2(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__2(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; -x_2 = l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__2___closed__2; +x_2 = l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__2___closed__2; lean_inc(x_1); x_3 = l_Lean_Syntax_isOfKind(x_1, x_2); if (x_3 == 0) @@ -1112,7 +1116,7 @@ return x_14; } } } -LEAN_EXPORT lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__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_Linter_suspiciousUnexpanderPatterns___elambda__1___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; size_t x_7; size_t x_8; lean_object* x_9; lean_object* x_10; @@ -1121,7 +1125,7 @@ x_7 = lean_usize_of_nat(x_6); lean_dec(x_6); x_8 = 0; x_9 = lean_box(0); -x_10 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4(x_1, x_7, x_8, x_9, x_3, x_4, x_5); +x_10 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4(x_1, x_7, x_8, x_9, x_3, x_4, x_5); if (lean_obj_tag(x_10) == 0) { uint8_t x_11; @@ -1170,7 +1174,7 @@ return x_18; } } } -static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__1() { +static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__1() { _start: { lean_object* x_1; @@ -1178,19 +1182,19 @@ x_1 = lean_mk_string_from_bytes("attributes", 10); return x_1; } } -static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__2() { +static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__2() { _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_Linter_initFn____x40_Lean_Linter_Builtin___hyg_6____closed__7; -x_2 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__2; -x_3 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__3; -x_4 = l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__1; +x_2 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__2; +x_3 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__3; +x_4 = l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__3() { +static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__3() { _start: { lean_object* x_1; lean_object* x_2; @@ -1199,12 +1203,12 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__4() { +static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__4() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = 1; -x_2 = l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__3; +x_2 = l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__3; x_3 = lean_box(x_1); x_4 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_4, 0, x_3); @@ -1212,15 +1216,15 @@ lean_ctor_set(x_4, 1, x_2); return x_4; } } -static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__5() { +static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__5() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__1), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__1), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__6() { +static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__6() { _start: { lean_object* x_1; @@ -1228,7 +1232,7 @@ x_1 = lean_mk_string_from_bytes("Command", 7); return x_1; } } -static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__7() { +static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__7() { _start: { lean_object* x_1; @@ -1236,19 +1240,19 @@ x_1 = lean_mk_string_from_bytes("def", 3); return x_1; } } -static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__8() { +static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__8() { _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_Linter_initFn____x40_Lean_Linter_Builtin___hyg_6____closed__7; -x_2 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__2; -x_3 = l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__6; -x_4 = l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__7; +x_2 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__2; +x_3 = l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__6; +x_4 = l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__7; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__9() { +static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__9() { _start: { lean_object* x_1; @@ -1256,19 +1260,19 @@ x_1 = lean_mk_string_from_bytes("optDeclSig", 10); return x_1; } } -static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__10() { +static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__10() { _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_Linter_initFn____x40_Lean_Linter_Builtin___hyg_6____closed__7; -x_2 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__2; -x_3 = l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__6; -x_4 = l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__9; +x_2 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__2; +x_3 = l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__6; +x_4 = l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__9; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__11() { +static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__11() { _start: { lean_object* x_1; @@ -1276,19 +1280,19 @@ x_1 = lean_mk_string_from_bytes("typeSpec", 8); return x_1; } } -static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__12() { +static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__12() { _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_Linter_initFn____x40_Lean_Linter_Builtin___hyg_6____closed__7; -x_2 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__2; -x_3 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__3; -x_4 = l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__11; +x_2 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__2; +x_3 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__3; +x_4 = l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__11; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__13() { +static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__13() { _start: { lean_object* x_1; @@ -1296,19 +1300,19 @@ x_1 = lean_mk_string_from_bytes("declValEqns", 11); return x_1; } } -static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__14() { +static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__14() { _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_Linter_initFn____x40_Lean_Linter_Builtin___hyg_6____closed__7; -x_2 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__2; -x_3 = l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__6; -x_4 = l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__13; +x_2 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__2; +x_3 = l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__6; +x_4 = l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__13; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__15() { +static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__15() { _start: { lean_object* x_1; @@ -1316,19 +1320,19 @@ x_1 = lean_mk_string_from_bytes("matchAltsWhereDecls", 19); return x_1; } } -static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__16() { +static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__16() { _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_Linter_initFn____x40_Lean_Linter_Builtin___hyg_6____closed__7; -x_2 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__2; -x_3 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__3; -x_4 = l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__15; +x_2 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__2; +x_3 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__3; +x_4 = l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__15; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__17() { +static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__17() { _start: { lean_object* x_1; @@ -1336,27 +1340,27 @@ x_1 = lean_mk_string_from_bytes("matchAlts", 9); return x_1; } } -static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__18() { +static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__18() { _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_Linter_initFn____x40_Lean_Linter_Builtin___hyg_6____closed__7; -x_2 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__2; -x_3 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__3; -x_4 = l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__17; +x_2 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__2; +x_3 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__3; +x_4 = l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__17; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__19() { +static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__19() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__2), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__2), 1, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___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_EXPORT lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___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) { _start: { lean_object* x_7; lean_object* x_8; uint8_t x_9; @@ -1385,7 +1389,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); lean_dec(x_8); -x_14 = l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__2; +x_14 = l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__2; lean_inc(x_13); x_15 = l_Lean_Syntax_isOfKind(x_13, x_14); if (x_15 == 0) @@ -1416,7 +1420,7 @@ if (x_21 == 0) lean_object* x_121; lean_dec(x_20); lean_dec(x_19); -x_121 = l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__3; +x_121 = l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__3; x_22 = x_121; goto block_120; } @@ -1429,7 +1433,7 @@ if (x_122 == 0) lean_object* x_123; lean_dec(x_20); lean_dec(x_19); -x_123 = l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__3; +x_123 = l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__3; x_22 = x_123; goto block_120; } @@ -1439,7 +1443,7 @@ size_t x_124; size_t x_125; lean_object* x_126; lean_object* x_127; lean_object* x_124 = 0; x_125 = lean_usize_of_nat(x_20); lean_dec(x_20); -x_126 = l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__4; +x_126 = l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__4; x_127 = l_Array_foldlMUnsafe_fold___at_Lean_Syntax_SepArray_getElems___spec__1(x_19, x_124, x_125, x_126); lean_dec(x_19); x_128 = lean_ctor_get(x_127, 1); @@ -1452,8 +1456,8 @@ goto block_120; block_120: { lean_object* x_23; lean_object* x_24; -x_23 = l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__5; -x_24 = l_Array_sequenceMap___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__1(x_22, x_23); +x_23 = l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__5; +x_24 = l_Array_sequenceMap___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__1(x_22, x_23); lean_dec(x_22); if (lean_obj_tag(x_24) == 0) { @@ -1536,7 +1540,7 @@ else lean_object* x_43; lean_object* x_44; uint8_t x_45; x_43 = l_Lean_Syntax_getArg(x_2, x_7); lean_dec(x_2); -x_44 = l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__8; +x_44 = l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__8; lean_inc(x_43); x_45 = l_Lean_Syntax_isOfKind(x_43, x_44); if (x_45 == 0) @@ -1557,7 +1561,7 @@ else lean_object* x_48; lean_object* x_49; lean_object* x_50; uint8_t x_51; x_48 = lean_unsigned_to_nat(2u); x_49 = l_Lean_Syntax_getArg(x_43, x_48); -x_50 = l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__10; +x_50 = l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__10; lean_inc(x_49); x_51 = l_Lean_Syntax_isOfKind(x_49, x_50); if (x_51 == 0) @@ -1619,7 +1623,7 @@ else lean_object* x_62; lean_object* x_63; uint8_t x_64; x_62 = l_Lean_Syntax_getArg(x_58, x_12); lean_dec(x_58); -x_63 = l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__12; +x_63 = l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__12; x_64 = l_Lean_Syntax_isOfKind(x_62, x_63); if (x_64 == 0) { @@ -1638,7 +1642,7 @@ else { lean_object* x_67; lean_object* x_68; uint8_t x_69; x_67 = l_Lean_Syntax_getArg(x_43, x_28); -x_68 = l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__14; +x_68 = l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__14; lean_inc(x_67); x_69 = l_Lean_Syntax_isOfKind(x_67, x_68); if (x_69 == 0) @@ -1660,7 +1664,7 @@ else lean_object* x_72; lean_object* x_73; uint8_t x_74; x_72 = l_Lean_Syntax_getArg(x_67, x_12); lean_dec(x_67); -x_73 = l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__16; +x_73 = l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__16; lean_inc(x_72); x_74 = l_Lean_Syntax_isOfKind(x_72, x_73); if (x_74 == 0) @@ -1681,7 +1685,7 @@ else { lean_object* x_77; lean_object* x_78; uint8_t x_79; x_77 = l_Lean_Syntax_getArg(x_72, x_12); -x_78 = l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__18; +x_78 = l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__18; lean_inc(x_77); x_79 = l_Lean_Syntax_isOfKind(x_77, x_78); if (x_79 == 0) @@ -1706,7 +1710,7 @@ x_82 = l_Lean_Syntax_getArg(x_77, x_12); lean_dec(x_77); x_83 = l_Lean_Syntax_getArgs(x_82); lean_dec(x_82); -x_84 = l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__19; +x_84 = l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__19; x_85 = l_Array_sequenceMap___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__1(x_83, x_84); lean_dec(x_83); if (lean_obj_tag(x_85) == 0) @@ -1847,7 +1851,7 @@ size_t x_113; size_t x_114; uint8_t x_115; x_113 = 0; x_114 = lean_usize_of_nat(x_106); lean_dec(x_106); -x_115 = l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__5(x_27, x_113, x_114); +x_115 = l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__5(x_27, x_113, x_114); lean_dec(x_27); if (x_115 == 0) { @@ -1865,7 +1869,7 @@ else { lean_object* x_118; lean_object* x_119; x_118 = lean_box(0); -x_119 = l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__3(x_88, x_118, x_4, x_5, x_6); +x_119 = l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__3(x_88, x_118, x_4, x_5, x_6); lean_dec(x_88); return x_119; } @@ -1893,7 +1897,7 @@ return x_119; } } } -static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__5___closed__1() { +static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__5___closed__1() { _start: { lean_object* x_1; @@ -1901,19 +1905,19 @@ x_1 = lean_mk_string_from_bytes("declaration", 11); return x_1; } } -static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__5___closed__2() { +static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__5___closed__2() { _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_Linter_initFn____x40_Lean_Linter_Builtin___hyg_6____closed__7; -x_2 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__2; -x_3 = l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__6; -x_4 = l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__5___closed__1; +x_2 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__2; +x_3 = l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__6; +x_4 = l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__5___closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__5___closed__3() { +static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__5___closed__3() { _start: { lean_object* x_1; @@ -1921,19 +1925,19 @@ x_1 = lean_mk_string_from_bytes("declModifiers", 13); return x_1; } } -static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__5___closed__4() { +static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__5___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_Linter_initFn____x40_Lean_Linter_Builtin___hyg_6____closed__7; -x_2 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__2; -x_3 = l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__6; -x_4 = l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__5___closed__3; +x_2 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__2; +x_3 = l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__6; +x_4 = l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__5___closed__3; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__5___closed__5() { +static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__5___closed__5() { _start: { lean_object* x_1; @@ -1941,24 +1945,24 @@ x_1 = lean_mk_string_from_bytes("docComment", 10); return x_1; } } -static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__5___closed__6() { +static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__5___closed__6() { _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_Linter_initFn____x40_Lean_Linter_Builtin___hyg_6____closed__7; -x_2 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__2; -x_3 = l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__6; -x_4 = l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__5___closed__5; +x_2 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__2; +x_3 = l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__6; +x_4 = l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__5___closed__5; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__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_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__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; uint8_t x_7; lean_dec(x_2); -x_6 = l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__5___closed__2; +x_6 = l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__5___closed__2; lean_inc(x_1); x_7 = l_Lean_Syntax_isOfKind(x_1, x_6); if (x_7 == 0) @@ -1978,7 +1982,7 @@ 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_1, x_10); -x_12 = l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__5___closed__4; +x_12 = l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__5___closed__4; lean_inc(x_11); x_13 = l_Lean_Syntax_isOfKind(x_11, x_12); if (x_13 == 0) @@ -2024,7 +2028,7 @@ else lean_object* x_22; lean_object* x_23; uint8_t x_24; x_22 = l_Lean_Syntax_getArg(x_16, x_10); lean_dec(x_16); -x_23 = l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__5___closed__6; +x_23 = l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__5___closed__6; x_24 = l_Lean_Syntax_isOfKind(x_22, x_23); if (x_24 == 0) { @@ -2043,7 +2047,7 @@ else { lean_object* x_27; lean_object* x_28; x_27 = lean_box(0); -x_28 = l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4(x_11, x_1, x_27, x_3, x_4, x_5); +x_28 = l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4(x_11, x_1, x_27, x_3, x_4, x_5); return x_28; } } @@ -2053,14 +2057,14 @@ else lean_object* x_29; lean_object* x_30; lean_dec(x_16); x_29 = lean_box(0); -x_30 = l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4(x_11, x_1, x_29, x_3, x_4, x_5); +x_30 = l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4(x_11, x_1, x_29, x_3, x_4, x_5); return x_30; } } } } } -LEAN_EXPORT lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__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; @@ -2098,7 +2102,7 @@ else lean_object* x_16; lean_object* x_17; lean_free_object(x_5); x_16 = lean_box(0); -x_17 = l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__5(x_1, x_16, x_2, x_3, x_8); +x_17 = l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__5(x_1, x_16, x_2, x_3, x_8); return x_17; } } @@ -2138,53 +2142,92 @@ else { lean_object* x_28; lean_object* x_29; x_28 = lean_box(0); -x_29 = l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__5(x_1, x_28, x_2, x_3, x_19); +x_29 = l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__5(x_1, x_28, x_2, x_3, x_19); return x_29; } } } } -LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__2___boxed(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_Linter_suspiciousUnexpanderPatterns___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Linter_initFn____x40_Lean_Linter_Builtin___hyg_6____closed__7; +x_2 = l_Lean_Linter_initFn____x40_Lean_Linter_Builtin___hyg_6____closed__8; +x_3 = l_Lean_Linter_initFn____x40_Lean_Linter_Builtin___hyg_6____closed__2; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1), 4, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Linter_suspiciousUnexpanderPatterns___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Linter_suspiciousUnexpanderPatterns___closed__2; +x_2 = l_Lean_Linter_suspiciousUnexpanderPatterns___closed__1; +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_Linter_suspiciousUnexpanderPatterns() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Linter_suspiciousUnexpanderPatterns___closed__3; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__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_Linter_suspiciousUnexpanderPatterns___spec__2(x_1, x_2, x_3, x_4, x_5); +x_6 = l_Array_sequenceMap_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__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_Linter_suspiciousUnexpanderPatterns___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Array_sequenceMap___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Array_sequenceMap___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__1(x_1, x_2); +x_3 = l_Array_sequenceMap___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__1(x_1, x_2); lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___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_EXPORT lean_object* l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__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) { _start: { lean_object* x_7; -x_7 = l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__3(x_1, x_2, x_3, x_4, x_5, x_6); +x_7 = l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__3(x_1, x_2, x_3, x_4, x_5, x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); return x_7; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___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_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___lambda__1(x_1, x_2, x_3, x_4); +x_5 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___lambda__1(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_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___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_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__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_object* x_7) { _start: { size_t x_8; size_t x_9; lean_object* x_10; @@ -2192,12 +2235,12 @@ 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_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4(x_1, x_8, x_9, x_4, x_5, x_6, x_7); +x_10 = l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4(x_1, x_8, x_9, x_4, x_5, x_6, x_7); lean_dec(x_1); return x_10; } } -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___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; @@ -2205,35 +2248,27 @@ 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_Linter_suspiciousUnexpanderPatterns___spec__5(x_1, x_4, x_5); +x_6 = l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__5(x_1, x_4, x_5); lean_dec(x_1); x_7 = lean_box(x_6); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__3___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_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__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_Lean_Linter_suspiciousUnexpanderPatterns___lambda__3(x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__3(x_1, x_2, x_3, x_4, x_5); lean_dec(x_2); lean_dec(x_1); return x_6; } } -static lean_object* _init_l_Lean_Linter_initFn____x40_Lean_Linter_Builtin___hyg_2470____closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Linter_suspiciousUnexpanderPatterns), 4, 0); -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_Builtin___hyg_2470_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_Builtin___hyg_2485_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Linter_initFn____x40_Lean_Linter_Builtin___hyg_2470____closed__1; +x_2 = l_Lean_Linter_suspiciousUnexpanderPatterns; x_3 = l_Lean_Elab_Command_addLinter(x_2, x_1); return x_3; } @@ -2280,117 +2315,123 @@ lean_mark_persistent(l_Lean_Linter_linter_suspiciousUnexpanderPatterns); lean_dec_ref(res); }l_Lean_Linter_getLinterSuspiciousUnexpanderPatterns___closed__1 = _init_l_Lean_Linter_getLinterSuspiciousUnexpanderPatterns___closed__1(); lean_mark_persistent(l_Lean_Linter_getLinterSuspiciousUnexpanderPatterns___closed__1); -l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__3___closed__1 = _init_l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__3___closed__1(); -lean_mark_persistent(l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__3___closed__1); -l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__3___closed__2 = _init_l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__3___closed__2(); -lean_mark_persistent(l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__3___closed__2); -l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__3___closed__3 = _init_l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__3___closed__3(); -lean_mark_persistent(l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__3___closed__3); -l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__3___closed__4 = _init_l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__3___closed__4(); -lean_mark_persistent(l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__3___closed__4); -l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__3___closed__5 = _init_l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__3___closed__5(); -lean_mark_persistent(l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__3___closed__5); -l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___lambda__1___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___lambda__1___closed__1(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___lambda__1___closed__1); -l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___lambda__1___closed__2 = _init_l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___lambda__1___closed__2(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___lambda__1___closed__2); -l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___lambda__1___closed__3 = _init_l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___lambda__1___closed__3(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___lambda__1___closed__3); -l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___lambda__1___closed__4 = _init_l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___lambda__1___closed__4(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___lambda__1___closed__4); -l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__1(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__1); -l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__2 = _init_l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__2(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__2); -l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__3 = _init_l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__3(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__3); -l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__4 = _init_l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__4(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__4); -l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__5 = _init_l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__5(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__5); -l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__6 = _init_l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__6(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__6); -l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__7 = _init_l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__7(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__7); -l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__8 = _init_l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__8(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__8); -l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__9 = _init_l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__9(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__4___closed__9); -l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__5___closed__1 = _init_l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__5___closed__1(); -lean_mark_persistent(l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__5___closed__1); -l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__5___closed__2 = _init_l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__5___closed__2(); -lean_mark_persistent(l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__5___closed__2); -l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__5___closed__3 = _init_l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__5___closed__3(); -lean_mark_persistent(l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__5___closed__3); -l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__5___closed__4 = _init_l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__5___closed__4(); -lean_mark_persistent(l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__5___closed__4); -l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__5___closed__5 = _init_l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__5___closed__5(); -lean_mark_persistent(l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___spec__5___closed__5); -l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__1___closed__1 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__1___closed__1(); -lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__1___closed__1); -l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__1___closed__2 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__1___closed__2(); -lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__1___closed__2); -l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__1___closed__3 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__1___closed__3(); -lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__1___closed__3); -l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__1___closed__4 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__1___closed__4(); -lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__1___closed__4); -l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__2___closed__1 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__2___closed__1(); -lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__2___closed__1); -l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__2___closed__2 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__2___closed__2(); -lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__2___closed__2); -l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__1 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__1(); -lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__1); -l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__2 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__2(); -lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__2); -l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__3 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__3(); -lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__3); -l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__4 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__4(); -lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__4); -l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__5 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__5(); -lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__5); -l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__6 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__6(); -lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__6); -l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__7 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__7(); -lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__7); -l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__8 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__8(); -lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__8); -l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__9 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__9(); -lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__9); -l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__10 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__10(); -lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__10); -l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__11 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__11(); -lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__11); -l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__12 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__12(); -lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__12); -l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__13 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__13(); -lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__13); -l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__14 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__14(); -lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__14); -l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__15 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__15(); -lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__15); -l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__16 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__16(); -lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__16); -l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__17 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__17(); -lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__17); -l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__18 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__18(); -lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__18); -l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__19 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__19(); -lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__4___closed__19); -l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__5___closed__1 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__5___closed__1(); -lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__5___closed__1); -l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__5___closed__2 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__5___closed__2(); -lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__5___closed__2); -l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__5___closed__3 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__5___closed__3(); -lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__5___closed__3); -l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__5___closed__4 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__5___closed__4(); -lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__5___closed__4); -l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__5___closed__5 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__5___closed__5(); -lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__5___closed__5); -l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__5___closed__6 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__5___closed__6(); -lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___lambda__5___closed__6); -l_Lean_Linter_initFn____x40_Lean_Linter_Builtin___hyg_2470____closed__1 = _init_l_Lean_Linter_initFn____x40_Lean_Linter_Builtin___hyg_2470____closed__1(); -lean_mark_persistent(l_Lean_Linter_initFn____x40_Lean_Linter_Builtin___hyg_2470____closed__1); -res = l_Lean_Linter_initFn____x40_Lean_Linter_Builtin___hyg_2470_(lean_io_mk_world()); +l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__3___closed__1 = _init_l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__3___closed__1(); +lean_mark_persistent(l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__3___closed__1); +l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__3___closed__2 = _init_l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__3___closed__2(); +lean_mark_persistent(l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__3___closed__2); +l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__3___closed__3 = _init_l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__3___closed__3(); +lean_mark_persistent(l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__3___closed__3); +l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__3___closed__4 = _init_l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__3___closed__4(); +lean_mark_persistent(l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__3___closed__4); +l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__3___closed__5 = _init_l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__3___closed__5(); +lean_mark_persistent(l_Lean_Linter_logLint___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__3___closed__5); +l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___lambda__1___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___lambda__1___closed__1(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___lambda__1___closed__1); +l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___lambda__1___closed__2 = _init_l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___lambda__1___closed__2(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___lambda__1___closed__2); +l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___lambda__1___closed__3 = _init_l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___lambda__1___closed__3(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___lambda__1___closed__3); +l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___lambda__1___closed__4 = _init_l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___lambda__1___closed__4(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___lambda__1___closed__4); +l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__1(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__1); +l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__2 = _init_l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__2(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__2); +l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__3 = _init_l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__3(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__3); +l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__4 = _init_l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__4(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__4); +l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__5 = _init_l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__5(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__5); +l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__6 = _init_l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__6(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__6); +l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__7 = _init_l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__7(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__7); +l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__8 = _init_l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__8(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__8); +l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__9 = _init_l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__9(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__4___closed__9); +l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__5___closed__1 = _init_l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__5___closed__1(); +lean_mark_persistent(l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__5___closed__1); +l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__5___closed__2 = _init_l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__5___closed__2(); +lean_mark_persistent(l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__5___closed__2); +l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__5___closed__3 = _init_l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__5___closed__3(); +lean_mark_persistent(l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__5___closed__3); +l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__5___closed__4 = _init_l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__5___closed__4(); +lean_mark_persistent(l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__5___closed__4); +l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__5___closed__5 = _init_l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__5___closed__5(); +lean_mark_persistent(l_Array_anyMUnsafe_any___at_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___spec__5___closed__5); +l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__1___closed__1 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__1___closed__1); +l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__1___closed__2 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__1___closed__2(); +lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__1___closed__2); +l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__1___closed__3 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__1___closed__3(); +lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__1___closed__3); +l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__1___closed__4 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__1___closed__4(); +lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__1___closed__4); +l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__2___closed__1 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__2___closed__1); +l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__2___closed__2 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__2___closed__2(); +lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__2___closed__2); +l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__1 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__1); +l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__2 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__2(); +lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__2); +l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__3 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__3(); +lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__3); +l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__4 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__4(); +lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__4); +l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__5 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__5(); +lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__5); +l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__6 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__6(); +lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__6); +l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__7 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__7(); +lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__7); +l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__8 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__8(); +lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__8); +l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__9 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__9(); +lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__9); +l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__10 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__10(); +lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__10); +l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__11 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__11(); +lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__11); +l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__12 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__12(); +lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__12); +l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__13 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__13(); +lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__13); +l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__14 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__14(); +lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__14); +l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__15 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__15(); +lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__15); +l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__16 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__16(); +lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__16); +l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__17 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__17(); +lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__17); +l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__18 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__18(); +lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__18); +l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__19 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__19(); +lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__4___closed__19); +l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__5___closed__1 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__5___closed__1(); +lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__5___closed__1); +l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__5___closed__2 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__5___closed__2(); +lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__5___closed__2); +l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__5___closed__3 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__5___closed__3(); +lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__5___closed__3); +l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__5___closed__4 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__5___closed__4(); +lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__5___closed__4); +l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__5___closed__5 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__5___closed__5(); +lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__5___closed__5); +l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__5___closed__6 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__5___closed__6(); +lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___elambda__1___lambda__5___closed__6); +l_Lean_Linter_suspiciousUnexpanderPatterns___closed__1 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___closed__1(); +lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___closed__1); +l_Lean_Linter_suspiciousUnexpanderPatterns___closed__2 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___closed__2(); +lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___closed__2); +l_Lean_Linter_suspiciousUnexpanderPatterns___closed__3 = _init_l_Lean_Linter_suspiciousUnexpanderPatterns___closed__3(); +lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns___closed__3); +l_Lean_Linter_suspiciousUnexpanderPatterns = _init_l_Lean_Linter_suspiciousUnexpanderPatterns(); +lean_mark_persistent(l_Lean_Linter_suspiciousUnexpanderPatterns); +res = l_Lean_Linter_initFn____x40_Lean_Linter_Builtin___hyg_2485_(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/Linter/Deprecated.c b/stage0/stdlib/Lean/Linter/Deprecated.c index 0b2895b9a4e3..5c50d2ea2efa 100644 --- a/stage0/stdlib/Lean/Linter/Deprecated.c +++ b/stage0/stdlib/Lean/Linter/Deprecated.c @@ -101,7 +101,6 @@ extern lean_object* l_Std_Format_defWidth; LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoTree___at_Lean_Linter_initFn____x40_Lean_Linter_Deprecated___hyg_47____spec__17___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_pushInfoLeaf___at_Lean_Linter_initFn____x40_Lean_Linter_Deprecated___hyg_47____spec__16___closed__3; static lean_object* l_Lean_resolveGlobalConst___at_Lean_Linter_initFn____x40_Lean_Linter_Deprecated___hyg_47____spec__4___closed__3; -lean_object* l_Lean_Option_register___at_Std_Format_initFn____x40_Lean_Data_Format___hyg_95____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Linter_isDeprecated(lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofExpr(lean_object*); LEAN_EXPORT lean_object* l_Lean_resolveGlobalName___at_Lean_Linter_initFn____x40_Lean_Linter_Deprecated___hyg_47____spec__8(lean_object*, lean_object*, lean_object*, lean_object*); @@ -126,6 +125,7 @@ static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_Deprecated___hyg_47_ LEAN_EXPORT lean_object* l_Lean_Linter_checkDeprecated___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_isNone(lean_object*); lean_object* l_Lean_ResolveName_resolveGlobalName(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_Deprecated___hyg_47____lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_Deprecated___hyg_47____closed__8; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Linter_initFn____x40_Lean_Linter_Deprecated___hyg_47____spec__11(lean_object*, lean_object*, lean_object*, lean_object*); @@ -247,7 +247,7 @@ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_2 = l_Lean_Linter_initFn____x40_Lean_Linter_Deprecated___hyg_6____closed__3; x_3 = l_Lean_Linter_initFn____x40_Lean_Linter_Deprecated___hyg_6____closed__6; x_4 = l_Lean_Linter_initFn____x40_Lean_Linter_Deprecated___hyg_6____closed__9; -x_5 = l_Lean_Option_register___at_Std_Format_initFn____x40_Lean_Data_Format___hyg_95____spec__1(x_2, x_3, x_4, x_1); +x_5 = l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__1(x_2, x_3, x_4, x_1); return x_5; } } diff --git a/stage0/stdlib/Lean/Linter/MissingDocs.c b/stage0/stdlib/Lean/Linter/MissingDocs.c index e175c339aa56..38f9d4ded235 100644 --- a/stage0/stdlib/Lean/Linter/MissingDocs.c +++ b/stage0/stdlib/Lean/Linter/MissingDocs.c @@ -13,75 +13,75 @@ #ifdef __cplusplus extern "C" { #endif -static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__13; static lean_object* l_Lean_Linter_MissingDocs_lintNamed___closed__1; lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_checkDecl___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Linter_MissingDocs_handleMutual(lean_object*); static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkSyntax___closed__1; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Linter_MissingDocs_handleIn___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Linter_MissingDocs_missingDocs___closed__1; static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkRegisterBuiltinOption___closed__4; +static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__7; lean_object* l_Lean_Elab_Command_withScope___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__4___closed__2; -static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__4___closed__3; -LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Linter_MissingDocs_checkDecl___spec__7___at_Lean_Linter_MissingDocs_checkDecl___spec__8(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_handleMutual___closed__3; static lean_object* l_Lean_Linter_MissingDocs_declModifiersPubNoDoc___closed__3; static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_559____closed__3; LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_checkDecl___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkRegisterBuiltinOption___closed__1; +LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__7(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_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_559____closed__6; static lean_object* l_Lean_Linter_MissingDocs_checkNotation___closed__2; size_t lean_usize_shift_right(size_t, size_t); +static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__22; LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_checkDecl(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1178_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_559_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_checkClassAbbrev(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_MissingDocs_lintDeclHead___closed__4; static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkRegisterSimpAttr___closed__3; static lean_object* l_Lean_Linter_MissingDocs_lintDeclHead___closed__18; -static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1___closed__2; +static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__2___closed__1; static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkMacro___closed__1; lean_object* l_Lean_logAt___at_Lean_Elab_Command_elabCommand___spec__4(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__2___closed__5; static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkNotation___closed__1; LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_checkRegisterOption___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkDecl(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_missingDocs___elambda__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_MissingDocs_mkHandlerUnsafe___closed__1; -static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__5; +LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ConstantInfo_type(lean_object*); -static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__22; static lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_MissingDocs_hasInheritDoc___spec__1___closed__3; static lean_object* l_Lean_Elab_elabSetOption___at_Lean_Linter_MissingDocs_handleIn___spec__1___closed__1; LEAN_EXPORT lean_object* l_IO_ofExcept___at_Lean_Linter_MissingDocs_mkHandlerUnsafe___spec__2(lean_object*, lean_object*); lean_object* l_Lean_ConstantInfo_levelParams(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__2___closed__4; static lean_object* l_Lean_Linter_MissingDocs_checkRegisterSimpAttr___closed__1; -static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__12; uint64_t lean_uint64_of_nat(lean_object*); lean_object* l_Lean_mkAppB(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkNotation(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkSyntaxCat(lean_object*); static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_559____closed__1; size_t l___private_Lean_Data_HashSet_0__Lean_HashSetImp_mkIdx(lean_object*, uint64_t, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_MissingDocs___hyg_8____closed__1; -LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_MissingDocs_handleIn___rarg___closed__1; -static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1___closed__1; +static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__10; lean_object* l_Lean_Name_toString(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_checkSyntaxAbbrev(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashSetImp_contains___at_Lean_Linter_MissingDocs_checkDecl___spec__16___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__11; static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkRegisterOption___closed__3; lean_object* l_Lean_Syntax_getId(lean_object*); static lean_object* l_Lean_Linter_MissingDocs_lintDeclHead___closed__6; LEAN_EXPORT lean_object* l_IO_ofExcept___at_Lean_Linter_MissingDocs_mkHandlerUnsafe___spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Linter_MissingDocs_missingDocs___spec__1___boxed(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkInit___closed__1; -static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__9; +static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__8; lean_object* l_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_559____lambda__1___closed__1; +static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__9; static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkMacro___closed__3; lean_object* l_Lean_mkHashSetImp___rarg(lean_object*); static lean_object* l_Lean_Linter_MissingDocs_lintDeclHead___closed__10; @@ -94,18 +94,15 @@ static lean_object* l_Lean_Elab_elabSetOption___at_Lean_Linter_MissingDocs_handl uint8_t lean_usize_dec_eq(size_t, size_t); static lean_object* l_Lean_Elab_elabSetOption_setOption___at_Lean_Linter_MissingDocs_handleIn___spec__4___closed__1; LEAN_EXPORT lean_object* l_Lean_Linter_logLint___at_Lean_Linter_MissingDocs_lint___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getArgs(lean_object*); static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkElab___closed__3; -static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__10; static lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at_Lean_Linter_MissingDocs_checkDecl___spec__12___closed__1; static lean_object* l_Lean_Linter_logLint___at_Lean_Linter_MissingDocs_lint___spec__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_addBuiltinHandler(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__8___closed__1; lean_object* lean_array_fget(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_checkClassAbbrev___boxed(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_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Linter_MissingDocs_missingDocs___closed__2; LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_mkSimpleHandler___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_559____lambda__6___closed__2; LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_checkSyntax(lean_object*, lean_object*, lean_object*, lean_object*); @@ -113,9 +110,11 @@ static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkSyntax___closed_ static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_handleIn___closed__1; lean_object* lean_environment_find(lean_object*, lean_object*); static lean_object* l_Lean_Linter_logLint___at_Lean_Linter_MissingDocs_lint___spec__1___closed__2; +LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkRegisterBuiltinOption___closed__5; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Linter_MissingDocs_hasInheritDoc___spec__1(lean_object*, size_t, size_t); uint8_t l_Lean_Linter_getLinterValue(lean_object*, lean_object*); +static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__3; lean_object* l_Lean_Elab_InfoTree_foldInfo___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_MissingDocs_checkDecl___spec__17___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkNotation___closed__2; @@ -125,44 +124,45 @@ lean_object* l_Lean_Name_mkStr5(lean_object*, lean_object*, lean_object*, lean_o static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_559____closed__10; static lean_object* l_Lean_Linter_MissingDocs_lintDeclHead___closed__9; lean_object* l_Lean_stringToMessageData(lean_object*); -static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__5___closed__2; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_MissingDocs_checkDecl___spec__18___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_Linter_MissingDocs_missingDocsExt; static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_handleIn___closed__2; static lean_object* l_Lean_Linter_MissingDocs_lintDeclHead___closed__17; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_MissingDocs___hyg_8____closed__7; lean_object* l_Lean_instInhabitedPersistentArrayNode(lean_object*); -static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__7; LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at_Lean_Linter_MissingDocs_checkDecl___spec__12(lean_object*, size_t, size_t, lean_object*); -static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__20; -LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_KVMap_insertCore(lean_object*, lean_object*, lean_object*); lean_object* l_List_head_x21___rarg(lean_object*, lean_object*); uint8_t lean_string_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_MissingDocs_checkDecl___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1___closed__3; LEAN_EXPORT uint8_t l_Lean_Linter_getLinterMissingDocs(lean_object*); +static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__6; LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_checkSyntaxCat(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_MissingDocs_checkDecl___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_512_(lean_object*); static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_559____closed__4; static lean_object* l_Lean_Linter_MissingDocs_lintDeclHead___closed__16; LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_mkHandlerUnsafe___elambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1___closed__4; +static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__24; +static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__8___closed__2; LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_checkSimpLike(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1163_(lean_object*); static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkMixfix___closed__1; LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_checkMacro___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_checkSimpLike___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_MissingDocs_lintDeclHead___closed__8; +static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__5___closed__2; lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkClassAbbrev___closed__1; LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_checkNotation(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__1___closed__4; static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkMacro___closed__2; size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_checkElab___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__18; +static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__1___closed__2; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_MissingDocs___hyg_8____closed__9; LEAN_EXPORT lean_object* l_Lean_Linter_getLinterMissingDocs___boxed(lean_object*); +static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__4; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_MissingDocs_checkDecl___spec__17(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_MissingDocs_checkSyntax___closed__1; static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkMixfix___closed__2; @@ -174,44 +174,46 @@ LEAN_EXPORT lean_object* l_Lean_Elab_elabSetOption___at_Lean_Linter_MissingDocs_ LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_checkInit(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Linter_MissingDocs_handleIn___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); -static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__2___closed__3; LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_hasInheritDoc___boxed(lean_object*); +static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__4___closed__3; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l_Lean_Linter_MissingDocs_mkHandlerUnsafe___closed__6; static lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_MissingDocs_hasInheritDoc___spec__1___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkInit(lean_object*); static lean_object* l_Lean_Linter_MissingDocs_lintDeclHead___closed__20; -static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__17; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_MissingDocs_checkDecl___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_expr_eqv(lean_object*, lean_object*); lean_object* l_Lean_declareBuiltin(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkClassAbbrev(lean_object*); lean_object* l_Lean_registerPersistentEnvExtensionUnsafe___rarg(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkSyntaxAbbrev___closed__1; +static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__20; static lean_object* l_Lean_Linter_MissingDocs_lintDeclHead___closed__2; lean_object* l_Lean_MessageData_ofSyntax(lean_object*); lean_object* l_Lean_Syntax_getAtomVal(lean_object*); static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkElab___closed__2; -static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1___closed__8; +LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_MissingDocs_handleIn___rarg___closed__2; LEAN_EXPORT lean_object* l_List_replace___at_Lean_Linter_MissingDocs_checkDecl___spec__9(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_checkRegisterOption(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkElab(lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkMixfix(lean_object*); -static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__23; lean_object* l_Lean_PersistentEnvExtension_addEntry___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkRegisterBuiltinOption(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_missingDocs(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_missingDocs; lean_object* l_Lean_PersistentEnvExtension_getState___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__2___closed__3; lean_object* lean_eval_const(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_HashSetImp_contains___at_Lean_Linter_MissingDocs_checkDecl___spec__16(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_addCompletionInfo___at_Lean_Linter_MissingDocs_handleIn___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Command_getRef(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_MissingDocs_mkHandlerUnsafe___closed__7; static lean_object* l_Lean_Elab_elabSetOption_setOption___at_Lean_Linter_MissingDocs_handleIn___spec__4___closed__2; static lean_object* l_Lean_Linter_MissingDocs_checkInit___closed__1; lean_object* l_Lean_Syntax_getKind(lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_checkRegisterBuiltinOption___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__5___closed__1; lean_object* l_Lean_Elab_getBetterRef(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_559____lambda__4(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_checkMixfix___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -219,6 +221,8 @@ lean_object* l_Lean_setEnv___at_Lean_registerParametricAttribute___spec__3(lean_ static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkSyntaxCat___closed__3; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_MissingDocs_checkDecl___spec__17___lambda__1___closed__1; static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkSyntaxAbbrev___closed__2; +static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__14; +LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__5(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_lintStructField(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_handleIn___rarg(lean_object*, lean_object*, lean_object*, lean_object*); 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*); @@ -236,25 +240,23 @@ static lean_object* l_Lean_Linter_MissingDocs_lintDeclHead___closed__14; lean_object* l_Lean_Attribute_Builtin_getIdent(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_MissingDocs_checkDecl___spec__1___closed__1; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_559____spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_lintStructField___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_List_isEmpty___rarg(lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkRegisterOption___closed__2; -LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_isStrLit_x3f(lean_object*); LEAN_EXPORT lean_object* l_IO_ofExcept___at_Lean_Linter_MissingDocs_mkHandlerUnsafe___spec__2___boxed(lean_object*, lean_object*); lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_handleIn___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_lintNamed___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkSimpLike___closed__4; +static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__4___closed__4; LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_lintNamed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_559____lambda__6(lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkMacro(lean_object*); static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkClassAbbrev___closed__2; -static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__19; +LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Linter_MissingDocs_missingDocs___elambda__1___spec__1___boxed(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_handleMutual___closed__1; -LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Linter_MissingDocs_checkDecl___spec__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_MissingDocs___hyg_8____closed__4; static lean_object* l_Lean_Linter_MissingDocs_checkSimpLike___closed__1; @@ -266,20 +268,19 @@ static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingD static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_MissingDocs___hyg_8____closed__5; static lean_object* l_Lean_Linter_MissingDocs_checkMacro___closed__1; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_MissingDocs___hyg_8____closed__8; -static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__11; lean_object* l_List_lengthTRAux___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashSetImp_insert___at_Lean_Linter_MissingDocs_checkDecl___spec__3(lean_object*, lean_object*); lean_object* l_Lean_Syntax_getSepArgs(lean_object*); static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkNotation___closed__3; static lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_MissingDocs_hasInheritDoc___spec__1___closed__6; +static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__19; LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_checkSyntaxAbbrev___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_MissingDocs_lintDeclHead___closed__11; -static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1___closed__7; +static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__1; static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkDecl___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlMAux___at_Lean_Linter_MissingDocs_checkDecl___spec__13(lean_object*, lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_559____lambda__2(lean_object*, lean_object*); -static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__5___closed__1; static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkRegisterBuiltinOption___closed__2; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_linter_missingDocs; @@ -297,11 +298,12 @@ LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_getHandlers___boxed(lean_obje LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_lintField(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Linter_MissingDocs_handleMutual___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_559____closed__5; +static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__1___closed__1; static lean_object* l_Lean_Linter_MissingDocs_mkHandlerUnsafe___closed__4; LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_checkRegisterSimpAttr___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__4___closed__1; -static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__6; +LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Linter_MissingDocs_missingDocs___elambda__1___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_lint(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkRegisterSimpAttr___closed__2; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_559____spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Linter_MissingDocs_hasInheritDoc(lean_object*); @@ -309,14 +311,13 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_MissingDocs_c static lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_MissingDocs_hasInheritDoc___spec__1___closed__5; static lean_object* l_Lean_Linter_MissingDocs_lintField___closed__1; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_MissingDocs___hyg_8____closed__3; -static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__1; lean_object* l_Lean_Elab_Command_addLinter(lean_object*, lean_object*); lean_object* lean_usize_to_nat(size_t); -LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220_(lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_MissingDocs_hasInheritDoc___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_MissingDocs_checkDecl___spec__17___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_Linter_MissingDocs_handleMutual(uint8_t); +static lean_object* l_Lean_Linter_MissingDocs_missingDocs___closed__3; static lean_object* l_Lean_Linter_MissingDocs_lintDeclHead___closed__7; static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkRegisterBuiltinOption___closed__3; static lean_object* l_Lean_Linter_MissingDocs_checkClassAbbrev___closed__1; @@ -324,42 +325,42 @@ LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_handleIn(uint8_t); static lean_object* l_Lean_Linter_MissingDocs_checkNotation___closed__3; static lean_object* l_Lean_Linter_logLint___at_Lean_Linter_MissingDocs_lint___spec__1___closed__4; LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_checkNotation___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Linter_MissingDocs_missingDocs___spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__4___closed__2; static lean_object* l_Lean_Linter_MissingDocs_lintDeclHead___closed__12; LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_checkRegisterSimpAttr(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__16; static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkSyntax___closed__2; LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_lintField___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_throwError___at_Lean_AttributeImpl_erase___default___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__2___closed__2; -static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__8; static lean_object* l_Lean_Elab_elabSetOption___at_Lean_Linter_MissingDocs_handleIn___spec__1___closed__5; lean_object* l_Lean_Syntax_setArgs(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_MissingDocs___hyg_8_(lean_object*); -static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1163____closed__1; static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkSyntaxCat___closed__2; static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkRegisterSimpAttr___closed__4; static lean_object* l_Lean_Linter_logLint___at_Lean_Linter_MissingDocs_lint___spec__1___closed__5; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_MissingDocs_checkDecl___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_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_559____lambda__5(lean_object*); lean_object* l_Lean_Expr_app___override(lean_object*, lean_object*); +static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__1___closed__3; LEAN_EXPORT lean_object* l_Lean_mkHashSet___at_Lean_Linter_MissingDocs_checkDecl___spec__10___boxed(lean_object*); -static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1___closed__6; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_MissingDocs_checkDecl___spec__18___closed__2; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_logLint___at_Lean_Linter_MissingDocs_lint___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__15; +LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__4(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_checkMacro(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_mkHandlerUnsafe___elambda__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); +static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__8___closed__1; +static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__13; +static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__1___closed__6; static lean_object* l_Lean_Linter_MissingDocs_addHandler___closed__1; -static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__4; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Linter_MissingDocs_declModifiersPubNoDoc(lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_handleMutual___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkElab___closed__1; static lean_object* l_Lean_Linter_logLint___at_Lean_Linter_MissingDocs_lint___spec__1___closed__1; static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkRegisterOption___closed__4; +static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__1___closed__5; +static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__16; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Linter_MissingDocs_checkDecl___spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkSyntaxAbbrev___closed__4; static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkInit___closed__3; @@ -370,37 +371,40 @@ LEAN_EXPORT lean_object* l_Lean_Elab_elabSetOption_setOption___at_Lean_Linter_Mi static lean_object* l_Lean_Linter_MissingDocs_declModifiersPubNoDoc___closed__1; LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_559____lambda__3___boxed(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_isNone(lean_object*); -static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__4___closed__4; -LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__21; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_MissingDocs_checkDecl___spec__2___closed__1; +LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__9(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_MissingDocs_lint___closed__2; lean_object* lean_nat_sub(lean_object*, lean_object*); -static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__8___closed__2; +static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__17; lean_object* lean_nat_mul(lean_object*, lean_object*); +static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__2___closed__5; LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_handleIn___rarg___lambda__1(lean_object*, lean_object*); -static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__2; static lean_object* l_Lean_Linter_MissingDocs_lintDeclHead___closed__19; lean_object* l_Lean_Syntax_getRange_x3f(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_559____spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__15; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Linter_MissingDocs_checkDecl___spec__15(lean_object*, size_t, size_t, lean_object*); uint8_t l_Lean_DataValue_sameCtor(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_lintDeclHead___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Linter_MissingDocs_checkDecl___spec__15___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_checkDecl___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_elem___at_Lean_Linter_MissingDocs_checkDecl___spec__4___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__24; static lean_object* l_Lean_Linter_MissingDocs_mkHandlerUnsafe___closed__5; lean_object* l_Lean_getOptionDecl(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_559____lambda__6___boxed(lean_object*); +static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__2; lean_object* l_Array_ofSubarray___rarg(lean_object*); LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Linter_MissingDocs_checkDecl___spec__7(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Environment_getModuleIdxFor_x3f(lean_object*, lean_object*); lean_object* lean_erase_macro_scopes(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Nat_decEq___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_mkHandlerUnsafe(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_registerBuiltinAttribute(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_559____lambda__1(lean_object*); lean_object* l_List_reverse___rarg(lean_object*); +static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__23; lean_object* l___private_Lean_ToExpr_0__Lean_Name_toExprAux(lean_object*); static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkSimpLike___closed__3; size_t lean_usize_sub(size_t, size_t); @@ -430,18 +434,18 @@ lean_object* lean_array_uget(lean_object*, size_t); static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_559____closed__8; LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_handleIn___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_handleMutual___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231_(lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_MissingDocs_checkDecl___spec__18___closed__3; -LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__4(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); lean_object* lean_io_error_to_string(lean_object*); static lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_MissingDocs_hasInheritDoc___spec__1___closed__4; static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_559____lambda__6___closed__1; -LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__5(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_Linter_MissingDocs_checkSyntaxCat___closed__1; static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkClassAbbrev___closed__3; static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkMixfix___closed__4; size_t lean_usize_shift_left(size_t, size_t); lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__2___closed__2; lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_instBEq___rarg(lean_object*, lean_object*, lean_object*); lean_object* l_List_redLength___rarg(lean_object*); @@ -450,24 +454,23 @@ static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkSimpLike___close static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkRegisterSimpAttr___closed__1; lean_object* lean_string_append(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkSyntaxCat___closed__1; -static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__18; uint8_t l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_367_(uint8_t, uint8_t); static lean_object* l_Lean_Linter_MissingDocs_lintDeclHead___closed__13; +LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); lean_object* l_Lean_Syntax_isNatLit_x3f(lean_object*); lean_object* l_Lean_getConstInfo___at_Lean_Linter_initFn____x40_Lean_Linter_Deprecated___hyg_47____spec__14(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkInit___closed__2; -static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1___closed__5; static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkMixfix___closed__3; uint8_t lean_nat_dec_le(lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkSimpLike(lean_object*); -static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__21; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Linter_MissingDocs_handleIn(lean_object*); -static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__14; +LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_List_elem___at_Lean_Linter_MissingDocs_checkDecl___spec__4(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkRegisterOption___closed__1; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Linter_MissingDocs_checkDecl___spec__14(lean_object*, size_t, size_t, lean_object*); +static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__1___closed__8; lean_object* lean_nat_add(lean_object*, lean_object*); lean_object* l_EStateM_bind___rarg(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_resolveGlobalConstNoOverloadWithInfo___at_Lean_Linter_initFn____x40_Lean_Linter_Deprecated___hyg_47____spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -476,9 +479,10 @@ LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_declModifiersPubNoDoc___boxed LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_SimpleHandler_toHandler___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_559____lambda__3(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_MissingDocs_lintDeclHead___closed__5; -static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__3; +static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__5; static lean_object* l_Lean_Linter_MissingDocs_getHandlers___closed__1; LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_SimpleHandler_toHandler(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__4___closed__1; lean_object* l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Command_expandDeclId___spec__11(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Linter_MissingDocs_handleMutual___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_handleMutual___closed__2; @@ -488,18 +492,17 @@ static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkDecl___closed__3 lean_object* l_Lean_MessageData_ofName(lean_object*); lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabSetOption_setOption___at_Lean_Linter_MissingDocs_handleIn___spec__4___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__9(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_MissingDocs_lint___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at_Lean_Linter_MissingDocs_checkDecl___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_checkMixfix(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__2___closed__4; lean_object* l_Nat_repr(lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_checkInit___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_builtinHandlersRef; LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_checkRegisterBuiltinOption(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__12; size_t lean_usize_land(size_t, size_t); -LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_toArrayAux___rarg(lean_object*, lean_object*); +static lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__1___closed__7; LEAN_EXPORT lean_object* l_Lean_PersistentArray_foldlM___at_Lean_Linter_MissingDocs_checkDecl___spec__11(lean_object*, lean_object*, lean_object*); static lean_object* _init_l_Lean_Linter_initFn____x40_Lean_Linter_MissingDocs___hyg_8____closed__1() { _start: @@ -2017,7 +2020,7 @@ lean_dec(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Linter_MissingDocs_missingDocs___spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Linter_MissingDocs_missingDocs___elambda__1___spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -2057,7 +2060,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_missingDocs(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_missingDocs___elambda__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; @@ -2075,7 +2078,7 @@ x_10 = l_Lean_Linter_MissingDocs_getHandlers(x_9); lean_dec(x_9); lean_inc(x_1); x_11 = l_Lean_Syntax_getKind(x_1); -x_12 = l_Lean_RBNode_find___at_Lean_Linter_MissingDocs_missingDocs___spec__1(x_10, x_11); +x_12 = l_Lean_RBNode_find___at_Lean_Linter_MissingDocs_missingDocs___elambda__1___spec__1(x_10, x_11); lean_dec(x_11); lean_dec(x_10); if (lean_obj_tag(x_12) == 0) @@ -2133,7 +2136,7 @@ x_29 = l_Lean_Linter_MissingDocs_getHandlers(x_28); lean_dec(x_28); lean_inc(x_1); x_30 = l_Lean_Syntax_getKind(x_1); -x_31 = l_Lean_RBNode_find___at_Lean_Linter_MissingDocs_missingDocs___spec__1(x_29, x_30); +x_31 = l_Lean_RBNode_find___at_Lean_Linter_MissingDocs_missingDocs___elambda__1___spec__1(x_29, x_30); lean_dec(x_30); lean_dec(x_29); if (lean_obj_tag(x_31) == 0) @@ -2179,29 +2182,61 @@ return x_45; } } } -LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Linter_MissingDocs_missingDocs___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_Linter_MissingDocs_missingDocs___closed__1() { _start: { -lean_object* x_3; -x_3 = l_Lean_RBNode_find___at_Lean_Linter_MissingDocs_missingDocs___spec__1(x_1, x_2); -lean_dec(x_2); -lean_dec(x_1); +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Linter_initFn____x40_Lean_Linter_MissingDocs___hyg_8____closed__7; +x_2 = l_Lean_Linter_initFn____x40_Lean_Linter_MissingDocs___hyg_8____closed__8; +x_3 = l_Lean_Linter_MissingDocs_mkHandlerUnsafe___closed__5; +x_4 = l_Lean_Linter_initFn____x40_Lean_Linter_MissingDocs___hyg_8____closed__2; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Linter_MissingDocs_missingDocs___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Linter_MissingDocs_missingDocs___elambda__1), 4, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Linter_MissingDocs_missingDocs___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Linter_MissingDocs_missingDocs___closed__2; +x_2 = l_Lean_Linter_MissingDocs_missingDocs___closed__1; +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_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1163____closed__1() { +static lean_object* _init_l_Lean_Linter_MissingDocs_missingDocs() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Linter_MissingDocs_missingDocs), 4, 0); +x_1 = l_Lean_Linter_MissingDocs_missingDocs___closed__3; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1163_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Linter_MissingDocs_missingDocs___elambda__1___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_RBNode_find___at_Lean_Linter_MissingDocs_missingDocs___elambda__1___spec__1(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1178_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1163____closed__1; +x_2 = l_Lean_Linter_MissingDocs_missingDocs; x_3 = l_Lean_Elab_Command_addLinter(x_2, x_1); return x_3; } @@ -2239,7 +2274,7 @@ return x_13; } } } -static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1___closed__1() { +static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__1___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; @@ -2251,17 +2286,17 @@ x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1___closed__2() { +static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____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_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1___closed__1; +x_2 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__1___closed__1; x_3 = l_Lean_Expr_const___override(x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1___closed__3() { +static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__1___closed__3() { _start: { lean_object* x_1; @@ -2269,29 +2304,29 @@ x_1 = lean_mk_string_from_bytes("addBuiltinHandler", 17); return x_1; } } -static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1___closed__4() { +static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__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_Linter_initFn____x40_Lean_Linter_MissingDocs___hyg_8____closed__7; x_2 = l_Lean_Linter_initFn____x40_Lean_Linter_MissingDocs___hyg_8____closed__8; x_3 = l_Lean_Linter_MissingDocs_mkHandlerUnsafe___closed__5; -x_4 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1___closed__3; +x_4 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__1___closed__3; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1___closed__5() { +static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____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_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1___closed__4; +x_2 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__1___closed__4; x_3 = l_Lean_Expr_const___override(x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1___closed__6() { +static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__1___closed__6() { _start: { lean_object* x_1; @@ -2299,7 +2334,7 @@ x_1 = lean_mk_string_from_bytes("toHandler", 9); return x_1; } } -static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1___closed__7() { +static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__1___closed__7() { _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; @@ -2307,28 +2342,28 @@ x_1 = l_Lean_Linter_initFn____x40_Lean_Linter_MissingDocs___hyg_8____closed__7; x_2 = l_Lean_Linter_initFn____x40_Lean_Linter_MissingDocs___hyg_8____closed__8; x_3 = l_Lean_Linter_MissingDocs_mkHandlerUnsafe___closed__5; x_4 = l_Lean_Linter_MissingDocs_mkHandlerUnsafe___closed__6; -x_5 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1___closed__6; +x_5 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__1___closed__6; x_6 = l_Lean_Name_mkStr5(x_1, x_2, x_3, x_4, x_5); return x_6; } } -static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1___closed__8() { +static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____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_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1___closed__7; +x_2 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__1___closed__7; x_3 = l_Lean_Expr_const___override(x_2, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____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_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____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; lean_object* x_9; lean_object* x_10; uint8_t x_11; lean_object* x_12; x_8 = l_Lean_ConstantInfo_type(x_1); x_9 = lean_box(0); -x_10 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1___closed__2; +x_10 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__1___closed__2; x_11 = lean_expr_eqv(x_8, x_10); lean_dec(x_8); x_12 = l___private_Lean_ToExpr_0__Lean_Name_toExprAux(x_2); @@ -2337,7 +2372,7 @@ if (x_11 == 0) lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_inc(x_3); x_13 = l_Lean_Expr_const___override(x_3, x_9); -x_14 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1___closed__5; +x_14 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__1___closed__5; x_15 = l_Lean_mkAppB(x_14, x_12, x_13); x_16 = l_Lean_declareBuiltin(x_3, x_15, x_5, x_6, x_7); return x_16; @@ -2347,16 +2382,16 @@ 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_inc(x_3); x_17 = l_Lean_Expr_const___override(x_3, x_9); -x_18 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1___closed__8; +x_18 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__1___closed__8; x_19 = l_Lean_Expr_app___override(x_18, x_17); -x_20 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1___closed__5; +x_20 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__1___closed__5; x_21 = l_Lean_mkAppB(x_20, x_12, x_19); x_22 = l_Lean_declareBuiltin(x_3, x_21, x_5, x_6, x_7); return x_22; } } } -static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__2___closed__1() { +static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__2___closed__1() { _start: { lean_object* x_1; lean_object* x_2; @@ -2365,7 +2400,7 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__2___closed__2() { +static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__2___closed__2() { _start: { lean_object* x_1; lean_object* x_2; @@ -2374,7 +2409,7 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__2___closed__3() { +static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__2___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; @@ -2386,27 +2421,27 @@ x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__2___closed__4() { +static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____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_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__2___closed__3; +x_2 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__2___closed__3; x_3 = l_Lean_Expr_const___override(x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__2___closed__5() { +static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__2___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1___closed__1; +x_2 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__1___closed__1; x_3 = l_Lean_Expr_const___override(x_2, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____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_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____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; @@ -2453,11 +2488,11 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean lean_dec(x_15); lean_dec(x_8); x_19 = l_Lean_MessageData_ofName(x_1); -x_20 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__2___closed__1; +x_20 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__2___closed__1; x_21 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_21, 0, x_20); lean_ctor_set(x_21, 1, x_19); -x_22 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__2___closed__2; +x_22 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__2___closed__2; x_23 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_23, 0, x_21); lean_ctor_set(x_23, 1, x_22); @@ -2487,12 +2522,12 @@ else { lean_object* x_29; lean_object* x_30; uint8_t x_31; x_29 = l_Lean_ConstantInfo_type(x_8); -x_30 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__2___closed__4; +x_30 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__2___closed__4; x_31 = lean_expr_eqv(x_29, x_30); if (x_31 == 0) { lean_object* x_32; uint8_t x_33; -x_32 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__2___closed__5; +x_32 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__2___closed__5; x_33 = lean_expr_eqv(x_29, x_32); lean_dec(x_29); if (x_33 == 0) @@ -2501,11 +2536,11 @@ lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean lean_dec(x_15); lean_dec(x_8); x_34 = l_Lean_MessageData_ofName(x_1); -x_35 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__2___closed__1; +x_35 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__2___closed__1; x_36 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_36, 0, x_35); lean_ctor_set(x_36, 1, x_34); -x_37 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__2___closed__2; +x_37 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__2___closed__2; x_38 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_38, 0, x_36); lean_ctor_set(x_38, 1, x_37); @@ -2535,7 +2570,7 @@ else { lean_object* x_44; lean_object* x_45; x_44 = lean_box(0); -x_45 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1(x_8, x_15, x_1, x_44, x_4, x_5, x_16); +x_45 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__1(x_8, x_15, x_1, x_44, x_4, x_5, x_16); lean_dec(x_5); lean_dec(x_4); lean_dec(x_8); @@ -2547,7 +2582,7 @@ else lean_object* x_46; lean_object* x_47; lean_dec(x_29); x_46 = lean_box(0); -x_47 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1(x_8, x_15, x_1, x_46, x_4, x_5, x_16); +x_47 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__1(x_8, x_15, x_1, x_46, x_4, x_5, x_16); lean_dec(x_5); lean_dec(x_4); lean_dec(x_8); @@ -2637,7 +2672,7 @@ return x_59; } } } -LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____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_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____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) { _start: { lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; @@ -2647,11 +2682,11 @@ 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_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__2(x_1, x_2, x_9, x_4, x_5, x_8); +x_10 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__2(x_1, x_2, x_9, x_4, x_5, x_8); return x_10; } } -static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__4___closed__1() { +static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__4___closed__1() { _start: { lean_object* x_1; @@ -2659,16 +2694,16 @@ x_1 = lean_mk_string_from_bytes("invalid attribute '", 19); return x_1; } } -static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__4___closed__2() { +static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__4___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__4___closed__1; +x_1 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__4___closed__1; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__4___closed__3() { +static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__4___closed__3() { _start: { lean_object* x_1; @@ -2676,16 +2711,16 @@ x_1 = lean_mk_string_from_bytes("', must be global", 17); return x_1; } } -static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__4___closed__4() { +static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__4___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__4___closed__3; +x_1 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__4___closed__3; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__4(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_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__4(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) { _start: { uint8_t x_8; uint8_t x_9; @@ -2697,11 +2732,11 @@ lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean lean_dec(x_3); lean_dec(x_2); x_10 = l_Lean_MessageData_ofName(x_1); -x_11 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__4___closed__2; +x_11 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__4___closed__2; x_12 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_12, 0, x_11); lean_ctor_set(x_12, 1, x_10); -x_13 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__4___closed__4; +x_13 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__4___closed__4; x_14 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_14, 0, x_12); lean_ctor_set(x_14, 1, x_13); @@ -2732,12 +2767,12 @@ else lean_object* x_20; lean_object* x_21; lean_dec(x_1); x_20 = lean_box(0); -x_21 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__3(x_2, x_3, x_20, x_5, x_6, x_7); +x_21 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__3(x_2, x_3, x_20, x_5, x_6, x_7); return x_21; } } } -static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__5___closed__1() { +static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__5___closed__1() { _start: { lean_object* x_1; @@ -2745,25 +2780,25 @@ x_1 = lean_mk_string_from_bytes("attribute cannot be erased", 26); return x_1; } } -static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__5___closed__2() { +static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__5___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__5___closed__1; +x_1 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__5___closed__1; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____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; -x_5 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__5___closed__2; +x_5 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__5___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; } } -LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____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_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____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) { _start: { 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; @@ -2851,7 +2886,7 @@ return x_35; } } } -LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__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_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__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) { _start: { lean_object* x_8; @@ -2899,11 +2934,11 @@ lean_dec(x_16); lean_dec(x_9); lean_dec(x_3); x_20 = l_Lean_MessageData_ofName(x_1); -x_21 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__2___closed__1; +x_21 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__2___closed__1; x_22 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_22, 0, x_21); lean_ctor_set(x_22, 1, x_20); -x_23 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__2___closed__2; +x_23 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__2___closed__2; x_24 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_24, 0, x_22); lean_ctor_set(x_24, 1, x_23); @@ -2934,12 +2969,12 @@ else lean_object* x_30; lean_object* x_31; uint8_t x_32; x_30 = l_Lean_ConstantInfo_type(x_9); lean_dec(x_9); -x_31 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__2___closed__4; +x_31 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__2___closed__4; x_32 = lean_expr_eqv(x_30, x_31); if (x_32 == 0) { lean_object* x_33; uint8_t x_34; -x_33 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__2___closed__5; +x_33 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__2___closed__5; x_34 = lean_expr_eqv(x_30, x_33); lean_dec(x_30); if (x_34 == 0) @@ -2948,11 +2983,11 @@ lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean lean_dec(x_16); lean_dec(x_3); x_35 = l_Lean_MessageData_ofName(x_1); -x_36 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__2___closed__1; +x_36 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__2___closed__1; x_37 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_37, 0, x_36); lean_ctor_set(x_37, 1, x_35); -x_38 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__2___closed__2; +x_38 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__2___closed__2; x_39 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_39, 0, x_37); lean_ctor_set(x_39, 1, x_38); @@ -2982,7 +3017,7 @@ else { lean_object* x_45; lean_object* x_46; x_45 = lean_box(0); -x_46 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__6(x_1, x_16, x_3, x_45, x_5, x_6, x_17); +x_46 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__6(x_1, x_16, x_3, x_45, x_5, x_6, x_17); lean_dec(x_6); lean_dec(x_5); return x_46; @@ -2993,7 +3028,7 @@ else lean_object* x_47; lean_object* x_48; lean_dec(x_30); x_47 = lean_box(0); -x_48 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__6(x_1, x_16, x_3, x_47, x_5, x_6, x_17); +x_48 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__6(x_1, x_16, x_3, x_47, x_5, x_6, x_17); lean_dec(x_6); lean_dec(x_5); return x_48; @@ -3085,7 +3120,7 @@ return x_60; } } } -static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__8___closed__1() { +static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__8___closed__1() { _start: { lean_object* x_1; @@ -3093,16 +3128,16 @@ x_1 = lean_mk_string_from_bytes("', declaration is in an imported module", 39); return x_1; } } -static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__8___closed__2() { +static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__8___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__8___closed__1; +x_1 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__8___closed__1; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____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_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____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) { _start: { lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; @@ -3124,7 +3159,7 @@ if (lean_obj_tag(x_12) == 0) lean_object* x_13; lean_object* x_14; lean_dec(x_3); x_13 = lean_box(0); -x_14 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__7(x_1, x_2, x_11, x_13, x_5, x_6, x_10); +x_14 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__7(x_1, x_2, x_11, x_13, x_5, x_6, x_10); return x_14; } else @@ -3135,11 +3170,11 @@ lean_dec(x_11); lean_dec(x_2); lean_dec(x_1); x_15 = l_Lean_MessageData_ofName(x_3); -x_16 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__4___closed__2; +x_16 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__4___closed__2; x_17 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_17, 0, x_16); lean_ctor_set(x_17, 1, x_15); -x_18 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__8___closed__2; +x_18 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__8___closed__2; x_19 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_19, 0, x_17); lean_ctor_set(x_19, 1, x_18); @@ -3167,7 +3202,7 @@ return x_24; } } } -LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__9(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_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__9(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) { _start: { uint8_t x_8; uint8_t x_9; @@ -3179,11 +3214,11 @@ lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean lean_dec(x_3); lean_dec(x_2); x_10 = l_Lean_MessageData_ofName(x_1); -x_11 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__4___closed__2; +x_11 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__4___closed__2; x_12 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_12, 0, x_11); lean_ctor_set(x_12, 1, x_10); -x_13 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__4___closed__4; +x_13 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__4___closed__4; x_14 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_14, 0, x_12); lean_ctor_set(x_14, 1, x_13); @@ -3213,12 +3248,12 @@ else { lean_object* x_20; lean_object* x_21; x_20 = lean_box(0); -x_21 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__8(x_2, x_3, x_1, x_20, x_5, x_6, x_7); +x_21 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__8(x_2, x_3, x_1, x_20, x_5, x_6, x_7); return x_21; } } } -static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__1() { +static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__1() { _start: { lean_object* x_1; @@ -3226,17 +3261,17 @@ x_1 = lean_mk_string_from_bytes("builtin_missing_docs_handler", 28); return x_1; } } -static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__2() { +static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__1; +x_2 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__3() { +static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -3246,27 +3281,27 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__4() { +static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__3; +x_1 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__3; x_2 = l_Lean_Linter_initFn____x40_Lean_Linter_MissingDocs___hyg_8____closed__8; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__5() { +static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__4; +x_1 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__4; x_2 = l_Lean_Linter_MissingDocs_mkHandlerUnsafe___closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__6() { +static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__6() { _start: { lean_object* x_1; @@ -3274,17 +3309,17 @@ x_1 = lean_mk_string_from_bytes("initFn", 6); return x_1; } } -static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__7() { +static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__5; -x_2 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__6; +x_1 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__5; +x_2 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__8() { +static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__8() { _start: { lean_object* x_1; @@ -3292,47 +3327,47 @@ x_1 = lean_mk_string_from_bytes("_@", 2); return x_1; } } -static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__9() { +static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__7; -x_2 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__8; +x_1 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__7; +x_2 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__8; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__10() { +static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__9; +x_1 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__9; x_2 = l_Lean_Linter_initFn____x40_Lean_Linter_MissingDocs___hyg_8____closed__7; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__11() { +static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__10; +x_1 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__10; x_2 = l_Lean_Linter_initFn____x40_Lean_Linter_MissingDocs___hyg_8____closed__8; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__12() { +static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__11; +x_1 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__11; x_2 = l_Lean_Linter_MissingDocs_mkHandlerUnsafe___closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__13() { +static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__13() { _start: { lean_object* x_1; @@ -3340,27 +3375,27 @@ x_1 = lean_mk_string_from_bytes("_hyg", 4); return x_1; } } -static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__14() { +static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__12; -x_2 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__13; +x_1 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__12; +x_2 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__13; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__15() { +static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__14; -x_2 = lean_unsigned_to_nat(1220u); +x_1 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__14; +x_2 = lean_unsigned_to_nat(1231u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__16() { +static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__16() { _start: { lean_object* x_1; @@ -3368,7 +3403,7 @@ x_1 = lean_mk_string_from_bytes("(builtin) ", 10); return x_1; } } -static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__17() { +static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__17() { _start: { lean_object* x_1; @@ -3376,23 +3411,23 @@ x_1 = lean_mk_string_from_bytes("adds a syntax traversal for the missing docs li return x_1; } } -static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__18() { +static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__16; -x_2 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__17; +x_1 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__16; +x_2 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__17; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__19() { +static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__19() { _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_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__15; -x_2 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__2; -x_3 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__18; +x_1 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__15; +x_2 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__2; +x_3 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__18; x_4 = 1; x_5 = lean_alloc_ctor(0, 3, 1); lean_ctor_set(x_5, 0, x_1); @@ -3402,15 +3437,15 @@ lean_ctor_set_uint8(x_5, sizeof(void*)*3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__20() { +static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__20() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__5___boxed), 4, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__5___boxed), 4, 0); return x_1; } } -static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__21() { +static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__21() { _start: { lean_object* x_1; @@ -3418,33 +3453,33 @@ x_1 = lean_mk_string_from_bytes("missing_docs_handler", 20); return x_1; } } -static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__22() { +static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__22() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__21; +x_2 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__21; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__23() { +static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Linter_initFn____x40_Lean_Linter_MissingDocs___hyg_8____closed__4; -x_2 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__17; +x_2 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__17; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__24() { +static lean_object* _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__24() { _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_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__15; -x_2 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__22; -x_3 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__23; +x_1 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__15; +x_2 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__22; +x_3 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__23; x_4 = 1; x_5 = lean_alloc_ctor(0, 3, 1); lean_ctor_set(x_5, 0, x_1); @@ -3454,15 +3489,15 @@ lean_ctor_set_uint8(x_5, sizeof(void*)*3, x_4); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231_(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; -x_2 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__2; -x_3 = lean_alloc_closure((void*)(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__4___boxed), 7, 1); +x_2 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__2; +x_3 = lean_alloc_closure((void*)(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__4___boxed), 7, 1); lean_closure_set(x_3, 0, x_2); -x_4 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__19; -x_5 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__20; +x_4 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__19; +x_5 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__20; x_6 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_6, 0, x_4); lean_ctor_set(x_6, 1, x_3); @@ -3474,10 +3509,10 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_o x_8 = lean_ctor_get(x_7, 1); lean_inc(x_8); lean_dec(x_7); -x_9 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__22; -x_10 = lean_alloc_closure((void*)(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__9___boxed), 7, 1); +x_9 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__22; +x_10 = lean_alloc_closure((void*)(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__9___boxed), 7, 1); lean_closure_set(x_10, 0, x_9); -x_11 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__24; +x_11 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__24; x_12 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_12, 0, x_11); lean_ctor_set(x_12, 1, x_10); @@ -3509,11 +3544,11 @@ return x_17; } } } -LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____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_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____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_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +x_8 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); @@ -3521,45 +3556,45 @@ lean_dec(x_1); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____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_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____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) { _start: { uint8_t x_8; lean_object* x_9; x_8 = lean_unbox(x_4); lean_dec(x_4); -x_9 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__4(x_1, x_2, x_3, x_8, x_5, x_6, x_7); +x_9 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__4(x_1, x_2, x_3, x_8, x_5, x_6, x_7); return x_9; } } -LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__5___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_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__5(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__5(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_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__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_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__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) { _start: { lean_object* x_8; -x_8 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +x_8 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____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_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____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) { _start: { uint8_t x_8; lean_object* x_9; x_8 = lean_unbox(x_4); lean_dec(x_4); -x_9 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__9(x_1, x_2, x_3, x_8, x_5, x_6, x_7); +x_9 = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__9(x_1, x_2, x_3, x_8, x_5, x_6, x_7); return x_9; } } @@ -8157,7 +8192,7 @@ lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_dec(x_6); x_10 = lean_unsigned_to_nat(2u); x_11 = l_Lean_Syntax_getArg(x_1, x_10); -x_12 = l_Lean_Linter_MissingDocs_missingDocs(x_11, x_2, x_3, x_4); +x_12 = l_Lean_Linter_MissingDocs_missingDocs___elambda__1(x_11, x_2, x_3, x_4); return x_12; } else @@ -8182,7 +8217,7 @@ lean_dec(x_17); x_20 = lean_alloc_closure((void*)(l_Lean_Linter_MissingDocs_handleIn___rarg___lambda__1), 2, 1); lean_closure_set(x_20, 0, x_18); x_21 = l_Lean_Syntax_getArg(x_1, x_15); -x_22 = lean_alloc_closure((void*)(l_Lean_Linter_MissingDocs_missingDocs), 4, 1); +x_22 = lean_alloc_closure((void*)(l_Lean_Linter_MissingDocs_missingDocs___elambda__1), 4, 1); lean_closure_set(x_22, 0, x_21); x_23 = l_Lean_Elab_Command_withScope___rarg(x_20, x_22, x_2, x_3, x_19); return x_23; @@ -8330,7 +8365,7 @@ lean_dec(x_4); x_9 = lean_array_uget(x_1, x_2); lean_inc(x_6); lean_inc(x_5); -x_10 = l_Lean_Linter_MissingDocs_missingDocs(x_9, x_5, x_6, x_7); +x_10 = l_Lean_Linter_MissingDocs_missingDocs___elambda__1(x_9, x_5, x_6, x_7); if (lean_obj_tag(x_10) == 0) { lean_object* x_11; lean_object* x_12; size_t x_13; size_t x_14; @@ -8612,102 +8647,108 @@ lean_dec_ref(res); lean_mark_persistent(l_Lean_Linter_MissingDocs_addHandler___closed__1); l_Lean_Linter_MissingDocs_getHandlers___closed__1 = _init_l_Lean_Linter_MissingDocs_getHandlers___closed__1(); lean_mark_persistent(l_Lean_Linter_MissingDocs_getHandlers___closed__1); -l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1163____closed__1 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1163____closed__1(); -lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1163____closed__1); -res = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1163_(lean_io_mk_world()); +l_Lean_Linter_MissingDocs_missingDocs___closed__1 = _init_l_Lean_Linter_MissingDocs_missingDocs___closed__1(); +lean_mark_persistent(l_Lean_Linter_MissingDocs_missingDocs___closed__1); +l_Lean_Linter_MissingDocs_missingDocs___closed__2 = _init_l_Lean_Linter_MissingDocs_missingDocs___closed__2(); +lean_mark_persistent(l_Lean_Linter_MissingDocs_missingDocs___closed__2); +l_Lean_Linter_MissingDocs_missingDocs___closed__3 = _init_l_Lean_Linter_MissingDocs_missingDocs___closed__3(); +lean_mark_persistent(l_Lean_Linter_MissingDocs_missingDocs___closed__3); +l_Lean_Linter_MissingDocs_missingDocs = _init_l_Lean_Linter_MissingDocs_missingDocs(); +lean_mark_persistent(l_Lean_Linter_MissingDocs_missingDocs); +res = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1178_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1___closed__1 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1___closed__1(); -lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1___closed__1); -l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1___closed__2 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1___closed__2(); -lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1___closed__2); -l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1___closed__3 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1___closed__3(); -lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1___closed__3); -l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1___closed__4 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1___closed__4(); -lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1___closed__4); -l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1___closed__5 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1___closed__5(); -lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1___closed__5); -l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1___closed__6 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1___closed__6(); -lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1___closed__6); -l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1___closed__7 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1___closed__7(); -lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1___closed__7); -l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1___closed__8 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1___closed__8(); -lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__1___closed__8); -l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__2___closed__1 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__2___closed__1(); -lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__2___closed__1); -l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__2___closed__2 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__2___closed__2(); -lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__2___closed__2); -l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__2___closed__3 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__2___closed__3(); -lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__2___closed__3); -l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__2___closed__4 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__2___closed__4(); -lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__2___closed__4); -l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__2___closed__5 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__2___closed__5(); -lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__2___closed__5); -l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__4___closed__1 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__4___closed__1(); -lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__4___closed__1); -l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__4___closed__2 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__4___closed__2(); -lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__4___closed__2); -l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__4___closed__3 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__4___closed__3(); -lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__4___closed__3); -l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__4___closed__4 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__4___closed__4(); -lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__4___closed__4); -l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__5___closed__1 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__5___closed__1(); -lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__5___closed__1); -l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__5___closed__2 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__5___closed__2(); -lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__5___closed__2); -l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__8___closed__1 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__8___closed__1(); -lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__8___closed__1); -l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__8___closed__2 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__8___closed__2(); -lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____lambda__8___closed__2); -l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__1 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__1(); -lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__1); -l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__2 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__2(); -lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__2); -l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__3 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__3(); -lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__3); -l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__4 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__4(); -lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__4); -l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__5 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__5(); -lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__5); -l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__6 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__6(); -lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__6); -l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__7 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__7(); -lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__7); -l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__8 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__8(); -lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__8); -l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__9 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__9(); -lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__9); -l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__10 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__10(); -lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__10); -l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__11 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__11(); -lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__11); -l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__12 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__12(); -lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__12); -l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__13 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__13(); -lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__13); -l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__14 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__14(); -lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__14); -l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__15 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__15(); -lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__15); -l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__16 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__16(); -lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__16); -l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__17 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__17(); -lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__17); -l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__18 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__18(); -lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__18); -l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__19 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__19(); -lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__19); -l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__20 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__20(); -lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__20); -l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__21 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__21(); -lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__21); -l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__22 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__22(); -lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__22); -l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__23 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__23(); -lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__23); -l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__24 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__24(); -lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220____closed__24); -res = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1220_(lean_io_mk_world()); +l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__1___closed__1 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__1___closed__1); +l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__1___closed__2 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__1___closed__2(); +lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__1___closed__2); +l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__1___closed__3 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__1___closed__3(); +lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__1___closed__3); +l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__1___closed__4 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__1___closed__4(); +lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__1___closed__4); +l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__1___closed__5 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__1___closed__5(); +lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__1___closed__5); +l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__1___closed__6 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__1___closed__6(); +lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__1___closed__6); +l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__1___closed__7 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__1___closed__7(); +lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__1___closed__7); +l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__1___closed__8 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__1___closed__8(); +lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__1___closed__8); +l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__2___closed__1 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__2___closed__1); +l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__2___closed__2 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__2___closed__2(); +lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__2___closed__2); +l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__2___closed__3 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__2___closed__3(); +lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__2___closed__3); +l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__2___closed__4 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__2___closed__4(); +lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__2___closed__4); +l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__2___closed__5 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__2___closed__5(); +lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__2___closed__5); +l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__4___closed__1 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__4___closed__1(); +lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__4___closed__1); +l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__4___closed__2 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__4___closed__2(); +lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__4___closed__2); +l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__4___closed__3 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__4___closed__3(); +lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__4___closed__3); +l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__4___closed__4 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__4___closed__4(); +lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__4___closed__4); +l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__5___closed__1 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__5___closed__1(); +lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__5___closed__1); +l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__5___closed__2 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__5___closed__2(); +lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__5___closed__2); +l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__8___closed__1 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__8___closed__1(); +lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__8___closed__1); +l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__8___closed__2 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__8___closed__2(); +lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____lambda__8___closed__2); +l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__1 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__1(); +lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__1); +l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__2 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__2(); +lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__2); +l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__3 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__3(); +lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__3); +l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__4 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__4(); +lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__4); +l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__5 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__5(); +lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__5); +l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__6 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__6(); +lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__6); +l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__7 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__7(); +lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__7); +l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__8 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__8(); +lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__8); +l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__9 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__9(); +lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__9); +l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__10 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__10(); +lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__10); +l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__11 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__11(); +lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__11); +l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__12 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__12(); +lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__12); +l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__13 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__13(); +lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__13); +l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__14 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__14(); +lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__14); +l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__15 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__15(); +lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__15); +l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__16 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__16(); +lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__16); +l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__17 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__17(); +lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__17); +l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__18 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__18(); +lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__18); +l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__19 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__19(); +lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__19); +l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__20 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__20(); +lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__20); +l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__21 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__21(); +lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__21); +l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__22 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__22(); +lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__22); +l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__23 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__23(); +lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__23); +l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__24 = _init_l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__24(); +lean_mark_persistent(l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231____closed__24); +res = l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_MissingDocs___hyg_1231_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); l_Lean_Linter_logLint___at_Lean_Linter_MissingDocs_lint___spec__1___closed__1 = _init_l_Lean_Linter_logLint___at_Lean_Linter_MissingDocs_lint___spec__1___closed__1(); diff --git a/stage0/stdlib/Lean/Linter/UnusedVariables.c b/stage0/stdlib/Lean/Linter/UnusedVariables.c index 6972ddc043b9..07b104b274ff 100644 --- a/stage0/stdlib/Lean/Linter/UnusedVariables.c +++ b/stage0/stdlib/Lean/Linter/UnusedVariables.c @@ -14,55 +14,55 @@ extern "C" { #endif static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_499____lambda__1___closed__8; -LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Linter_unusedVariables___spec__36___lambda__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Linter_unusedVariables___elambda__1___spec__36___lambda__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_357____closed__1; -static lean_object* l_Array_qsort_sort___at_Lean_Linter_unusedVariables___spec__36___closed__1; LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2367____lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_685____lambda__3___closed__2; +LEAN_EXPORT lean_object* l_Lean_Linter_unusedVariables___elambda__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_7____closed__5; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__38___closed__1; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_1871____lambda__1___closed__9; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_89____closed__3; -static lean_object* l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___spec__37___closed__1; -LEAN_EXPORT lean_object* l_Lean_mkHashSet___at_Lean_Linter_unusedVariables___spec__2(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Linter_unusedVariables___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_foldr___at_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2206____spec__1___closed__2; lean_object* lean_mk_empty_array_with_capacity(lean_object*); uint8_t l_Lean_HashSetImp_contains___at_Lean_MVarId_getNondepPropHyps___spec__19(lean_object*, lean_object*); +static lean_object* l_Lean_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___elambda__1___spec__25___lambda__2___closed__1; LEAN_EXPORT uint8_t l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_685____lambda__5(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_50____closed__1; LEAN_EXPORT uint8_t l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2038____lambda__1(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_matchesIdent(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35___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*); +static lean_object* l_Array_qsort_sort___at_Lean_Linter_unusedVariables___elambda__1___spec__36___closed__1; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Linter_unusedVariables___elambda__1___spec__16(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2038____lambda__1___closed__3; +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_foldr___at_Lean_Linter_unusedVariables___elambda__1___spec__31___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2300____closed__5; -static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_5176____closed__1; LEAN_EXPORT lean_object* l_List_foldr___at_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_395____spec__1___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Linter_unusedVariables___elambda__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentArray_toArray___rarg(lean_object*); -static lean_object* l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___spec__27___closed__3; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_685____lambda__5___closed__10; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_395____lambda__1___closed__14; +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___elambda__1___spec__32(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_395____lambda__1___closed__5; lean_object* l_Lean_ConstantInfo_type(lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_7____closed__6; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2146____lambda__1___closed__7; LEAN_EXPORT lean_object* l_Lean_Linter_getUnusedVariablesIgnoreFnsImpl(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_685____spec__3(lean_object*, size_t, size_t); +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___elambda__1___spec__42___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t l___private_Lean_Data_HashSet_0__Lean_HashSetImp_mkIdx(lean_object*, uint64_t, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_1907____lambda__1___closed__3; -LEAN_EXPORT lean_object* l_Lean_HashMapImp_moveEntries___at_Lean_Linter_unusedVariables___spec__11(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___spec__39___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___spec__32(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_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2206____lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_7_(lean_object*); LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_685____spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35___lambda__2(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_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Linter_unusedVariables___elambda__1___spec__41___lambda__2(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_7____closed__8; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2038____lambda__1___closed__1; lean_object* l_Lean_Name_toString(lean_object*, uint8_t); +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___elambda__1___spec__30___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_499____closed__1; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2146____lambda__1___closed__2; lean_object* l_Lean_Syntax_getId(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_5176_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2367____lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_1871____lambda__1___closed__4; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2367____closed__9; @@ -71,92 +71,88 @@ static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hy LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_685____lambda__2___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_685____lambda__5___closed__7; -LEAN_EXPORT lean_object* l_Lean_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___spec__25___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___elambda__1___spec__27___closed__1; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2146____lambda__1___closed__9; lean_object* l_Array_qpartition___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___elambda__1___spec__37___closed__4; lean_object* l_Lean_mkHashSetImp___rarg(lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_395____closed__1; -static lean_object* l_Array_qsort_sort___at_Lean_Linter_unusedVariables___spec__36___lambda__1___closed__1; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_357____lambda__1___closed__7; LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2300____lambda__1(lean_object*, lean_object*); lean_object* lean_mk_array(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_HashSetImp_expand___at_Lean_Linter_unusedVariables___spec__4(lean_object*, lean_object*); static lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_685____spec__3___closed__2; uint8_t lean_usize_dec_eq(size_t, size_t); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_685____lambda__5___closed__11; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_592____lambda__1___closed__10; -static lean_object* l_panic___at_Lean_Linter_unusedVariables___spec__28___closed__1; lean_object* l_Lean_Syntax_getArgs(lean_object*); -LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_checked___at_Lean_Linter_unusedVariables___spec__20___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_685____lambda__5___closed__9; +LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visited___at_Lean_Linter_unusedVariables___elambda__1___spec__19___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_String_Range_includes(lean_object*, lean_object*); +static lean_object* l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___elambda__1___spec__37___closed__1; +LEAN_EXPORT lean_object* l_Lean_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___elambda__1___spec__25___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2206____lambda__1(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getPos_x3f(lean_object*, uint8_t); lean_object* lean_array_fget(lean_object*, lean_object*); lean_object* l_Lean_Elab_InfoTree_hasSorry(lean_object*, lean_object*); extern lean_object* l_Lean_ForEachExprWhere_initCache; -LEAN_EXPORT lean_object* l_Lean_mkHashSet___at_Lean_Linter_unusedVariables___spec__2___boxed(lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35(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_array_fset(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Linter_unusedVariables___spec__23___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2367____lambda__2___closed__4; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_685____lambda__1___closed__4; LEAN_EXPORT lean_object* l_Lean_Linter_unusedVariables_skipDeclIdIfPresent(lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_284____lambda__1___closed__1; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_1871____closed__1; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2367____lambda__4___closed__2; +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35___lambda__4(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*); lean_object* l_Lean_Expr_fvarId_x21(lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2367____closed__17; -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___spec__30___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l_Lean_AssocList_contains___at_Lean_Linter_unusedVariables___spec__9(lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_357____lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Linter_getLinterUnusedVariablesFunArgs___boxed(lean_object*); +LEAN_EXPORT uint8_t l_Array_qsort_sort___at_Lean_Linter_unusedVariables___elambda__1___spec__36___lambda__1(lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_50____closed__4; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_685____lambda__5___closed__4; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_1871____lambda__1___closed__2; lean_object* l_Lean_Attribute_Builtin_ensureNoArgs(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Linter_unusedVariables___spec__6(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___spec__32___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_qsort_sort___at_Lean_Linter_unusedVariables___spec__36___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Linter_unusedVariables___elambda__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_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2300____lambda__2___boxed(lean_object*); -LEAN_EXPORT lean_object* l_List_foldr___at_Lean_Linter_unusedVariables___spec__31___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_7____closed__2; -LEAN_EXPORT uint8_t l_List_foldr___at_Lean_Linter_unusedVariables___spec__31___lambda__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_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___spec__27___closed__1; +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35___lambda__2(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*); uint8_t l_Lean_Linter_getLinterValue(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldr___at_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2038____spec__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_685____lambda__4___closed__3; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_7____closed__9; -LEAN_EXPORT lean_object* l_Lean_HashSetImp_moveEntries___at_Lean_Linter_unusedVariables___spec__5(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_685____lambda__5___closed__13; +LEAN_EXPORT uint8_t l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35___lambda__3(lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_1907____lambda__1___closed__4; lean_object* l_Lean_Elab_InfoTree_foldInfo___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_50____closed__5; -LEAN_EXPORT uint8_t l_Array_qsort_sort___at_Lean_Linter_unusedVariables___spec__36___lambda__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Linter_unusedVariables___elambda__1___spec__40___boxed(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_object* l_Lean_Name_mkStr5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_HashSetImp_insert___at_Lean_CollectLevelParams_visitExpr___spec__3(lean_object*, lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_395____lambda__1___closed__4; -LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___spec__39(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_592____lambda__1___closed__1; lean_object* l_Lean_PersistentArray_toList___rarg(lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_395____lambda__1___closed__6; -static lean_object* l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___spec__37___closed__4; -static lean_object* l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___spec__39___closed__1; -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35___lambda__6(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*, lean_object*, lean_object*); lean_object* l_List_head_x21___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35___lambda__5(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_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_685____lambda__5___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_312____lambda__1___closed__7; +LEAN_EXPORT lean_object* l_Lean_mkHashSet___at_Lean_Linter_unusedVariables___elambda__1___spec__2(lean_object*); +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___elambda__1___spec__34___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2367____closed__3; -static lean_object* l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___spec__37___closed__2; +static lean_object* l_Array_qsort_sort___at_Lean_Linter_unusedVariables___elambda__1___spec__36___lambda__1___closed__1; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_685____lambda__5___closed__3; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2206____closed__1; LEAN_EXPORT uint8_t l_Lean_Linter_getLinterUnusedVariablesFunArgs(lean_object*); LEAN_EXPORT lean_object* l_Array_sequenceMap___at_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_685____spec__1___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_HashMap_insert___at_Lean_Linter_unusedVariables___spec__8(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___spec__33(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Linter_unusedVariables___spec__40(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_replace___at_Lean_Linter_unusedVariables___elambda__1___spec__7___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Linter_unusedVariables___elambda__1___spec__41___lambda__2___closed__1; +LEAN_EXPORT lean_object* l_Lean_HashMap_insert___at_Lean_Linter_unusedVariables___elambda__1___spec__8(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___elambda__1___spec__27___closed__3; +static lean_object* l_panic___at_Lean_Linter_unusedVariables___elambda__1___spec__28___closed__1; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_592____closed__1; LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_284_(lean_object*); LEAN_EXPORT lean_object* l_Lean_ofExcept___at_Lean_Linter_getUnusedVariablesIgnoreFnsImpl___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*); @@ -164,96 +160,105 @@ static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hy static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2300____closed__3; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_395____lambda__1___closed__13; LEAN_EXPORT uint8_t l_List_foldr___at_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_592____spec__1(lean_object*, uint8_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___elambda__1___spec__39(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___elambda__1___spec__39___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_685____lambda__4___closed__4; LEAN_EXPORT uint8_t l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_592____lambda__1(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___elambda__1___spec__27___closed__4; LEAN_EXPORT lean_object* l_Array_sequenceMap___at_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_685____spec__1(lean_object*, lean_object*); extern lean_object* l_instInhabitedPUnit; -LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Linter_unusedVariables___spec__41(lean_object*, size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Linter_unusedVariables___spec__36(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit_go___at_Lean_Linter_unusedVariables___elambda__1___spec__18___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_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_5191_(lean_object*); LEAN_EXPORT lean_object* l_Lean_evalConstCheck___at_Lean_Linter_getUnusedVariablesIgnoreFnsImpl___spec__1(lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_357____lambda__1___closed__5; size_t lean_ptr_addr(lean_object*); lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); static lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_685____spec__3___closed__1; lean_object* l_Lean_Server_findModuleRefs(lean_object*, lean_object*, uint8_t, uint8_t); -LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Linter_unusedVariables___spec__29(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___elambda__1___spec__25(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_HashSetImp_expand___at_Lean_Linter_unusedVariables___elambda__1___spec__4(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashSetImp_contains___at_Lean_Linter_unusedVariables_isTopLevelDecl___spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_592____lambda__1___closed__7; size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT uint8_t l_Lean_Linter_getLinterUnusedVariablesPatternVars(lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_1871____lambda__1___closed__8; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_89____closed__2; +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35___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*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_312____lambda__1___closed__12; LEAN_EXPORT lean_object* l_List_elem___at_Lean_Linter_unusedVariables_isTopLevelDecl___spec__2___boxed(lean_object*, lean_object*); lean_object* l_panic___at_Lean_Elab_InfoTree_smallestInfo_x3f___spec__1(lean_object*); -LEAN_EXPORT lean_object* l_List_replace___at_Lean_Linter_unusedVariables___spec__7(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_AssocList_replace___at_Lean_Linter_unusedVariables___elambda__1___spec__13(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_685____spec__3___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___elambda__1___spec__25___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_499____lambda__1___closed__4; lean_object* l_List_filterMap___at_Lean_Linter_collectMacroExpansions_x3f_go___spec__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Linter_getUnusedVariablesIgnoreFnsImpl___spec__3(lean_object*); -LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___spec__39___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_unusedVariablesIgnoreFnsExt; -static lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Linter_unusedVariables___spec__41___lambda__2___closed__1; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_395____lambda__1___closed__2; lean_object* lean_local_ctx_find(lean_object*, lean_object*); -static lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35___lambda__4___closed__1; LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_685____spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); +static lean_object* l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___elambda__1___spec__37___closed__2; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visited___at_Lean_Linter_unusedVariables___elambda__1___spec__19(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_685____closed__1; +static lean_object* l_Lean_Linter_unusedVariables___elambda__1___lambda__1___closed__1; uint8_t l_Lean_AssocList_contains___at_Lean_Meta_AbstractMVars_abstractExprMVars___spec__5(lean_object*, lean_object*); +static lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Linter_unusedVariables___elambda__1___spec__41___closed__1; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_592____lambda__1___closed__5; lean_object* l_Lean_HashMap_insert___at_Lean_Meta_AbstractMVars_abstractExprMVars___spec__4(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Linter_unusedVariables___elambda__1___spec__23(lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_HashMap_toList___at_Lean_Linter_unusedVariables___elambda__1___spec__21(lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2367____lambda__2___closed__3; LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Linter_getUnusedVariablesIgnoreFnsImpl___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_qsort_sort___at_Lean_Linter_unusedVariables___spec__36___lambda__1___closed__2; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2146____lambda__1___closed__4; lean_object* l_Lean_Elab_Info_updateContext_x3f(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_HashSetImp_contains___at_Lean_Linter_unusedVariables_isTopLevelDecl___spec__1(lean_object*, lean_object*); lean_object* l_Lean_PersistentEnvExtension_addEntry___rarg(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___spec__25___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_Stack_matches(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Linter_unusedVariables___elambda__1___spec__41___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Linter_unusedVariables___elambda__1___spec__36(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_499____lambda__1___closed__10; -LEAN_EXPORT lean_object* l_Lean_Linter_unusedVariables___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___elambda__1___spec__27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___elambda__1___spec__25___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_linter_unusedVariables_patternVars; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2367____closed__7; lean_object* l_Lean_Elab_Command_getRef(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_685____lambda__1(lean_object*); LEAN_EXPORT uint8_t l_List_foldr___at_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_395____spec__1(lean_object*, uint8_t, lean_object*); +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___elambda__1___spec__32___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_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_395____lambda__1___closed__10; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_499____lambda__1___closed__11; +static lean_object* l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___elambda__1___spec__27___closed__2; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_395____lambda__1___closed__3; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_1871____lambda__1___closed__1; -LEAN_EXPORT lean_object* l_Lean_HashMapImp_expand___at_Lean_Linter_unusedVariables___spec__10(lean_object*, lean_object*); lean_object* l_Lean_Elab_getBetterRef(lean_object*, lean_object*); static lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_685____spec__3___closed__6; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_7____closed__7; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2367____closed__6; lean_object* l_Lean_setEnv___at_Lean_registerParametricAttribute___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Linter_collectMacroExpansions_x3f___at_Lean_Linter_unusedVariables___spec__24(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2367____closed__4; LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2146____lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_357____lambda__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_312____lambda__1___closed__11; -LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___spec__27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_685____lambda__5___closed__8; +LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___elambda__1___spec__39___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_395____lambda__1___closed__7; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_499____lambda__1___closed__3; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___spec__38___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___elambda__1___spec__25___closed__1; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_7____closed__1; 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*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2146____lambda__1___closed__6; -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35___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_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35___lambda__4___closed__1; lean_object* lean_st_ref_get(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Linter_unusedVariables___spec__41___lambda__1(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_throwError___at_Lean_registerTagAttribute___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_MetavarContext_getExprAssignmentDomain___spec__2___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_685____lambda__5___closed__6; -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_312____lambda__1___closed__6; LEAN_EXPORT uint8_t l_Lean_Linter_unusedVariables_isTopLevelDecl(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__38___closed__4; +static lean_object* l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___elambda__1___spec__37___closed__5; LEAN_EXPORT lean_object* l_Lean_ofExcept___at_Lean_Linter_getUnusedVariablesIgnoreFnsImpl___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_1907____lambda__1___closed__5; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2367____closed__20; @@ -265,49 +270,53 @@ lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); uint8_t l_String_startsWith(lean_object*, lean_object*); lean_object* l_instInhabited___rarg(lean_object*, lean_object*); lean_object* l_Lean_HashSetImp_insert___at_Lean_MVarId_getNondepPropHyps___spec__9(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_HashMapImp_moveEntries___at_Lean_Linter_unusedVariables___elambda__1___spec__11(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_312____lambda__1___closed__4; -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___spec__38___closed__4; +LEAN_EXPORT lean_object* l_Lean_HashMapImp_expand___at_Lean_Linter_unusedVariables___elambda__1___spec__10(lean_object*, lean_object*); uint8_t l___private_Lean_Syntax_0__String_beqRange____x40_Lean_Syntax___hyg_109_(lean_object*, lean_object*); -static lean_object* l_Lean_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___spec__25___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_685____lambda__4___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_1907_(lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_592____lambda__1___closed__8; -static lean_object* l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___spec__37___closed__5; +LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit___at_Lean_Linter_unusedVariables___elambda__1___spec__17(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_unusedVariables_isTopLevelDecl___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_685____lambda__3___closed__1; LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2367____lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_checked___at_Lean_Linter_unusedVariables___elambda__1___spec__20(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_312____lambda__1___closed__1; -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35___lambda__6___boxed(lean_object**); -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___spec__38(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35___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 uint8_t l_List_foldr___at_Lean_Linter_unusedVariables___elambda__1___spec__31___lambda__2(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_List_foldr___at_Lean_Linter_unusedVariables___elambda__1___spec__31___lambda__1(lean_object*, lean_object*); +static lean_object* l_Array_qsort_sort___at_Lean_Linter_unusedVariables___elambda__1___spec__36___lambda__1___closed__3; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_592____lambda__1___closed__11; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2146____lambda__1___closed__3; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_312____lambda__1___closed__8; +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Linter_unusedVariables___elambda__1___spec__41(lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_1907____closed__1; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_89____closed__4; +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35___lambda__6(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*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_1871____lambda__1___closed__5; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_89____closed__5; extern lean_object* l_Lean_Elab_Command_instMonadCommandElabM; +LEAN_EXPORT lean_object* l_List_foldr___at_Lean_Linter_unusedVariables___elambda__1___spec__31___lambda__2___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_mkHashSet___at_Lean_Linter_unusedVariables___elambda__1___spec__2___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2038____lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_1907____lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ofExcept___at_Lean_Linter_getUnusedVariablesIgnoreFnsImpl___spec__2(lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Lean_HashMapImp_contains___at_Lean_Linter_unusedVariables___elambda__1___spec__15(lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2367____closed__11; -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___spec__42___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35___lambda__3___boxed(lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35___lambda__5(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_EXPORT uint8_t l_Lean_AssocList_contains___at_Lean_Linter_unusedVariables___elambda__1___spec__9(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_MessageData_isUnusedVariableWarning___lambda__1(lean_object*); -LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visited___at_Lean_Linter_unusedVariables___spec__19___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldr___at_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_499____spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Linter_getUnusedVariablesIgnoreFnsImpl___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_685____lambda__1___closed__5; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Linter_unusedVariables___elambda__1___spec__40(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2367____lambda__4___closed__1; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_592____lambda__1___closed__3; LEAN_EXPORT uint8_t l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_685____lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_foldr___at_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2206____spec__1___closed__5; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Linter_unusedVariables___spec__16(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_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___spec__34___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_addBuiltinUnusedVariablesIgnoreFn(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_List_elem___at_Lean_Linter_unusedVariables_isTopLevelDecl___spec__2(lean_object*, lean_object*); @@ -318,118 +327,108 @@ static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hy uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); static lean_object* l_Lean_Linter_addBuiltinUnusedVariablesIgnoreFn___closed__1; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_1871____lambda__1___closed__3; -LEAN_EXPORT lean_object* l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___spec__37___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_builtinUnusedVariablesIgnoreFnsRef; LEAN_EXPORT uint8_t l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_685____lambda__4(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Linter_unusedVariables___spec__23(lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_7____closed__4; uint8_t l_Lean_Syntax_isIdent(lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_1871____lambda__1___closed__7; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_50____closed__2; LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_1871_(lean_object*); -static lean_object* l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___spec__27___closed__2; +LEAN_EXPORT lean_object* l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___elambda__1___spec__37(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_592____lambda__1___closed__6; static lean_object* l_List_foldr___at_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2206____spec__1___closed__4; -LEAN_EXPORT lean_object* l_Lean_Linter_unusedVariables___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2367____closed__19; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_685____lambda__1___closed__2; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_357____lambda__1___closed__4; -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35___lambda__3___boxed(lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2038____lambda__1___closed__2; lean_object* l_Lean_SimplePersistentEnvExtension_getEntries___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_499____lambda__1___closed__9; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_395____lambda__1___closed__9; static lean_object* l_Lean_Linter_getLinterUnusedVariables___closed__1; -static lean_object* l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___spec__27___closed__4; -static lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Linter_unusedVariables___spec__41___closed__1; LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_312_(lean_object*); -static lean_object* l_Array_qsort_sort___at_Lean_Linter_unusedVariables___spec__36___lambda__1___closed__3; +LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___elambda__1___spec__22(lean_object*, lean_object*); size_t lean_usize_mod(size_t, size_t); -LEAN_EXPORT lean_object* l_Lean_HashMapImp_contains___at_Lean_Linter_unusedVariables___spec__15___boxed(lean_object*, lean_object*); lean_object* l_Lean_Elab_Command_addLinter(lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_1907____lambda__1___closed__1; size_t lean_data_hashmap_mk_idx(lean_object*, uint64_t); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_162____closed__1; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_312____lambda__1___closed__10; +LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Linter_unusedVariables___elambda__1___spec__6(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_685____lambda__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_592____lambda__1___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35___lambda__4(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_Linter_unusedVariables___lambda__1___closed__1; +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__38___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__14(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2367_(lean_object*); LEAN_EXPORT uint8_t l_Lean_MessageData_isUnusedVariableWarning(lean_object*); uint8_t l_Lean_PersistentArray_anyM___at_Lean_MessageLog_hasErrors___spec__1(lean_object*); -LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Linter_unusedVariables___spec__41___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_LocalDecl_userName(lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_1907____lambda__1___closed__7; static lean_object* l_Lean_Linter_unusedVariables_skipDeclIdIfPresent___closed__2; -LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Linter_unusedVariables___spec__41___lambda__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_50_(lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_395____lambda__1___closed__8; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_357____lambda__1___closed__6; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_685____lambda__5___closed__12; -LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_checked___at_Lean_Linter_unusedVariables___spec__20(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2300____lambda__1___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit_go___at_Lean_Linter_unusedVariables___spec__18___lambda__1(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_Lean_Linter_unusedVariables___spec__31(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_checked___at_Lean_Linter_unusedVariables___elambda__1___spec__20___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_357_(lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_685____lambda__4___closed__1; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2146____lambda__1___closed__8; +LEAN_EXPORT uint8_t l_List_foldr___at_Lean_Linter_unusedVariables___elambda__1___spec__31(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_1907____lambda__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2367____closed__12; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2367____lambda__3___closed__2; uint8_t l_Lean_Name_hasMacroScopes(lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_89_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_AssocList_contains___at_Lean_Linter_unusedVariables___spec__9___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2367____closed__15; lean_object* l_Lean_throwError___at_Lean_AttributeImpl_erase___default___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___spec__38___closed__1; static lean_object* l_List_foldr___at_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2206____spec__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2300____lambda__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_linter_unusedVariables; uint8_t l_String_Range_contains(lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2206_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Linter_unusedVariables(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Linter_unusedVariables; +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2367____lambda__3(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_592_(lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_499____lambda__1___closed__12; -LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visited___at_Lean_Linter_unusedVariables___spec__19(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_162_(lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2146____lambda__1___closed__1; lean_object* l_Lean_Expr_isFVar___boxed(lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__38___closed__3; +static lean_object* l_Array_qsort_sort___at_Lean_Linter_unusedVariables___elambda__1___spec__36___lambda__1___closed__4; +LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___elambda__1___spec__22___boxed(lean_object*, lean_object*); uint8_t l_Lean_Syntax_hasArgs(lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_685____lambda__1___closed__3; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_89____closed__1; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Linter_getLinterUnusedVariables(lean_object*); -LEAN_EXPORT lean_object* l_Lean_AssocList_replace___at_Lean_Linter_unusedVariables___spec__13(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__38(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2367____lambda__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit_go___at_Lean_Linter_unusedVariables___elambda__1___spec__18(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_findStack_x3f(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MessageData_isUnusedVariableWarning___lambda__1___boxed(lean_object*); +static lean_object* l_Lean_Linter_unusedVariables___elambda__1___lambda__1___closed__2; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_395____lambda__1___closed__12; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l_Lean_MessageData_isUnusedVariableWarning___closed__1; -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__14(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___spec__22___boxed(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___elambda__1___spec__33(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t); +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35___lambda__6___boxed(lean_object**); static lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_685____spec__3___closed__3; LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_685_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_312____lambda__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___elambda__1___spec__39___closed__1; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2300_(lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_312____lambda__1___closed__9; static lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_685____spec__3___closed__4; -LEAN_EXPORT lean_object* l_Lean_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___spec__25(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_HashSetImp_insert___at_Lean_Linter_unusedVariables___spec__3(lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_312____closed__1; +LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_Linter_unusedVariables___elambda__1___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2367____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___spec__38___closed__3; -LEAN_EXPORT lean_object* l_panic___at_Lean_Linter_unusedVariables___spec__28(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_1907____lambda__1___closed__2; -LEAN_EXPORT lean_object* l_List_foldr___at_Lean_Linter_unusedVariables___spec__31___lambda__1___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___elambda__1___spec__39___closed__2; LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2146_(lean_object*); uint8_t l_Lean_Expr_isConstOf(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___spec__37(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_Linter_unusedVariables___spec__1___boxed(lean_object*); +static lean_object* l_Lean_Linter_unusedVariables___closed__2; LEAN_EXPORT lean_object* l_Lean_MessageData_isUnusedVariableWarning___boxed(lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_7____closed__3; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_499____lambda__1___closed__1; @@ -437,13 +436,10 @@ LEAN_EXPORT uint8_t l_List_foldr___at_Lean_Linter_initFn____x40_Lean_Linter_Unus uint8_t l_Lean_Syntax_isNone(lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); static lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_685____spec__3___closed__5; -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Linter_unusedVariables___spec__40___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Linter_unusedVariables___elambda__1___spec__41___lambda__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_592____lambda__1___closed__2; -LEAN_EXPORT uint8_t l_Lean_HashMapImp_contains___at_Lean_Linter_unusedVariables___spec__15(lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_HashMap_toList___at_Lean_Linter_unusedVariables___spec__21(lean_object*); lean_object* l_Array_insertAt_x21___rarg(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35(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*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_685____lambda__5___closed__2; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_312____lambda__1___closed__3; lean_object* lean_nat_mul(lean_object*, lean_object*); @@ -454,43 +450,51 @@ static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hy static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2300____closed__4; static lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_685____spec__3___closed__7; LEAN_EXPORT uint8_t l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_395____lambda__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_Linter_unusedVariables___elambda__1___spec__1___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_HashMapImp_contains___at_Lean_Linter_unusedVariables___elambda__1___spec__15___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Linter_getUnusedVariablesIgnoreFnsImpl___spec__3___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___spec__25___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Linter_unusedVariables___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_Environment_evalConstCheck___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2038____closed__1; -static lean_object* l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___spec__37___closed__3; +LEAN_EXPORT lean_object* l_panic___at_Lean_Linter_unusedVariables___elambda__1___spec__28(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_foldl___at_Array_appendList___spec__1___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2300____closed__6; -LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_visitM___at_Lean_Linter_unusedVariables___spec__26(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_qsort_sort___at_Lean_Linter_unusedVariables___elambda__1___spec__36___lambda__1___closed__2; lean_object* l_Lean_registerBuiltinAttribute(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___elambda__1___spec__33___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_395____lambda__1___closed__11; +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35___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_List_reverse___rarg(lean_object*); +static lean_object* l_Lean_Linter_unusedVariables___closed__3; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2300____closed__2; size_t lean_usize_sub(size_t, size_t); extern lean_object* l_Lean_Elab_Command_instInhabitedScope; +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_499____lambda__1___closed__6; lean_object* l_Lean_mkHashMapImp___rarg(lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_1871____lambda__1___closed__6; +LEAN_EXPORT lean_object* l_Lean_Linter_collectMacroExpansions_x3f___at_Lean_Linter_unusedVariables___elambda__1___spec__24(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldr___at_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2206____spec__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35___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*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2367____lambda__2___closed__2; +LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___elambda__1___spec__12(lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2367____closed__10; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_1907____lambda__1___closed__6; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2367____closed__14; uint64_t l___private_Lean_Expr_0__Lean_hashMVarId____x40_Lean_Expr___hyg_2302_(lean_object*); +LEAN_EXPORT lean_object* l_List_foldr___at_Lean_Linter_unusedVariables___elambda__1___spec__31___lambda__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2367____lambda__3___closed__1; LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_685____lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___spec__30(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t); +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___elambda__1___spec__34(lean_object*, lean_object*, size_t, size_t); lean_object* l_Lean_KVMap_findCore(lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); lean_object* l_Lean_Elab_Info_lctx(lean_object*); +LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__29(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2146____lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldr___at_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_592____spec__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_685____lambda__4___closed__2; static lean_object* l_Lean_Linter_unusedVariables_skipDeclIdIfPresent___closed__1; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_357____lambda__1___closed__1; -LEAN_EXPORT uint8_t l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35___lambda__3(lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_50____closed__3; lean_object* lean_io_error_to_string(lean_object*); LEAN_EXPORT lean_object* l_Lean_evalConstCheck___at_Lean_Linter_getUnusedVariablesIgnoreFnsImpl___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -503,23 +507,28 @@ static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hy lean_object* l_Lean_Name_mkStr4(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_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2367____lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_HashSetImp_moveEntries___at_Lean_Linter_unusedVariables___elambda__1___spec__5(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_685____lambda__5___closed__5; +LEAN_EXPORT lean_object* l_Lean_HashSetImp_insert___at_Lean_Linter_unusedVariables___elambda__1___spec__3(lean_object*, lean_object*); lean_object* l_Lean_HashMap_toList___at_Lean_Server_ModuleRefs_instCoeModuleRefsModuleRefs___spec__1(lean_object*); -static lean_object* l_Array_qsort_sort___at_Lean_Linter_unusedVariables___spec__36___lambda__1___closed__4; -static lean_object* l_Lean_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___spec__25___closed__1; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__38___closed__2; uint8_t l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_367_(uint8_t, uint8_t); LEAN_EXPORT uint8_t l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_357____lambda__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_357____lambda__1___closed__3; lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_SepArray_getElems___spec__1(lean_object*, size_t, size_t, lean_object*); lean_object* lean_array_get_size(lean_object*); -LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit_go___at_Lean_Linter_unusedVariables___spec__18(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_1871____lambda__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___elambda__1___spec__37___closed__3; lean_object* l_Lean_getConstInfo___at_Lean_Linter_initFn____x40_Lean_Linter_Deprecated___hyg_47____spec__14(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_List_foldr___at_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2206____spec__1(uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_395_(lean_object*); -LEAN_EXPORT uint8_t l_List_foldr___at_Lean_Linter_unusedVariables___spec__31___lambda__2(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_visitM___at_Lean_Linter_unusedVariables___elambda__1___spec__26(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Linter_unusedVariables___elambda__1___spec__23___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Linter_unusedVariables___closed__1; uint8_t lean_nat_dec_le(lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2367____closed__13; +LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Linter_unusedVariables___elambda__1___spec__36___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___elambda__1___spec__37___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_filterMap___at_Lean_Linter_collectMacroExpansions_x3f_go___spec__1(lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_312____lambda__1___closed__5; LEAN_EXPORT uint8_t l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_312____lambda__1(lean_object*, lean_object*, lean_object*); @@ -527,50 +536,44 @@ uint8_t lean_usize_dec_lt(size_t, size_t); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2367____closed__1; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_592____lambda__1___closed__9; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_284____closed__1; -LEAN_EXPORT lean_object* l_List_foldr___at_Lean_Linter_unusedVariables___spec__31___lambda__2___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_replace___at_Lean_Linter_unusedVariables___spec__7___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___spec__25___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_1871____lambda__1(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___spec__42(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_499____lambda__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_312____lambda__1___closed__2; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2367____closed__5; +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35___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_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2367____closed__8; -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___spec__33___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_List_foldr___at_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2038____spec__1(lean_object*, uint8_t, lean_object*); uint64_t l___private_Lean_Syntax_0__String_hashRange____x40_Lean_Syntax___hyg_202_(lean_object*); -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___spec__34(lean_object*, lean_object*, size_t, size_t); uint8_t l_Lean_HashSetImp_contains___at_Lean_CollectLevelParams_visitExpr___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_499____lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_284____lambda__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_AssocList_contains___at_Lean_Linter_unusedVariables___elambda__1___spec__9___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2367____lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2146____lambda__1___closed__5; +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___elambda__1___spec__30(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t); +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___elambda__1___spec__42(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2300____closed__1; LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2367____lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_499____lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Linter_linter_unusedVariables_funArgs; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); static lean_object* l_List_foldr___at_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2206____spec__1___closed__1; -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___spec__38___closed__2; lean_object* l_Lean_MessageData_ofName(lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_395____lambda__1___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Linter_unusedVariables___lambda__1___closed__2; +LEAN_EXPORT lean_object* l_Lean_Linter_unusedVariables___elambda__1___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2367____closed__18; -LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___spec__22(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35___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*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_1907____lambda__1___closed__8; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2367____lambda__2___closed__1; lean_object* l_List_get_x3f___rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit___at_Lean_Linter_unusedVariables___spec__17(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_replace___at_Lean_Linter_unusedVariables___elambda__1___spec__7(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_395____lambda__1___closed__1; lean_object* l___private_Lean_Data_HashMap_0__Lean_numBucketsForCapacity(lean_object*); -LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___spec__12(lean_object*, lean_object*); uint8_t l_Lean_Elab_Info_contains(lean_object*, lean_object*, uint8_t); -static lean_object* l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___spec__39___closed__2; +LEAN_EXPORT lean_object* l_Lean_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___elambda__1___spec__25___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Linter_unusedVariables___elambda__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Array_isEmpty___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_499_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_Linter_unusedVariables___spec__1(lean_object*); static lean_object* _init_l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_7____closed__1() { _start: { @@ -5251,7 +5254,7 @@ x_6 = lean_box(x_5); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_Linter_unusedVariables___spec__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_Linter_unusedVariables___elambda__1___spec__1(lean_object* x_1) { _start: { lean_object* x_2; @@ -5259,7 +5262,7 @@ x_2 = l_Lean_mkHashMapImp___rarg(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_mkHashSet___at_Lean_Linter_unusedVariables___spec__2(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_mkHashSet___at_Lean_Linter_unusedVariables___elambda__1___spec__2(lean_object* x_1) { _start: { lean_object* x_2; @@ -5267,7 +5270,7 @@ x_2 = l_Lean_mkHashSetImp___rarg(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Linter_unusedVariables___spec__6(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Linter_unusedVariables___elambda__1___spec__6(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -5316,7 +5319,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_HashSetImp_moveEntries___at_Lean_Linter_unusedVariables___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_HashSetImp_moveEntries___at_Lean_Linter_unusedVariables___elambda__1___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; uint8_t x_5; @@ -5335,7 +5338,7 @@ lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_obj x_6 = lean_array_fget(x_2, x_1); x_7 = lean_box(0); x_8 = lean_array_fset(x_2, x_1, x_7); -x_9 = l_List_foldl___at_Lean_Linter_unusedVariables___spec__6(x_3, x_6); +x_9 = l_List_foldl___at_Lean_Linter_unusedVariables___elambda__1___spec__6(x_3, x_6); x_10 = lean_unsigned_to_nat(1u); x_11 = lean_nat_add(x_1, x_10); lean_dec(x_1); @@ -5346,7 +5349,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_HashSetImp_expand___at_Lean_Linter_unusedVariables___spec__4(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_HashSetImp_expand___at_Lean_Linter_unusedVariables___elambda__1___spec__4(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; @@ -5357,14 +5360,14 @@ lean_dec(x_3); x_6 = lean_box(0); x_7 = lean_mk_array(x_5, x_6); x_8 = lean_unsigned_to_nat(0u); -x_9 = l_Lean_HashSetImp_moveEntries___at_Lean_Linter_unusedVariables___spec__5(x_8, x_2, x_7); +x_9 = l_Lean_HashSetImp_moveEntries___at_Lean_Linter_unusedVariables___elambda__1___spec__5(x_8, x_2, x_7); x_10 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_10, 0, x_1); lean_ctor_set(x_10, 1, x_9); return x_10; } } -LEAN_EXPORT lean_object* l_List_replace___at_Lean_Linter_unusedVariables___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_List_replace___at_Lean_Linter_unusedVariables___elambda__1___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_1) == 0) @@ -5387,7 +5390,7 @@ x_8 = l___private_Lean_Syntax_0__String_beqRange____x40_Lean_Syntax___hyg_109_(x if (x_8 == 0) { lean_object* x_9; -x_9 = l_List_replace___at_Lean_Linter_unusedVariables___spec__7(x_7, x_2, x_3); +x_9 = l_List_replace___at_Lean_Linter_unusedVariables___elambda__1___spec__7(x_7, x_2, x_3); lean_ctor_set(x_1, 1, x_9); return x_1; } @@ -5410,7 +5413,7 @@ x_12 = l___private_Lean_Syntax_0__String_beqRange____x40_Lean_Syntax___hyg_109_( if (x_12 == 0) { lean_object* x_13; lean_object* x_14; -x_13 = l_List_replace___at_Lean_Linter_unusedVariables___spec__7(x_11, x_2, x_3); +x_13 = l_List_replace___at_Lean_Linter_unusedVariables___elambda__1___spec__7(x_11, x_2, x_3); x_14 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_14, 0, x_10); lean_ctor_set(x_14, 1, x_13); @@ -5429,7 +5432,7 @@ return x_15; } } } -LEAN_EXPORT lean_object* l_Lean_HashSetImp_insert___at_Lean_Linter_unusedVariables___spec__3(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_HashSetImp_insert___at_Lean_Linter_unusedVariables___elambda__1___spec__3(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; @@ -5461,7 +5464,7 @@ if (x_15 == 0) { lean_object* x_16; lean_free_object(x_1); -x_16 = l_Lean_HashSetImp_expand___at_Lean_Linter_unusedVariables___spec__4(x_12, x_14); +x_16 = l_Lean_HashSetImp_expand___at_Lean_Linter_unusedVariables___elambda__1___spec__4(x_12, x_14); return x_16; } else @@ -5476,7 +5479,7 @@ else lean_object* x_17; lean_object* x_18; lean_dec(x_6); lean_inc(x_2); -x_17 = l_List_replace___at_Lean_Linter_unusedVariables___spec__7(x_9, x_2, x_2); +x_17 = l_List_replace___at_Lean_Linter_unusedVariables___elambda__1___spec__7(x_9, x_2, x_2); lean_dec(x_2); x_18 = lean_array_uset(x_5, x_8, x_17); lean_ctor_set(x_1, 1, x_18); @@ -5512,7 +5515,7 @@ lean_dec(x_21); if (x_30 == 0) { lean_object* x_31; -x_31 = l_Lean_HashSetImp_expand___at_Lean_Linter_unusedVariables___spec__4(x_27, x_29); +x_31 = l_Lean_HashSetImp_expand___at_Lean_Linter_unusedVariables___elambda__1___spec__4(x_27, x_29); return x_31; } else @@ -5529,7 +5532,7 @@ else lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_dec(x_21); lean_inc(x_2); -x_33 = l_List_replace___at_Lean_Linter_unusedVariables___spec__7(x_24, x_2, x_2); +x_33 = l_List_replace___at_Lean_Linter_unusedVariables___elambda__1___spec__7(x_24, x_2, x_2); lean_dec(x_2); x_34 = lean_array_uset(x_20, x_23, x_33); x_35 = lean_alloc_ctor(0, 2, 0); @@ -5540,7 +5543,7 @@ return x_35; } } } -LEAN_EXPORT uint8_t l_Lean_AssocList_contains___at_Lean_Linter_unusedVariables___spec__9(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l_Lean_AssocList_contains___at_Lean_Linter_unusedVariables___elambda__1___spec__9(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -5569,7 +5572,7 @@ return x_8; } } } -LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___spec__12(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___elambda__1___spec__12(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -5621,7 +5624,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_HashMapImp_moveEntries___at_Lean_Linter_unusedVariables___spec__11(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_HashMapImp_moveEntries___at_Lean_Linter_unusedVariables___elambda__1___spec__11(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; uint8_t x_5; @@ -5640,7 +5643,7 @@ lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_obj x_6 = lean_array_fget(x_2, x_1); x_7 = lean_box(0); x_8 = lean_array_fset(x_2, x_1, x_7); -x_9 = l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___spec__12(x_3, x_6); +x_9 = l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___elambda__1___spec__12(x_3, x_6); x_10 = lean_unsigned_to_nat(1u); x_11 = lean_nat_add(x_1, x_10); lean_dec(x_1); @@ -5651,7 +5654,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_HashMapImp_expand___at_Lean_Linter_unusedVariables___spec__10(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_HashMapImp_expand___at_Lean_Linter_unusedVariables___elambda__1___spec__10(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; @@ -5662,14 +5665,14 @@ lean_dec(x_3); x_6 = lean_box(0); x_7 = lean_mk_array(x_5, x_6); x_8 = lean_unsigned_to_nat(0u); -x_9 = l_Lean_HashMapImp_moveEntries___at_Lean_Linter_unusedVariables___spec__11(x_8, x_2, x_7); +x_9 = l_Lean_HashMapImp_moveEntries___at_Lean_Linter_unusedVariables___elambda__1___spec__11(x_8, x_2, x_7); x_10 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_10, 0, x_1); lean_ctor_set(x_10, 1, x_9); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_AssocList_replace___at_Lean_Linter_unusedVariables___spec__13(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_AssocList_replace___at_Lean_Linter_unusedVariables___elambda__1___spec__13(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_3) == 0) @@ -5694,7 +5697,7 @@ x_9 = lean_name_eq(x_6, x_1); if (x_9 == 0) { lean_object* x_10; -x_10 = l_Lean_AssocList_replace___at_Lean_Linter_unusedVariables___spec__13(x_1, x_2, x_8); +x_10 = l_Lean_AssocList_replace___at_Lean_Linter_unusedVariables___elambda__1___spec__13(x_1, x_2, x_8); lean_ctor_set(x_3, 2, x_10); return x_3; } @@ -5721,7 +5724,7 @@ x_14 = lean_name_eq(x_11, x_1); if (x_14 == 0) { lean_object* x_15; lean_object* x_16; -x_15 = l_Lean_AssocList_replace___at_Lean_Linter_unusedVariables___spec__13(x_1, x_2, x_13); +x_15 = l_Lean_AssocList_replace___at_Lean_Linter_unusedVariables___elambda__1___spec__13(x_1, x_2, x_13); x_16 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_16, 0, x_11); lean_ctor_set(x_16, 1, x_12); @@ -5743,7 +5746,7 @@ return x_17; } } } -LEAN_EXPORT lean_object* l_Lean_HashMap_insert___at_Lean_Linter_unusedVariables___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_HashMap_insert___at_Lean_Linter_unusedVariables___elambda__1___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -5758,7 +5761,7 @@ x_8 = l___private_Lean_Expr_0__Lean_hashFVarId____x40_Lean_Expr___hyg_1933_(x_2) lean_inc(x_7); x_9 = lean_data_hashmap_mk_idx(x_7, x_8); x_10 = lean_array_uget(x_6, x_9); -x_11 = l_Lean_AssocList_contains___at_Lean_Linter_unusedVariables___spec__9(x_2, x_10); +x_11 = l_Lean_AssocList_contains___at_Lean_Linter_unusedVariables___elambda__1___spec__9(x_2, x_10); if (x_11 == 0) { lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; @@ -5778,7 +5781,7 @@ if (x_17 == 0) { lean_object* x_18; lean_free_object(x_1); -x_18 = l_Lean_HashMapImp_expand___at_Lean_Linter_unusedVariables___spec__10(x_13, x_15); +x_18 = l_Lean_HashMapImp_expand___at_Lean_Linter_unusedVariables___elambda__1___spec__10(x_13, x_15); return x_18; } else @@ -5792,7 +5795,7 @@ else { lean_object* x_19; lean_object* x_20; lean_dec(x_7); -x_19 = l_Lean_AssocList_replace___at_Lean_Linter_unusedVariables___spec__13(x_2, x_3, x_10); +x_19 = l_Lean_AssocList_replace___at_Lean_Linter_unusedVariables___elambda__1___spec__13(x_2, x_3, x_10); x_20 = lean_array_uset(x_6, x_9, x_19); lean_ctor_set(x_1, 1, x_20); return x_1; @@ -5811,7 +5814,7 @@ x_24 = l___private_Lean_Expr_0__Lean_hashFVarId____x40_Lean_Expr___hyg_1933_(x_2 lean_inc(x_23); x_25 = lean_data_hashmap_mk_idx(x_23, x_24); x_26 = lean_array_uget(x_22, x_25); -x_27 = l_Lean_AssocList_contains___at_Lean_Linter_unusedVariables___spec__9(x_2, x_26); +x_27 = l_Lean_AssocList_contains___at_Lean_Linter_unusedVariables___elambda__1___spec__9(x_2, x_26); if (x_27 == 0) { lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; @@ -5830,7 +5833,7 @@ lean_dec(x_32); if (x_33 == 0) { lean_object* x_34; -x_34 = l_Lean_HashMapImp_expand___at_Lean_Linter_unusedVariables___spec__10(x_29, x_31); +x_34 = l_Lean_HashMapImp_expand___at_Lean_Linter_unusedVariables___elambda__1___spec__10(x_29, x_31); return x_34; } else @@ -5846,7 +5849,7 @@ else { lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_dec(x_23); -x_36 = l_Lean_AssocList_replace___at_Lean_Linter_unusedVariables___spec__13(x_2, x_3, x_26); +x_36 = l_Lean_AssocList_replace___at_Lean_Linter_unusedVariables___elambda__1___spec__13(x_2, x_3, x_26); x_37 = lean_array_uset(x_22, x_25, x_36); x_38 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_38, 0, x_21); @@ -5856,7 +5859,7 @@ return x_38; } } } -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__14(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_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__14(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { if (lean_obj_tag(x_1) == 0) @@ -5942,7 +5945,7 @@ 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 = l_Lean_HashSetImp_insert___at_Lean_Linter_unusedVariables___spec__3(x_20, x_26); +x_27 = l_Lean_HashSetImp_insert___at_Lean_Linter_unusedVariables___elambda__1___spec__3(x_20, x_26); lean_ctor_set(x_2, 0, x_27); x_1 = x_18; goto _start; @@ -5980,7 +5983,7 @@ lean_object* x_37; lean_object* x_38; lean_object* x_39; x_37 = lean_ctor_get(x_34, 0); lean_inc(x_37); lean_dec(x_34); -x_38 = l_Lean_HashSetImp_insert___at_Lean_Linter_unusedVariables___spec__3(x_29, x_37); +x_38 = l_Lean_HashSetImp_insert___at_Lean_Linter_unusedVariables___elambda__1___spec__3(x_29, x_37); x_39 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_39, 0, x_38); lean_ctor_set(x_39, 1, x_30); @@ -6008,7 +6011,7 @@ x_44 = lean_ctor_get(x_2, 1); x_45 = lean_ctor_get(x_8, 0); lean_inc(x_45); lean_dec(x_8); -x_46 = l_Lean_HashMap_insert___at_Lean_Linter_unusedVariables___spec__8(x_44, x_45, x_42); +x_46 = l_Lean_HashMap_insert___at_Lean_Linter_unusedVariables___elambda__1___spec__8(x_44, x_45, x_42); lean_ctor_set(x_2, 1, x_46); x_1 = x_41; goto _start; @@ -6024,7 +6027,7 @@ lean_dec(x_2); x_50 = lean_ctor_get(x_8, 0); lean_inc(x_50); lean_dec(x_8); -x_51 = l_Lean_HashMap_insert___at_Lean_Linter_unusedVariables___spec__8(x_49, x_50, x_42); +x_51 = l_Lean_HashMap_insert___at_Lean_Linter_unusedVariables___elambda__1___spec__8(x_49, x_50, x_42); x_52 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_52, 0, x_48); lean_ctor_set(x_52, 1, x_51); @@ -6036,7 +6039,7 @@ goto _start; } } } -LEAN_EXPORT uint8_t l_Lean_HashMapImp_contains___at_Lean_Linter_unusedVariables___spec__15(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l_Lean_HashMapImp_contains___at_Lean_Linter_unusedVariables___elambda__1___spec__15(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; uint64_t x_5; size_t x_6; lean_object* x_7; uint8_t x_8; @@ -6050,7 +6053,7 @@ lean_dec(x_7); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Linter_unusedVariables___spec__16(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Linter_unusedVariables___elambda__1___spec__16(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; @@ -6061,7 +6064,7 @@ x_5 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_MetavarContext_getExprAssignm return x_5; } } -LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visited___at_Lean_Linter_unusedVariables___spec__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_EXPORT lean_object* l_Lean_ForEachExprWhere_visited___at_Lean_Linter_unusedVariables___elambda__1___spec__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) { _start: { lean_object* x_7; uint8_t x_8; @@ -6273,7 +6276,7 @@ return x_76; } } } -LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_checked___at_Lean_Linter_unusedVariables___spec__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_EXPORT lean_object* l_Lean_ForEachExprWhere_checked___at_Lean_Linter_unusedVariables___elambda__1___spec__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) { _start: { lean_object* x_7; uint8_t x_8; @@ -6473,7 +6476,7 @@ return x_66; } } } -LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit_go___at_Lean_Linter_unusedVariables___spec__18___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_EXPORT lean_object* l_Lean_ForEachExprWhere_visit_go___at_Lean_Linter_unusedVariables___elambda__1___spec__18___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_dec(x_4); @@ -6491,7 +6494,7 @@ lean_inc(x_7); lean_inc(x_5); lean_inc(x_3); lean_inc(x_2); -x_12 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Linter_unusedVariables___spec__18(x_2, x_3, x_10, x_5, x_6, x_7, x_8, x_9); +x_12 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Linter_unusedVariables___elambda__1___spec__18(x_2, x_3, x_10, 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; lean_object* x_16; @@ -6503,7 +6506,7 @@ lean_dec(x_12); x_15 = lean_ctor_get(x_13, 1); lean_inc(x_15); lean_dec(x_13); -x_16 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Linter_unusedVariables___spec__18(x_2, x_3, x_11, x_5, x_15, x_7, x_8, x_14); +x_16 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Linter_unusedVariables___elambda__1___spec__18(x_2, x_3, x_11, x_5, x_15, x_7, x_8, x_14); return x_16; } else @@ -6548,7 +6551,7 @@ lean_inc(x_7); lean_inc(x_5); lean_inc(x_3); lean_inc(x_2); -x_23 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Linter_unusedVariables___spec__18(x_2, x_3, x_21, x_5, x_6, x_7, x_8, x_9); +x_23 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Linter_unusedVariables___elambda__1___spec__18(x_2, x_3, x_21, x_5, x_6, x_7, x_8, x_9); if (lean_obj_tag(x_23) == 0) { lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; @@ -6560,7 +6563,7 @@ lean_dec(x_23); x_26 = lean_ctor_get(x_24, 1); lean_inc(x_26); lean_dec(x_24); -x_27 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Linter_unusedVariables___spec__18(x_2, x_3, x_22, x_5, x_26, x_7, x_8, x_25); +x_27 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Linter_unusedVariables___elambda__1___spec__18(x_2, x_3, x_22, x_5, x_26, x_7, x_8, x_25); return x_27; } else @@ -6605,7 +6608,7 @@ lean_inc(x_7); lean_inc(x_5); lean_inc(x_3); lean_inc(x_2); -x_34 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Linter_unusedVariables___spec__18(x_2, x_3, x_32, x_5, x_6, x_7, x_8, x_9); +x_34 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Linter_unusedVariables___elambda__1___spec__18(x_2, x_3, x_32, x_5, x_6, x_7, x_8, x_9); if (lean_obj_tag(x_34) == 0) { lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; @@ -6617,7 +6620,7 @@ lean_dec(x_34); x_37 = lean_ctor_get(x_35, 1); lean_inc(x_37); lean_dec(x_35); -x_38 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Linter_unusedVariables___spec__18(x_2, x_3, x_33, x_5, x_37, x_7, x_8, x_36); +x_38 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Linter_unusedVariables___elambda__1___spec__18(x_2, x_3, x_33, x_5, x_37, x_7, x_8, x_36); return x_38; } else @@ -6664,7 +6667,7 @@ lean_inc(x_7); lean_inc(x_5); lean_inc(x_3); lean_inc(x_2); -x_46 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Linter_unusedVariables___spec__18(x_2, x_3, x_43, x_5, x_6, x_7, x_8, x_9); +x_46 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Linter_unusedVariables___elambda__1___spec__18(x_2, x_3, x_43, x_5, x_6, x_7, x_8, x_9); if (lean_obj_tag(x_46) == 0) { lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; @@ -6681,7 +6684,7 @@ lean_inc(x_7); lean_inc(x_5); lean_inc(x_3); lean_inc(x_2); -x_50 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Linter_unusedVariables___spec__18(x_2, x_3, x_44, x_5, x_49, x_7, x_8, x_48); +x_50 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Linter_unusedVariables___elambda__1___spec__18(x_2, x_3, x_44, x_5, x_49, x_7, x_8, x_48); if (lean_obj_tag(x_50) == 0) { lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; @@ -6693,7 +6696,7 @@ lean_dec(x_50); x_53 = lean_ctor_get(x_51, 1); lean_inc(x_53); lean_dec(x_51); -x_54 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Linter_unusedVariables___spec__18(x_2, x_3, x_45, x_5, x_53, x_7, x_8, x_52); +x_54 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Linter_unusedVariables___elambda__1___spec__18(x_2, x_3, x_45, x_5, x_53, x_7, x_8, x_52); return x_54; } else @@ -6761,7 +6764,7 @@ lean_object* x_63; lean_object* x_64; x_63 = lean_ctor_get(x_1, 1); lean_inc(x_63); lean_dec(x_1); -x_64 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Linter_unusedVariables___spec__18(x_2, x_3, x_63, x_5, x_6, x_7, x_8, x_9); +x_64 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Linter_unusedVariables___elambda__1___spec__18(x_2, x_3, x_63, x_5, x_6, x_7, x_8, x_9); return x_64; } case 11: @@ -6770,7 +6773,7 @@ lean_object* x_65; lean_object* x_66; x_65 = lean_ctor_get(x_1, 2); lean_inc(x_65); lean_dec(x_1); -x_66 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Linter_unusedVariables___spec__18(x_2, x_3, x_65, x_5, x_6, x_7, x_8, x_9); +x_66 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Linter_unusedVariables___elambda__1___spec__18(x_2, x_3, x_65, x_5, x_6, x_7, x_8, x_9); return x_66; } default: @@ -6794,12 +6797,12 @@ return x_69; } } } -LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit_go___at_Lean_Linter_unusedVariables___spec__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_EXPORT lean_object* l_Lean_ForEachExprWhere_visit_go___at_Lean_Linter_unusedVariables___elambda__1___spec__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) { _start: { lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; lean_inc(x_3); -x_9 = l_Lean_ForEachExprWhere_visited___at_Lean_Linter_unusedVariables___spec__19(x_3, x_4, x_5, x_6, x_7, x_8); +x_9 = l_Lean_ForEachExprWhere_visited___at_Lean_Linter_unusedVariables___elambda__1___spec__19(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_10, 0); @@ -6824,14 +6827,14 @@ if (x_16 == 0) { lean_object* x_17; lean_object* x_18; x_17 = lean_box(0); -x_18 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Linter_unusedVariables___spec__18___lambda__1(x_3, x_1, x_2, x_17, x_4, x_14, x_6, x_7, x_13); +x_18 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Linter_unusedVariables___elambda__1___spec__18___lambda__1(x_3, x_1, x_2, x_17, x_4, x_14, x_6, x_7, x_13); return x_18; } else { lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; lean_inc(x_3); -x_19 = l_Lean_ForEachExprWhere_checked___at_Lean_Linter_unusedVariables___spec__20(x_3, x_4, x_14, x_6, x_7, x_13); +x_19 = l_Lean_ForEachExprWhere_checked___at_Lean_Linter_unusedVariables___elambda__1___spec__20(x_3, x_4, x_14, x_6, x_7, x_13); x_20 = lean_ctor_get(x_19, 0); lean_inc(x_20); x_21 = lean_ctor_get(x_20, 0); @@ -6865,7 +6868,7 @@ lean_inc(x_28); x_29 = lean_ctor_get(x_26, 1); lean_inc(x_29); lean_dec(x_26); -x_30 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Linter_unusedVariables___spec__18___lambda__1(x_3, x_1, x_2, x_28, x_4, x_29, x_6, x_7, x_27); +x_30 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Linter_unusedVariables___elambda__1___spec__18___lambda__1(x_3, x_1, x_2, x_28, x_4, x_29, x_6, x_7, x_27); return x_30; } else @@ -6907,7 +6910,7 @@ x_36 = lean_ctor_get(x_20, 1); lean_inc(x_36); lean_dec(x_20); x_37 = lean_box(0); -x_38 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Linter_unusedVariables___spec__18___lambda__1(x_3, x_1, x_2, x_37, x_4, x_36, x_6, x_7, x_35); +x_38 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Linter_unusedVariables___elambda__1___spec__18___lambda__1(x_3, x_1, x_2, x_37, x_4, x_36, x_6, x_7, x_35); return x_38; } } @@ -6983,7 +6986,7 @@ return x_52; } } } -LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit___at_Lean_Linter_unusedVariables___spec__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_EXPORT lean_object* l_Lean_ForEachExprWhere_visit___at_Lean_Linter_unusedVariables___elambda__1___spec__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) { _start: { lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; @@ -6995,7 +6998,7 @@ x_11 = lean_ctor_get(x_9, 1); lean_inc(x_11); lean_dec(x_9); lean_inc(x_10); -x_12 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Linter_unusedVariables___spec__18(x_1, x_2, x_3, x_10, x_4, x_5, x_6, x_11); +x_12 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Linter_unusedVariables___elambda__1___spec__18(x_1, x_2, x_3, x_10, x_4, x_5, x_6, x_11); if (lean_obj_tag(x_12) == 0) { lean_object* x_13; lean_object* x_14; uint8_t x_15; @@ -7089,7 +7092,7 @@ return x_31; } } } -LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___spec__22(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___elambda__1___spec__22(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -7116,7 +7119,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Linter_unusedVariables___spec__23(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_Linter_unusedVariables___elambda__1___spec__23(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { _start: { uint8_t x_5; @@ -7125,7 +7128,7 @@ 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 = l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___spec__22(x_4, x_6); +x_7 = l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___elambda__1___spec__22(x_4, x_6); lean_dec(x_6); x_8 = 1; x_9 = lean_usize_add(x_2, x_8); @@ -7139,7 +7142,7 @@ return x_4; } } } -LEAN_EXPORT lean_object* l_Lean_HashMap_toList___at_Lean_Linter_unusedVariables___spec__21(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_HashMap_toList___at_Lean_Linter_unusedVariables___elambda__1___spec__21(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; @@ -7172,14 +7175,14 @@ size_t x_8; size_t x_9; lean_object* x_10; x_8 = 0; x_9 = lean_usize_of_nat(x_4); lean_dec(x_4); -x_10 = l_Array_foldlMUnsafe_fold___at_Lean_Linter_unusedVariables___spec__23(x_3, x_8, x_9, x_2); +x_10 = l_Array_foldlMUnsafe_fold___at_Lean_Linter_unusedVariables___elambda__1___spec__23(x_3, x_8, x_9, x_2); lean_dec(x_3); return x_10; } } } } -static lean_object* _init_l_panic___at_Lean_Linter_unusedVariables___spec__28___closed__1() { +static lean_object* _init_l_panic___at_Lean_Linter_unusedVariables___elambda__1___spec__28___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -7189,17 +7192,17 @@ x_3 = l_instInhabited___rarg(x_2, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_panic___at_Lean_Linter_unusedVariables___spec__28(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_panic___at_Lean_Linter_unusedVariables___elambda__1___spec__28(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; -x_5 = l_panic___at_Lean_Linter_unusedVariables___spec__28___closed__1; +x_5 = l_panic___at_Lean_Linter_unusedVariables___elambda__1___spec__28___closed__1; x_6 = lean_panic_fn(x_5, x_1); x_7 = lean_apply_3(x_6, x_2, x_3, x_4); return x_7; } } -LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Linter_unusedVariables___spec__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_EXPORT lean_object* l_List_mapM_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__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) { _start: { if (lean_obj_tag(x_4) == 0) @@ -7230,7 +7233,7 @@ lean_inc(x_6); lean_inc(x_3); lean_inc(x_2); lean_inc(x_1); -x_14 = l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___spec__27(x_1, x_2, x_3, x_12, x_6, x_7, x_8); +x_14 = l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___elambda__1___spec__27(x_1, x_2, x_3, x_12, x_6, x_7, x_8); if (lean_obj_tag(x_14) == 0) { lean_object* x_15; lean_object* x_16; @@ -7295,7 +7298,7 @@ lean_inc(x_6); lean_inc(x_3); lean_inc(x_2); lean_inc(x_1); -x_24 = l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___spec__27(x_1, x_2, x_3, x_22, x_6, x_7, x_8); +x_24 = l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___elambda__1___spec__27(x_1, x_2, x_3, x_22, x_6, x_7, x_8); if (lean_obj_tag(x_24) == 0) { lean_object* x_25; lean_object* x_26; lean_object* x_27; @@ -7347,7 +7350,7 @@ return x_32; } } } -static lean_object* _init_l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___spec__27___closed__1() { +static lean_object* _init_l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___elambda__1___spec__27___closed__1() { _start: { lean_object* x_1; @@ -7355,7 +7358,7 @@ x_1 = lean_mk_string_from_bytes("Lean.Server.InfoUtils", 21); return x_1; } } -static lean_object* _init_l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___spec__27___closed__2() { +static lean_object* _init_l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___elambda__1___spec__27___closed__2() { _start: { lean_object* x_1; @@ -7363,7 +7366,7 @@ x_1 = lean_mk_string_from_bytes("Lean.Elab.InfoTree.visitM.go", 28); return x_1; } } -static lean_object* _init_l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___spec__27___closed__3() { +static lean_object* _init_l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___elambda__1___spec__27___closed__3() { _start: { lean_object* x_1; @@ -7371,20 +7374,20 @@ x_1 = lean_mk_string_from_bytes("unexpected context-free info tree node", 38); return x_1; } } -static lean_object* _init_l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___spec__27___closed__4() { +static lean_object* _init_l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___elambda__1___spec__27___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_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___spec__27___closed__1; -x_2 = l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___spec__27___closed__2; +x_1 = l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___elambda__1___spec__27___closed__1; +x_2 = l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___elambda__1___spec__27___closed__2; x_3 = lean_unsigned_to_nat(50u); x_4 = lean_unsigned_to_nat(21u); -x_5 = l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___spec__27___closed__3; +x_5 = l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___elambda__1___spec__27___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___spec__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_EXPORT lean_object* l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___elambda__1___spec__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) { _start: { if (lean_obj_tag(x_3) == 0) @@ -7410,8 +7413,8 @@ lean_object* x_12; lean_object* x_13; lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_12 = l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___spec__27___closed__4; -x_13 = l_panic___at_Lean_Linter_unusedVariables___spec__28(x_12, x_5, x_6, x_7); +x_12 = l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___elambda__1___spec__27___closed__4; +x_13 = l_panic___at_Lean_Linter_unusedVariables___elambda__1___spec__28(x_12, x_5, x_6, x_7); return x_13; } default: @@ -7501,7 +7504,7 @@ x_33 = lean_box(0); lean_inc(x_6); lean_inc(x_5); lean_inc(x_2); -x_34 = l_List_mapM_loop___at_Lean_Linter_unusedVariables___spec__29(x_1, x_2, x_31, x_32, x_33, x_5, x_6, x_30); +x_34 = l_List_mapM_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__29(x_1, x_2, x_31, x_32, x_33, x_5, x_6, x_30); if (lean_obj_tag(x_34) == 0) { lean_object* x_35; lean_object* x_36; lean_object* x_37; @@ -7657,7 +7660,7 @@ x_65 = lean_box(0); lean_inc(x_6); lean_inc(x_5); lean_inc(x_2); -x_66 = l_List_mapM_loop___at_Lean_Linter_unusedVariables___spec__29(x_1, x_2, x_63, x_64, x_65, x_5, x_6, x_61); +x_66 = l_List_mapM_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__29(x_1, x_2, x_63, x_64, x_65, x_5, x_6, x_61); if (lean_obj_tag(x_66) == 0) { lean_object* x_67; lean_object* x_68; lean_object* x_69; @@ -7801,16 +7804,16 @@ return x_88; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_visitM___at_Lean_Linter_unusedVariables___spec__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_EXPORT lean_object* l_Lean_Elab_InfoTree_visitM___at_Lean_Linter_unusedVariables___elambda__1___spec__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) { _start: { lean_object* x_7; lean_object* x_8; x_7 = lean_box(0); -x_8 = l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___spec__27(x_1, x_2, x_7, x_3, x_4, x_5, x_6); +x_8 = l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___elambda__1___spec__27(x_1, x_2, x_7, x_3, x_4, x_5, x_6); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___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_EXPORT lean_object* l_Lean_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___elambda__1___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) { _start: { lean_object* x_7; lean_object* x_8; @@ -7821,7 +7824,7 @@ lean_ctor_set(x_8, 1, x_6); return x_8; } } -static lean_object* _init_l_Lean_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___spec__25___lambda__2___closed__1() { +static lean_object* _init_l_Lean_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___elambda__1___spec__25___lambda__2___closed__1() { _start: { lean_object* x_1; lean_object* x_2; @@ -7831,7 +7834,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___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_EXPORT lean_object* l_Lean_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___elambda__1___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) { _start: { lean_object* x_9; lean_object* x_10; @@ -7888,7 +7891,7 @@ return x_25; else { lean_object* x_26; lean_object* x_27; -x_26 = l_Lean_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___spec__25___lambda__2___closed__1; +x_26 = l_Lean_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___elambda__1___spec__25___lambda__2___closed__1; x_27 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_27, 0, x_26); lean_ctor_set(x_27, 1, x_8); @@ -7955,30 +7958,30 @@ return x_41; } } } -static lean_object* _init_l_Lean_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___spec__25___closed__1() { +static lean_object* _init_l_Lean_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___elambda__1___spec__25___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___spec__25___lambda__1___boxed), 6, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___elambda__1___spec__25___lambda__1___boxed), 6, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___spec__25(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_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___elambda__1___spec__25(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 = lean_alloc_closure((void*)(l_Lean_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___spec__25___lambda__2___boxed), 8, 1); +x_6 = lean_alloc_closure((void*)(l_Lean_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___elambda__1___spec__25___lambda__2___boxed), 8, 1); lean_closure_set(x_6, 0, x_1); -x_7 = l_Lean_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___spec__25___closed__1; -x_8 = l_Lean_Elab_InfoTree_visitM___at_Lean_Linter_unusedVariables___spec__26(x_7, x_6, x_2, x_3, x_4, x_5); +x_7 = l_Lean_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___elambda__1___spec__25___closed__1; +x_8 = l_Lean_Elab_InfoTree_visitM___at_Lean_Linter_unusedVariables___elambda__1___spec__26(x_7, x_6, x_2, x_3, x_4, x_5); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Linter_collectMacroExpansions_x3f___at_Lean_Linter_unusedVariables___spec__24(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_Linter_collectMacroExpansions_x3f___at_Lean_Linter_unusedVariables___elambda__1___spec__24(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_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___spec__25(x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___elambda__1___spec__25(x_1, x_2, x_3, x_4, x_5); if (lean_obj_tag(x_6) == 0) { lean_object* x_7; @@ -8128,7 +8131,7 @@ return x_38; } } } -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___spec__30(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_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___elambda__1___spec__30(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6) { _start: { uint8_t x_7; @@ -8172,7 +8175,7 @@ return x_15; } } } -LEAN_EXPORT uint8_t l_List_foldr___at_Lean_Linter_unusedVariables___spec__31___lambda__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l_List_foldr___at_Lean_Linter_unusedVariables___elambda__1___spec__31___lambda__1(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; @@ -8196,7 +8199,7 @@ return x_7; } } } -LEAN_EXPORT uint8_t l_List_foldr___at_Lean_Linter_unusedVariables___spec__31___lambda__2(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l_List_foldr___at_Lean_Linter_unusedVariables___elambda__1___spec__31___lambda__2(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; @@ -8232,7 +8235,7 @@ return x_9; } } } -LEAN_EXPORT uint8_t l_List_foldr___at_Lean_Linter_unusedVariables___spec__31(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_EXPORT uint8_t l_List_foldr___at_Lean_Linter_unusedVariables___elambda__1___spec__31(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: { if (lean_obj_tag(x_6) == 0) @@ -8253,14 +8256,14 @@ lean_dec(x_6); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_9 = l_List_foldr___at_Lean_Linter_unusedVariables___spec__31(x_1, x_2, x_3, x_4, x_5, x_8); +x_9 = l_List_foldr___at_Lean_Linter_unusedVariables___elambda__1___spec__31(x_1, x_2, x_3, x_4, x_5, x_8); x_10 = lean_ctor_get(x_7, 2); lean_inc(x_10); lean_dec(x_7); lean_inc(x_3); -x_11 = lean_alloc_closure((void*)(l_List_foldr___at_Lean_Linter_unusedVariables___spec__31___lambda__1___boxed), 2, 1); +x_11 = lean_alloc_closure((void*)(l_List_foldr___at_Lean_Linter_unusedVariables___elambda__1___spec__31___lambda__1___boxed), 2, 1); lean_closure_set(x_11, 0, x_3); -x_12 = lean_alloc_closure((void*)(l_List_foldr___at_Lean_Linter_unusedVariables___spec__31___lambda__2___boxed), 2, 1); +x_12 = lean_alloc_closure((void*)(l_List_foldr___at_Lean_Linter_unusedVariables___elambda__1___spec__31___lambda__2___boxed), 2, 1); lean_closure_set(x_12, 0, x_3); x_13 = l_Lean_Syntax_findStack_x3f(x_10, x_11, x_12); if (lean_obj_tag(x_13) == 0) @@ -8316,7 +8319,7 @@ size_t x_20; size_t x_21; uint8_t x_22; x_20 = 0; x_21 = lean_usize_of_nat(x_16); lean_dec(x_16); -x_22 = l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___spec__30(x_2, x_4, x_15, x_1, x_20, x_21); +x_22 = l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___elambda__1___spec__30(x_2, x_4, x_15, x_1, x_20, x_21); if (x_22 == 0) { return x_9; @@ -8334,7 +8337,7 @@ return x_23; } } } -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___spec__32(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_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___elambda__1___spec__32(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; @@ -8346,7 +8349,7 @@ x_12 = lean_array_uget(x_5, x_6); lean_inc(x_9); lean_inc(x_8); lean_inc(x_3); -x_13 = l_Lean_Linter_collectMacroExpansions_x3f___at_Lean_Linter_unusedVariables___spec__24(x_3, x_12, x_8, x_9, x_10); +x_13 = l_Lean_Linter_collectMacroExpansions_x3f___at_Lean_Linter_unusedVariables___elambda__1___spec__24(x_3, x_12, x_8, x_9, x_10); if (lean_obj_tag(x_13) == 0) { lean_object* x_14; @@ -8381,7 +8384,7 @@ x_23 = 0; lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_24 = l_List_foldr___at_Lean_Linter_unusedVariables___spec__31(x_1, x_2, x_3, x_4, x_23, x_22); +x_24 = l_List_foldr___at_Lean_Linter_unusedVariables___elambda__1___spec__31(x_1, x_2, x_3, x_4, x_23, x_22); if (x_24 == 0) { size_t x_25; size_t x_26; @@ -8419,7 +8422,7 @@ x_32 = 0; lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_33 = l_List_foldr___at_Lean_Linter_unusedVariables___spec__31(x_1, x_2, x_3, x_4, x_32, x_31); +x_33 = l_List_foldr___at_Lean_Linter_unusedVariables___elambda__1___spec__31(x_1, x_2, x_3, x_4, x_32, x_31); if (x_33 == 0) { size_t x_34; size_t x_35; @@ -8492,7 +8495,7 @@ return x_46; } } } -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___spec__33(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_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___elambda__1___spec__33(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6) { _start: { uint8_t x_7; @@ -8536,7 +8539,7 @@ return x_15; } } } -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___spec__34(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4) { +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___elambda__1___spec__34(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4) { _start: { uint8_t x_5; @@ -8588,7 +8591,7 @@ return x_16; } } } -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35___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_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35___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; lean_object* x_9; lean_object* x_10; lean_object* x_11; @@ -8604,7 +8607,7 @@ lean_ctor_set(x_11, 1, x_7); return x_11; } } -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35___lambda__2(uint8_t 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_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35___lambda__2(uint8_t 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: { if (x_1 == 0) @@ -8615,7 +8618,7 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); x_15 = lean_box(0); -x_16 = l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35___lambda__1(x_2, x_3, x_10, x_15, x_12, x_13, x_14); +x_16 = l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35___lambda__1(x_2, x_3, x_10, x_15, x_12, x_13, x_14); lean_dec(x_13); lean_dec(x_12); return x_16; @@ -8630,7 +8633,7 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); x_17 = lean_box(0); -x_18 = l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35___lambda__1(x_2, x_3, x_10, x_17, x_12, x_13, x_14); +x_18 = l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35___lambda__1(x_2, x_3, x_10, x_17, x_12, x_13, x_14); lean_dec(x_13); lean_dec(x_12); return x_18; @@ -8644,7 +8647,7 @@ lean_dec(x_5); lean_inc(x_13); lean_inc(x_12); lean_inc(x_2); -x_21 = l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___spec__32(x_6, x_2, x_7, x_8, x_9, x_19, x_20, x_12, x_13, x_14); +x_21 = l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___elambda__1___spec__32(x_6, x_2, x_7, x_8, x_9, x_19, x_20, x_12, x_13, x_14); lean_dec(x_6); if (lean_obj_tag(x_21) == 0) { @@ -8660,7 +8663,7 @@ x_24 = lean_ctor_get(x_21, 1); lean_inc(x_24); lean_dec(x_21); x_25 = lean_box(0); -x_26 = l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35___lambda__1(x_2, x_3, x_10, x_25, x_12, x_13, x_24); +x_26 = l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35___lambda__1(x_2, x_3, x_10, x_25, x_12, x_13, x_24); lean_dec(x_13); lean_dec(x_12); return x_26; @@ -8729,7 +8732,7 @@ return x_36; } } } -LEAN_EXPORT uint8_t l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35___lambda__3(lean_object* x_1) { +LEAN_EXPORT uint8_t l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35___lambda__3(lean_object* x_1) { _start: { uint8_t x_2; @@ -8748,23 +8751,23 @@ return x_4; } } } -static lean_object* _init_l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35___lambda__4___closed__1() { +static lean_object* _init_l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35___lambda__4___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35___lambda__3___boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35___lambda__3___boxed), 1, 0); return x_1; } } -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35___lambda__4(uint8_t 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_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35___lambda__4(uint8_t 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: { lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_dec(x_12); lean_inc(x_7); -x_16 = lean_alloc_closure((void*)(l_List_foldr___at_Lean_Linter_unusedVariables___spec__31___lambda__1___boxed), 2, 1); +x_16 = lean_alloc_closure((void*)(l_List_foldr___at_Lean_Linter_unusedVariables___elambda__1___spec__31___lambda__1___boxed), 2, 1); lean_closure_set(x_16, 0, x_7); -x_17 = l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35___lambda__4___closed__1; +x_17 = l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35___lambda__4___closed__1; x_18 = l_Lean_Syntax_findStack_x3f(x_10, x_16, x_17); if (lean_obj_tag(x_18) == 0) { @@ -8826,7 +8829,7 @@ if (x_27 == 0) lean_object* x_28; lean_object* x_29; lean_dec(x_25); x_28 = lean_box(0); -x_29 = l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_11, x_28, x_13, x_14, x_15); +x_29 = l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_11, x_28, x_13, x_14, x_15); return x_29; } else @@ -8841,7 +8844,7 @@ lean_object* x_33; lean_object* x_34; lean_dec(x_30); lean_dec(x_25); x_33 = lean_box(0); -x_34 = l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_11, x_33, x_13, x_14, x_15); +x_34 = l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_11, x_33, x_13, x_14, x_15); return x_34; } else @@ -8854,7 +8857,7 @@ lean_object* x_36; lean_object* x_37; lean_dec(x_30); lean_dec(x_25); x_36 = lean_box(0); -x_37 = l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_11, x_36, x_13, x_14, x_15); +x_37 = l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_11, x_36, x_13, x_14, x_15); return x_37; } else @@ -8865,12 +8868,12 @@ x_39 = lean_usize_of_nat(x_30); lean_dec(x_30); lean_inc(x_8); lean_inc(x_2); -x_40 = l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___spec__33(x_2, x_8, x_25, x_6, x_38, x_39); +x_40 = l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___elambda__1___spec__33(x_2, x_8, x_25, x_6, x_38, x_39); if (x_40 == 0) { lean_object* x_41; lean_object* x_42; x_41 = lean_box(0); -x_42 = l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_11, x_41, x_13, x_14, x_15); +x_42 = l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_11, x_41, x_13, x_14, x_15); return x_42; } else @@ -8898,7 +8901,7 @@ return x_44; } } } -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35___lambda__5(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, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35___lambda__5(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, 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; @@ -8934,12 +8937,12 @@ else { lean_object* x_22; lean_object* x_23; x_22 = lean_box(0); -x_23 = l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35___lambda__4(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_17, x_9, x_10, x_11, x_22, x_13, x_14, x_15); +x_23 = l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35___lambda__4(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_17, x_9, x_10, x_11, x_22, x_13, x_14, x_15); return x_23; } } } -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35___lambda__6(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, 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_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35___lambda__6(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, 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_19; @@ -8987,7 +8990,7 @@ lean_dec(x_24); lean_dec(x_23); lean_dec(x_12); x_27 = lean_box(0); -x_28 = l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35___lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14, x_27, x_16, x_17, x_18); +x_28 = l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35___lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14, x_27, x_16, x_17, x_18); return x_28; } else @@ -9001,7 +9004,7 @@ lean_dec(x_24); lean_dec(x_23); lean_dec(x_12); x_30 = lean_box(0); -x_31 = l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35___lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14, x_30, x_16, x_17, x_18); +x_31 = l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35___lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14, x_30, x_16, x_17, x_18); return x_31; } else @@ -9010,13 +9013,13 @@ size_t x_32; size_t x_33; uint8_t x_34; x_32 = 0; x_33 = lean_usize_of_nat(x_24); lean_dec(x_24); -x_34 = l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___spec__34(x_12, x_23, x_32, x_33); +x_34 = l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___elambda__1___spec__34(x_12, x_23, x_32, x_33); lean_dec(x_23); if (x_34 == 0) { lean_object* x_35; lean_object* x_36; x_35 = lean_box(0); -x_36 = l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35___lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14, x_35, x_16, x_17, x_18); +x_36 = l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35___lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14, x_35, x_16, x_17, x_18); return x_36; } else @@ -9064,7 +9067,7 @@ return x_40; } } } -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t 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, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t 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, lean_object* x_12, lean_object* x_13) { _start: { if (lean_obj_tag(x_9) == 0) @@ -9190,7 +9193,7 @@ lean_inc(x_7); lean_inc(x_1); lean_inc(x_8); lean_inc(x_4); -x_41 = l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35___lambda__6(x_23, x_5, x_25, x_34, x_6, x_4, x_8, x_29, x_3, x_1, x_22, x_7, x_21, x_10, x_40, x_11, x_12, x_13); +x_41 = l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35___lambda__6(x_23, x_5, x_25, x_34, x_6, x_4, x_8, x_29, x_3, x_1, x_22, x_7, x_21, x_10, x_40, x_11, x_12, x_13); if (lean_obj_tag(x_41) == 0) { lean_object* x_42; @@ -9296,7 +9299,7 @@ goto _start; } } } -static lean_object* _init_l_Array_qsort_sort___at_Lean_Linter_unusedVariables___spec__36___lambda__1___closed__1() { +static lean_object* _init_l_Array_qsort_sort___at_Lean_Linter_unusedVariables___elambda__1___spec__36___lambda__1___closed__1() { _start: { lean_object* x_1; @@ -9304,7 +9307,7 @@ x_1 = lean_mk_string_from_bytes("Init.Data.Option.BasicAux", 25); return x_1; } } -static lean_object* _init_l_Array_qsort_sort___at_Lean_Linter_unusedVariables___spec__36___lambda__1___closed__2() { +static lean_object* _init_l_Array_qsort_sort___at_Lean_Linter_unusedVariables___elambda__1___spec__36___lambda__1___closed__2() { _start: { lean_object* x_1; @@ -9312,7 +9315,7 @@ x_1 = lean_mk_string_from_bytes("Option.get!", 11); return x_1; } } -static lean_object* _init_l_Array_qsort_sort___at_Lean_Linter_unusedVariables___spec__36___lambda__1___closed__3() { +static lean_object* _init_l_Array_qsort_sort___at_Lean_Linter_unusedVariables___elambda__1___spec__36___lambda__1___closed__3() { _start: { lean_object* x_1; @@ -9320,20 +9323,20 @@ x_1 = lean_mk_string_from_bytes("value is none", 13); return x_1; } } -static lean_object* _init_l_Array_qsort_sort___at_Lean_Linter_unusedVariables___spec__36___lambda__1___closed__4() { +static lean_object* _init_l_Array_qsort_sort___at_Lean_Linter_unusedVariables___elambda__1___spec__36___lambda__1___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_Array_qsort_sort___at_Lean_Linter_unusedVariables___spec__36___lambda__1___closed__1; -x_2 = l_Array_qsort_sort___at_Lean_Linter_unusedVariables___spec__36___lambda__1___closed__2; +x_1 = l_Array_qsort_sort___at_Lean_Linter_unusedVariables___elambda__1___spec__36___lambda__1___closed__1; +x_2 = l_Array_qsort_sort___at_Lean_Linter_unusedVariables___elambda__1___spec__36___lambda__1___closed__2; x_3 = lean_unsigned_to_nat(16u); x_4 = lean_unsigned_to_nat(14u); -x_5 = l_Array_qsort_sort___at_Lean_Linter_unusedVariables___spec__36___lambda__1___closed__3; +x_5 = l_Array_qsort_sort___at_Lean_Linter_unusedVariables___elambda__1___spec__36___lambda__1___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 uint8_t l_Array_qsort_sort___at_Lean_Linter_unusedVariables___spec__36___lambda__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l_Array_qsort_sort___at_Lean_Linter_unusedVariables___elambda__1___spec__36___lambda__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; uint8_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; @@ -9345,7 +9348,7 @@ x_7 = l_Lean_Syntax_getPos_x3f(x_6, x_4); if (lean_obj_tag(x_5) == 0) { lean_object* x_8; lean_object* x_9; -x_8 = l_Array_qsort_sort___at_Lean_Linter_unusedVariables___spec__36___lambda__1___closed__4; +x_8 = l_Array_qsort_sort___at_Lean_Linter_unusedVariables___elambda__1___spec__36___lambda__1___closed__4; x_9 = l_panic___at_Lean_Elab_InfoTree_smallestInfo_x3f___spec__1(x_8); if (lean_obj_tag(x_7) == 0) { @@ -9376,7 +9379,7 @@ lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; x_14 = lean_ctor_get(x_5, 0); lean_inc(x_14); lean_dec(x_5); -x_15 = l_Array_qsort_sort___at_Lean_Linter_unusedVariables___spec__36___lambda__1___closed__4; +x_15 = l_Array_qsort_sort___at_Lean_Linter_unusedVariables___elambda__1___spec__36___lambda__1___closed__4; x_16 = l_panic___at_Lean_Elab_InfoTree_smallestInfo_x3f___spec__1(x_15); x_17 = lean_nat_dec_lt(x_14, x_16); lean_dec(x_16); @@ -9400,15 +9403,15 @@ return x_20; } } } -static lean_object* _init_l_Array_qsort_sort___at_Lean_Linter_unusedVariables___spec__36___closed__1() { +static lean_object* _init_l_Array_qsort_sort___at_Lean_Linter_unusedVariables___elambda__1___spec__36___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Array_qsort_sort___at_Lean_Linter_unusedVariables___spec__36___lambda__1___boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l_Array_qsort_sort___at_Lean_Linter_unusedVariables___elambda__1___spec__36___lambda__1___boxed), 2, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Linter_unusedVariables___spec__36(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Linter_unusedVariables___elambda__1___spec__36(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -9421,7 +9424,7 @@ return x_1; else { lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; -x_5 = l_Array_qsort_sort___at_Lean_Linter_unusedVariables___spec__36___closed__1; +x_5 = l_Array_qsort_sort___at_Lean_Linter_unusedVariables___elambda__1___spec__36___closed__1; lean_inc(x_2); x_6 = l_Array_qpartition___rarg(x_1, x_5, x_2, x_3); x_7 = lean_ctor_get(x_6, 0); @@ -9433,7 +9436,7 @@ x_9 = lean_nat_dec_le(x_3, x_7); if (x_9 == 0) { lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = l_Array_qsort_sort___at_Lean_Linter_unusedVariables___spec__36(x_8, x_2, x_7); +x_10 = l_Array_qsort_sort___at_Lean_Linter_unusedVariables___elambda__1___spec__36(x_8, x_2, x_7); x_11 = lean_unsigned_to_nat(1u); x_12 = lean_nat_add(x_7, x_11); lean_dec(x_7); @@ -9450,7 +9453,7 @@ return x_8; } } } -static lean_object* _init_l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___spec__37___closed__1() { +static lean_object* _init_l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___elambda__1___spec__37___closed__1() { _start: { lean_object* x_1; lean_object* x_2; @@ -9459,7 +9462,7 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___spec__37___closed__2() { +static lean_object* _init_l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___elambda__1___spec__37___closed__2() { _start: { lean_object* x_1; @@ -9467,16 +9470,16 @@ x_1 = lean_mk_string_from_bytes(" [", 2); return x_1; } } -static lean_object* _init_l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___spec__37___closed__3() { +static lean_object* _init_l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___elambda__1___spec__37___closed__3() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___spec__37___closed__2; +x_1 = l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___elambda__1___spec__37___closed__2; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___spec__37___closed__4() { +static lean_object* _init_l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___elambda__1___spec__37___closed__4() { _start: { lean_object* x_1; @@ -9484,27 +9487,27 @@ x_1 = lean_mk_string_from_bytes("]", 1); return x_1; } } -static lean_object* _init_l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___spec__37___closed__5() { +static lean_object* _init_l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___elambda__1___spec__37___closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___spec__37___closed__4; +x_1 = l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___elambda__1___spec__37___closed__4; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___spec__37(lean_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_Linter_logLint___at_Lean_Linter_unusedVariables___elambda__1___spec__37(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; uint8_t x_17; lean_object* x_18; x_7 = lean_ctor_get(x_1, 0); lean_inc(x_7); lean_dec(x_1); -x_8 = l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___spec__37___closed__1; +x_8 = l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___elambda__1___spec__37___closed__1; x_9 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_9, 0, x_8); lean_ctor_set(x_9, 1, x_3); -x_10 = l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___spec__37___closed__3; +x_10 = l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___elambda__1___spec__37___closed__3; x_11 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_11, 0, x_9); lean_ctor_set(x_11, 1, x_10); @@ -9513,7 +9516,7 @@ x_12 = l_Lean_MessageData_ofName(x_7); x_13 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_13, 0, x_11); lean_ctor_set(x_13, 1, x_12); -x_14 = l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___spec__37___closed__5; +x_14 = l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___elambda__1___spec__37___closed__5; x_15 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_15, 0, x_13); lean_ctor_set(x_15, 1, x_14); @@ -9525,7 +9528,7 @@ x_18 = l_Lean_logAt___at_Lean_Elab_Command_runLinters___spec__2(x_2, x_16, x_17, return x_18; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___spec__38___closed__1() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__38___closed__1() { _start: { lean_object* x_1; @@ -9533,16 +9536,16 @@ x_1 = lean_mk_string_from_bytes("unused variable `", 17); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___spec__38___closed__2() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__38___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___spec__38___closed__1; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__38___closed__1; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___spec__38___closed__3() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__38___closed__3() { _start: { lean_object* x_1; @@ -9550,16 +9553,16 @@ x_1 = lean_mk_string_from_bytes("`", 1); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___spec__38___closed__4() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__38___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___spec__38___closed__3; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__38___closed__3; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___spec__38(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_forInUnsafe_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__38(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; @@ -9585,16 +9588,16 @@ lean_dec(x_10); x_13 = l_Lean_LocalDecl_userName(x_12); lean_dec(x_12); x_14 = l_Lean_MessageData_ofName(x_13); -x_15 = l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___spec__38___closed__2; +x_15 = l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__38___closed__2; x_16 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_16, 0, x_15); lean_ctor_set(x_16, 1, x_14); -x_17 = l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___spec__38___closed__4; +x_17 = l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__38___closed__4; x_18 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_18, 0, x_16); lean_ctor_set(x_18, 1, x_17); x_19 = l_Lean_Linter_getLinterUnusedVariables___closed__1; -x_20 = l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___spec__37(x_19, x_11, x_18, x_5, x_6, x_7); +x_20 = l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___elambda__1___spec__37(x_19, x_11, x_18, x_5, x_6, x_7); lean_dec(x_11); x_21 = lean_ctor_get(x_20, 1); lean_inc(x_21); @@ -9609,7 +9612,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___spec__39___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_AssocList_foldlM___at_Lean_Linter_unusedVariables___elambda__1___spec__39___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; @@ -9625,7 +9628,7 @@ lean_ctor_set(x_10, 1, x_5); return x_10; } } -static lean_object* _init_l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___spec__39___closed__1() { +static lean_object* _init_l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___elambda__1___spec__39___closed__1() { _start: { lean_object* x_1; @@ -9633,15 +9636,15 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Expr_isFVar___boxed), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___spec__39___closed__2() { +static lean_object* _init_l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___elambda__1___spec__39___closed__2() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___spec__39___lambda__1___boxed), 5, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___elambda__1___spec__39___lambda__1___boxed), 5, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___spec__39(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_AssocList_foldlM___at_Lean_Linter_unusedVariables___elambda__1___spec__39(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { if (lean_obj_tag(x_2) == 0) @@ -9662,11 +9665,11 @@ lean_inc(x_7); x_8 = lean_ctor_get(x_2, 2); lean_inc(x_8); lean_dec(x_2); -x_9 = l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___spec__39___closed__1; -x_10 = l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___spec__39___closed__2; +x_9 = l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___elambda__1___spec__39___closed__1; +x_10 = l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___elambda__1___spec__39___closed__2; lean_inc(x_4); lean_inc(x_3); -x_11 = l_Lean_ForEachExprWhere_visit___at_Lean_Linter_unusedVariables___spec__17(x_9, x_10, x_7, x_1, x_3, x_4, x_5); +x_11 = l_Lean_ForEachExprWhere_visit___at_Lean_Linter_unusedVariables___elambda__1___spec__17(x_9, x_10, x_7, x_1, x_3, x_4, x_5); if (lean_obj_tag(x_11) == 0) { lean_object* x_12; lean_object* x_13; lean_object* x_14; @@ -9711,7 +9714,7 @@ return x_19; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Linter_unusedVariables___spec__40(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_foldlMUnsafe_fold___at_Lean_Linter_unusedVariables___elambda__1___spec__40(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; @@ -9722,7 +9725,7 @@ lean_object* x_9; lean_object* x_10; x_9 = lean_array_uget(x_1, x_2); lean_inc(x_6); lean_inc(x_5); -x_10 = l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___spec__39(x_4, x_9, x_5, x_6, x_7); +x_10 = l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___elambda__1___spec__39(x_4, x_9, x_5, x_6, x_7); if (lean_obj_tag(x_10) == 0) { lean_object* x_11; lean_object* x_12; size_t x_13; size_t x_14; @@ -9775,11 +9778,11 @@ return x_20; } } } -LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Linter_unusedVariables___spec__41___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Linter_unusedVariables___elambda__1___spec__41___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { uint8_t x_4; -x_4 = l_Lean_HashMapImp_contains___at_Lean_Linter_unusedVariables___spec__15(x_1, x_2); +x_4 = l_Lean_HashMapImp_contains___at_Lean_Linter_unusedVariables___elambda__1___spec__15(x_1, x_2); if (x_4 == 0) { lean_object* x_5; @@ -9794,15 +9797,15 @@ return x_1; } } } -static lean_object* _init_l_Array_foldrMUnsafe_fold___at_Lean_Linter_unusedVariables___spec__41___lambda__2___closed__1() { +static lean_object* _init_l_Array_foldrMUnsafe_fold___at_Lean_Linter_unusedVariables___elambda__1___spec__41___lambda__2___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Array_foldrMUnsafe_fold___at_Lean_Linter_unusedVariables___spec__41___lambda__1), 3, 0); +x_1 = lean_alloc_closure((void*)(l_Array_foldrMUnsafe_fold___at_Lean_Linter_unusedVariables___elambda__1___spec__41___lambda__1), 3, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Linter_unusedVariables___spec__41___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Linter_unusedVariables___elambda__1___spec__41___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_dec(x_1); @@ -9818,8 +9821,8 @@ lean_dec(x_4); x_6 = lean_ctor_get(x_5, 7); lean_inc(x_6); lean_dec(x_5); -x_7 = l_Array_foldrMUnsafe_fold___at_Lean_Linter_unusedVariables___spec__41___lambda__2___closed__1; -x_8 = l_Lean_PersistentHashMap_foldlM___at_Lean_Linter_unusedVariables___spec__16(x_6, x_7, x_3); +x_7 = l_Array_foldrMUnsafe_fold___at_Lean_Linter_unusedVariables___elambda__1___spec__41___lambda__2___closed__1; +x_8 = l_Lean_PersistentHashMap_foldlM___at_Lean_Linter_unusedVariables___elambda__1___spec__16(x_6, x_7, x_3); return x_8; } else @@ -9829,15 +9832,15 @@ return x_3; } } } -static lean_object* _init_l_Array_foldrMUnsafe_fold___at_Lean_Linter_unusedVariables___spec__41___closed__1() { +static lean_object* _init_l_Array_foldrMUnsafe_fold___at_Lean_Linter_unusedVariables___elambda__1___spec__41___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Array_foldrMUnsafe_fold___at_Lean_Linter_unusedVariables___spec__41___lambda__2), 3, 0); +x_1 = lean_alloc_closure((void*)(l_Array_foldrMUnsafe_fold___at_Lean_Linter_unusedVariables___elambda__1___spec__41___lambda__2), 3, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Linter_unusedVariables___spec__41(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Linter_unusedVariables___elambda__1___spec__41(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { _start: { uint8_t x_5; @@ -9848,7 +9851,7 @@ size_t x_6; size_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; x_6 = 1; x_7 = lean_usize_sub(x_2, x_6); x_8 = lean_array_uget(x_1, x_7); -x_9 = l_Array_foldrMUnsafe_fold___at_Lean_Linter_unusedVariables___spec__41___closed__1; +x_9 = l_Array_foldrMUnsafe_fold___at_Lean_Linter_unusedVariables___elambda__1___spec__41___closed__1; x_10 = l_Lean_Elab_InfoTree_foldInfo___rarg(x_9, x_4, x_8); x_2 = x_7; x_4 = x_10; @@ -9860,7 +9863,7 @@ return x_4; } } } -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___spec__42(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_anyMUnsafe_any___at_Lean_Linter_unusedVariables___elambda__1___spec__42(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; @@ -9976,7 +9979,7 @@ return x_41; } } } -static lean_object* _init_l_Lean_Linter_unusedVariables___lambda__1___closed__1() { +static lean_object* _init_l_Lean_Linter_unusedVariables___elambda__1___lambda__1___closed__1() { _start: { lean_object* x_1; lean_object* x_2; @@ -9985,7 +9988,7 @@ x_2 = l_Lean_mkHashMapImp___rarg(x_1); return x_2; } } -static lean_object* _init_l_Lean_Linter_unusedVariables___lambda__1___closed__2() { +static lean_object* _init_l_Lean_Linter_unusedVariables___elambda__1___lambda__1___closed__2() { _start: { lean_object* x_1; lean_object* x_2; @@ -9994,7 +9997,7 @@ x_2 = l_Lean_mkHashSetImp___rarg(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Linter_unusedVariables___lambda__1(lean_object* 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_EXPORT lean_object* l_Lean_Linter_unusedVariables___elambda__1___lambda__1(lean_object* 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) { _start: { uint8_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; lean_object* x_21; lean_object* x_22; uint8_t x_23; lean_object* x_24; @@ -10008,7 +10011,7 @@ x_16 = l_Lean_HashMap_toList___at_Lean_Server_ModuleRefs_instCoeModuleRefsModule x_17 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_17, 0, x_15); lean_ctor_set(x_17, 1, x_14); -x_18 = l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__14(x_16, x_17, x_8, x_9, x_10); +x_18 = l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__14(x_16, x_17, x_8, x_9, x_10); x_19 = lean_ctor_get(x_18, 0); lean_inc(x_19); x_20 = lean_ctor_get(x_18, 1); @@ -10023,7 +10026,7 @@ x_23 = lean_nat_dec_le(x_3, x_3); if (x_6 == 0) { lean_object* x_78; -x_78 = l_Lean_Linter_unusedVariables___lambda__1___closed__1; +x_78 = l_Lean_Linter_unusedVariables___elambda__1___lambda__1___closed__1; x_24 = x_78; goto block_77; } @@ -10032,8 +10035,8 @@ else size_t x_79; size_t x_80; lean_object* x_81; lean_object* x_82; x_79 = lean_usize_of_nat(x_3); x_80 = 0; -x_81 = l_Lean_Linter_unusedVariables___lambda__1___closed__1; -x_82 = l_Array_foldrMUnsafe_fold___at_Lean_Linter_unusedVariables___spec__41(x_2, x_79, x_80, x_81); +x_81 = l_Lean_Linter_unusedVariables___elambda__1___lambda__1___closed__1; +x_82 = l_Array_foldrMUnsafe_fold___at_Lean_Linter_unusedVariables___elambda__1___spec__41(x_2, x_79, x_80, x_81); x_24 = x_82; goto block_77; } @@ -10051,7 +10054,7 @@ if (x_63 == 0) lean_object* x_64; lean_dec(x_61); lean_dec(x_60); -x_64 = l_Lean_Linter_unusedVariables___lambda__1___closed__2; +x_64 = l_Lean_Linter_unusedVariables___elambda__1___lambda__1___closed__2; x_25 = x_64; x_26 = x_20; goto block_59; @@ -10065,7 +10068,7 @@ if (x_65 == 0) lean_object* x_66; lean_dec(x_61); lean_dec(x_60); -x_66 = l_Lean_Linter_unusedVariables___lambda__1___closed__2; +x_66 = l_Lean_Linter_unusedVariables___elambda__1___lambda__1___closed__2; x_25 = x_66; x_26 = x_20; goto block_59; @@ -10076,10 +10079,10 @@ size_t x_67; size_t x_68; lean_object* x_69; lean_object* x_70; x_67 = 0; x_68 = lean_usize_of_nat(x_61); lean_dec(x_61); -x_69 = l_Lean_Linter_unusedVariables___lambda__1___closed__2; +x_69 = l_Lean_Linter_unusedVariables___elambda__1___lambda__1___closed__2; lean_inc(x_9); lean_inc(x_8); -x_70 = l_Array_foldlMUnsafe_fold___at_Lean_Linter_unusedVariables___spec__40(x_60, x_67, x_68, x_69, x_8, x_9, x_20); +x_70 = l_Array_foldlMUnsafe_fold___at_Lean_Linter_unusedVariables___elambda__1___spec__40(x_60, x_67, x_68, x_69, x_8, x_9, x_20); lean_dec(x_60); if (lean_obj_tag(x_70) == 0) { @@ -10142,11 +10145,11 @@ x_30 = lean_alloc_closure((void*)(l_Lean_Linter_unusedVariables_isTopLevelDecl__ lean_closure_set(x_30, 0, x_21); x_31 = lean_unsigned_to_nat(0u); x_32 = l_Array_insertAt_x21___rarg(x_28, x_31, x_30); -x_33 = l_Lean_HashMap_toList___at_Lean_Linter_unusedVariables___spec__21(x_22); +x_33 = l_Lean_HashMap_toList___at_Lean_Linter_unusedVariables___elambda__1___spec__21(x_22); x_34 = l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_162____closed__1; lean_inc(x_9); lean_inc(x_8); -x_35 = l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35(x_4, x_5, x_2, x_3, x_6, x_23, x_25, x_32, x_33, x_34, x_8, x_9, x_29); +x_35 = l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35(x_4, x_5, x_2, x_3, x_6, x_23, x_25, x_32, x_33, x_34, x_8, x_9, x_29); lean_dec(x_2); if (lean_obj_tag(x_35) == 0) { @@ -10160,14 +10163,14 @@ x_38 = lean_array_get_size(x_36); x_39 = lean_unsigned_to_nat(1u); x_40 = lean_nat_sub(x_38, x_39); lean_dec(x_38); -x_41 = l_Array_qsort_sort___at_Lean_Linter_unusedVariables___spec__36(x_36, x_31, x_40); +x_41 = l_Array_qsort_sort___at_Lean_Linter_unusedVariables___elambda__1___spec__36(x_36, x_31, x_40); lean_dec(x_40); x_42 = lean_array_get_size(x_41); x_43 = lean_usize_of_nat(x_42); lean_dec(x_42); x_44 = 0; x_45 = lean_box(0); -x_46 = l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___spec__38(x_41, x_43, x_44, x_45, x_8, x_9, x_37); +x_46 = l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__38(x_41, x_43, x_44, x_45, x_8, x_9, x_37); lean_dec(x_9); lean_dec(x_8); lean_dec(x_41); @@ -10251,7 +10254,7 @@ return x_58; } } } -LEAN_EXPORT lean_object* l_Lean_Linter_unusedVariables___lambda__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_Linter_unusedVariables___elambda__1___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { uint8_t x_6; lean_object* x_7; @@ -10323,7 +10326,7 @@ else size_t x_30; size_t x_31; lean_object* x_32; x_30 = 0; x_31 = lean_usize_of_nat(x_19); -x_32 = l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___spec__42(x_17, x_30, x_31, x_3, x_4, x_13); +x_32 = l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___elambda__1___spec__42(x_17, x_30, x_31, x_3, x_4, x_13); if (lean_obj_tag(x_32) == 0) { lean_object* x_33; lean_object* x_34; uint8_t x_35; @@ -10377,7 +10380,7 @@ if (x_22 == 0) lean_object* x_24; lean_object* x_25; lean_dec(x_14); x_24 = lean_box(0); -x_25 = l_Lean_Linter_unusedVariables___lambda__1(x_18, x_17, x_19, x_1, x_10, x_21, x_24, x_3, x_4, x_23); +x_25 = l_Lean_Linter_unusedVariables___elambda__1___lambda__1(x_18, x_17, x_19, x_1, x_10, x_21, x_24, x_3, x_4, x_23); lean_dec(x_10); return x_25; } @@ -10405,7 +10408,7 @@ return x_27; } } } -LEAN_EXPORT lean_object* l_Lean_Linter_unusedVariables___lambda__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_Linter_unusedVariables___elambda__1___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; uint8_t x_7; @@ -10426,7 +10429,7 @@ if (x_11 == 0) lean_object* x_12; lean_object* x_13; lean_free_object(x_6); x_12 = lean_box(0); -x_13 = l_Lean_Linter_unusedVariables___lambda__2(x_1, x_12, x_3, x_4, x_9); +x_13 = l_Lean_Linter_unusedVariables___elambda__1___lambda__2(x_1, x_12, x_3, x_4, x_9); return x_13; } else @@ -10456,7 +10459,7 @@ if (x_18 == 0) { lean_object* x_19; lean_object* x_20; x_19 = lean_box(0); -x_20 = l_Lean_Linter_unusedVariables___lambda__2(x_1, x_19, x_3, x_4, x_16); +x_20 = l_Lean_Linter_unusedVariables___elambda__1___lambda__2(x_1, x_19, x_3, x_4, x_16); return x_20; } else @@ -10474,7 +10477,7 @@ return x_22; } } } -LEAN_EXPORT lean_object* l_Lean_Linter_unusedVariables(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Linter_unusedVariables___elambda__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; @@ -10512,7 +10515,7 @@ else lean_object* x_16; lean_object* x_17; lean_free_object(x_5); x_16 = lean_box(0); -x_17 = l_Lean_Linter_unusedVariables___lambda__3(x_1, x_16, x_2, x_3, x_8); +x_17 = l_Lean_Linter_unusedVariables___elambda__1___lambda__3(x_1, x_16, x_2, x_3, x_8); return x_17; } } @@ -10552,103 +10555,142 @@ else { lean_object* x_28; lean_object* x_29; x_28 = lean_box(0); -x_29 = l_Lean_Linter_unusedVariables___lambda__3(x_1, x_28, x_2, x_3, x_19); +x_29 = l_Lean_Linter_unusedVariables___elambda__1___lambda__3(x_1, x_28, x_2, x_3, x_19); return x_29; } } } } -LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_Linter_unusedVariables___spec__1___boxed(lean_object* x_1) { +static lean_object* _init_l_Lean_Linter_unusedVariables___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_7____closed__7; +x_2 = l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_7____closed__8; +x_3 = l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_7____closed__2; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Linter_unusedVariables___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Linter_unusedVariables___elambda__1), 4, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Linter_unusedVariables___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Linter_unusedVariables___closed__2; +x_2 = l_Lean_Linter_unusedVariables___closed__1; +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_Linter_unusedVariables() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Linter_unusedVariables___closed__3; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_Linter_unusedVariables___elambda__1___spec__1___boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_mkHashMap___at_Lean_Linter_unusedVariables___spec__1(x_1); +x_2 = l_Lean_mkHashMap___at_Lean_Linter_unusedVariables___elambda__1___spec__1(x_1); lean_dec(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_mkHashSet___at_Lean_Linter_unusedVariables___spec__2___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_mkHashSet___at_Lean_Linter_unusedVariables___elambda__1___spec__2___boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_mkHashSet___at_Lean_Linter_unusedVariables___spec__2(x_1); +x_2 = l_Lean_mkHashSet___at_Lean_Linter_unusedVariables___elambda__1___spec__2(x_1); lean_dec(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_List_replace___at_Lean_Linter_unusedVariables___spec__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_List_replace___at_Lean_Linter_unusedVariables___elambda__1___spec__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_List_replace___at_Lean_Linter_unusedVariables___spec__7(x_1, x_2, x_3); +x_4 = l_List_replace___at_Lean_Linter_unusedVariables___elambda__1___spec__7(x_1, x_2, x_3); lean_dec(x_2); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_AssocList_contains___at_Lean_Linter_unusedVariables___spec__9___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_AssocList_contains___at_Lean_Linter_unusedVariables___elambda__1___spec__9___boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l_Lean_AssocList_contains___at_Lean_Linter_unusedVariables___spec__9(x_1, x_2); +x_3 = l_Lean_AssocList_contains___at_Lean_Linter_unusedVariables___elambda__1___spec__9(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_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__14___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_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__14___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_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__14(x_1, x_2, x_3, x_4, x_5); +x_6 = l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__14(x_1, x_2, x_3, x_4, x_5); lean_dec(x_4); lean_dec(x_3); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_HashMapImp_contains___at_Lean_Linter_unusedVariables___spec__15___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_HashMapImp_contains___at_Lean_Linter_unusedVariables___elambda__1___spec__15___boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l_Lean_HashMapImp_contains___at_Lean_Linter_unusedVariables___spec__15(x_1, x_2); +x_3 = l_Lean_HashMapImp_contains___at_Lean_Linter_unusedVariables___elambda__1___spec__15(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visited___at_Lean_Linter_unusedVariables___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_EXPORT lean_object* l_Lean_ForEachExprWhere_visited___at_Lean_Linter_unusedVariables___elambda__1___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) { _start: { lean_object* x_7; -x_7 = l_Lean_ForEachExprWhere_visited___at_Lean_Linter_unusedVariables___spec__19(x_1, x_2, x_3, x_4, x_5, x_6); +x_7 = l_Lean_ForEachExprWhere_visited___at_Lean_Linter_unusedVariables___elambda__1___spec__19(x_1, x_2, x_3, x_4, x_5, x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_checked___at_Lean_Linter_unusedVariables___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_EXPORT lean_object* l_Lean_ForEachExprWhere_checked___at_Lean_Linter_unusedVariables___elambda__1___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) { _start: { lean_object* x_7; -x_7 = l_Lean_ForEachExprWhere_checked___at_Lean_Linter_unusedVariables___spec__20(x_1, x_2, x_3, x_4, x_5, x_6); +x_7 = l_Lean_ForEachExprWhere_checked___at_Lean_Linter_unusedVariables___elambda__1___spec__20(x_1, x_2, x_3, x_4, x_5, x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___spec__22___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___elambda__1___spec__22___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___spec__22(x_1, x_2); +x_3 = l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___elambda__1___spec__22(x_1, x_2); lean_dec(x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Linter_unusedVariables___spec__23___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_Linter_unusedVariables___elambda__1___spec__23___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; @@ -10656,16 +10698,16 @@ 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_Linter_unusedVariables___spec__23(x_1, x_5, x_6, x_4); +x_7 = l_Array_foldlMUnsafe_fold___at_Lean_Linter_unusedVariables___elambda__1___spec__23(x_1, x_5, x_6, x_4); lean_dec(x_1); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___spec__25___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_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___elambda__1___spec__25___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_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___spec__25___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6); +x_7 = l_Lean_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___elambda__1___spec__25___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); @@ -10674,11 +10716,11 @@ lean_dec(x_1); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___spec__25___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_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___elambda__1___spec__25___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) { _start: { lean_object* x_9; -x_9 = l_Lean_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___spec__25___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_9 = l_Lean_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___elambda__1___spec__25___lambda__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_4); @@ -10688,7 +10730,7 @@ lean_dec(x_1); return x_9; } } -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___spec__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_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___elambda__1___spec__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) { _start: { size_t x_7; size_t x_8; uint8_t x_9; lean_object* x_10; @@ -10696,45 +10738,45 @@ x_7 = lean_unbox_usize(x_5); lean_dec(x_5); x_8 = lean_unbox_usize(x_6); lean_dec(x_6); -x_9 = l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___spec__30(x_1, x_2, x_3, x_4, x_7, x_8); +x_9 = l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___elambda__1___spec__30(x_1, x_2, x_3, x_4, x_7, x_8); lean_dec(x_4); x_10 = lean_box(x_9); return x_10; } } -LEAN_EXPORT lean_object* l_List_foldr___at_Lean_Linter_unusedVariables___spec__31___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_List_foldr___at_Lean_Linter_unusedVariables___elambda__1___spec__31___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l_List_foldr___at_Lean_Linter_unusedVariables___spec__31___lambda__1(x_1, x_2); +x_3 = l_List_foldr___at_Lean_Linter_unusedVariables___elambda__1___spec__31___lambda__1(x_1, x_2); lean_dec(x_1); x_4 = lean_box(x_3); return x_4; } } -LEAN_EXPORT lean_object* l_List_foldr___at_Lean_Linter_unusedVariables___spec__31___lambda__2___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_List_foldr___at_Lean_Linter_unusedVariables___elambda__1___spec__31___lambda__2___boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l_List_foldr___at_Lean_Linter_unusedVariables___spec__31___lambda__2(x_1, x_2); +x_3 = l_List_foldr___at_Lean_Linter_unusedVariables___elambda__1___spec__31___lambda__2(x_1, x_2); lean_dec(x_1); x_4 = lean_box(x_3); return x_4; } } -LEAN_EXPORT lean_object* l_List_foldr___at_Lean_Linter_unusedVariables___spec__31___boxed(lean_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_List_foldr___at_Lean_Linter_unusedVariables___elambda__1___spec__31___boxed(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; uint8_t x_8; lean_object* x_9; x_7 = lean_unbox(x_5); lean_dec(x_5); -x_8 = l_List_foldr___at_Lean_Linter_unusedVariables___spec__31(x_1, x_2, x_3, x_4, x_7, x_6); +x_8 = l_List_foldr___at_Lean_Linter_unusedVariables___elambda__1___spec__31(x_1, x_2, x_3, x_4, x_7, x_6); lean_dec(x_1); x_9 = lean_box(x_8); return x_9; } } -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___spec__32___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_anyMUnsafe_any___at_Lean_Linter_unusedVariables___elambda__1___spec__32___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, 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; @@ -10742,13 +10784,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_anyMUnsafe_any___at_Lean_Linter_unusedVariables___spec__32(x_1, x_2, x_3, x_4, x_5, x_11, x_12, x_8, x_9, x_10); +x_13 = l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___elambda__1___spec__32(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_1); return x_13; } } -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___spec__33___boxed(lean_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_anyMUnsafe_any___at_Lean_Linter_unusedVariables___elambda__1___spec__33___boxed(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; uint8_t x_9; lean_object* x_10; @@ -10756,13 +10798,13 @@ x_7 = lean_unbox_usize(x_5); lean_dec(x_5); x_8 = lean_unbox_usize(x_6); lean_dec(x_6); -x_9 = l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___spec__33(x_1, x_2, x_3, x_4, x_7, x_8); +x_9 = l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___elambda__1___spec__33(x_1, x_2, x_3, x_4, x_7, x_8); lean_dec(x_4); x_10 = lean_box(x_9); return x_10; } } -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___spec__34___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___elambda__1___spec__34___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; uint8_t x_7; lean_object* x_8; @@ -10770,24 +10812,24 @@ x_5 = lean_unbox_usize(x_3); lean_dec(x_3); x_6 = lean_unbox_usize(x_4); lean_dec(x_4); -x_7 = l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___spec__34(x_1, x_2, x_5, x_6); +x_7 = l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___elambda__1___spec__34(x_1, x_2, x_5, x_6); lean_dec(x_2); x_8 = lean_box(x_7); return x_8; } } -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35___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_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35___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_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +x_8 = l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); return x_8; } } -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35___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_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35___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: { uint8_t x_15; uint8_t x_16; lean_object* x_17; @@ -10795,23 +10837,23 @@ x_15 = lean_unbox(x_1); lean_dec(x_1); x_16 = lean_unbox(x_4); lean_dec(x_4); -x_17 = l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35___lambda__2(x_15, x_2, x_3, x_16, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +x_17 = l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35___lambda__2(x_15, x_2, x_3, x_16, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); lean_dec(x_11); lean_dec(x_9); return x_17; } } -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35___lambda__3___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35___lambda__3___boxed(lean_object* x_1) { _start: { uint8_t x_2; lean_object* x_3; -x_2 = l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35___lambda__3(x_1); +x_2 = l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35___lambda__3(x_1); lean_dec(x_1); x_3 = lean_box(x_2); return x_3; } } -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35___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, lean_object* x_15) { +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35___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, lean_object* x_15) { _start: { uint8_t x_16; uint8_t x_17; lean_object* x_18; @@ -10819,12 +10861,12 @@ x_16 = lean_unbox(x_1); lean_dec(x_1); x_17 = lean_unbox(x_4); lean_dec(x_4); -x_18 = l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35___lambda__4(x_16, x_2, x_3, x_17, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +x_18 = l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35___lambda__4(x_16, x_2, x_3, x_17, 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_9); return x_18; } } -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35___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) { +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35___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: { uint8_t x_16; uint8_t x_17; lean_object* x_18; @@ -10832,12 +10874,12 @@ x_16 = lean_unbox(x_2); lean_dec(x_2); x_17 = lean_unbox(x_5); lean_dec(x_5); -x_18 = l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35___lambda__5(x_1, x_16, x_3, x_4, x_17, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +x_18 = l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35___lambda__5(x_1, x_16, x_3, x_4, x_17, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); lean_dec(x_9); return x_18; } } -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35___lambda__6___boxed(lean_object** _args) { +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35___lambda__6___boxed(lean_object** _args) { lean_object* x_1 = _args[0]; lean_object* x_2 = _args[1]; lean_object* x_3 = _args[2]; @@ -10863,12 +10905,12 @@ x_19 = lean_unbox(x_2); lean_dec(x_2); x_20 = lean_unbox(x_5); lean_dec(x_5); -x_21 = l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35___lambda__6(x_1, x_19, x_3, x_4, x_20, 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_21 = l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35___lambda__6(x_1, x_19, x_3, x_4, x_20, 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_9); return x_21; } } -LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, 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_Linter_unusedVariables___elambda__1___spec__35___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, 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; uint8_t x_15; lean_object* x_16; @@ -10876,44 +10918,44 @@ x_14 = lean_unbox(x_5); lean_dec(x_5); x_15 = lean_unbox(x_6); lean_dec(x_6); -x_16 = l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35(x_1, x_2, x_3, x_4, x_14, x_15, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +x_16 = l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35(x_1, x_2, x_3, x_4, x_14, x_15, x_7, x_8, x_9, x_10, x_11, x_12, x_13); lean_dec(x_3); lean_dec(x_2); return x_16; } } -LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Linter_unusedVariables___spec__36___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Linter_unusedVariables___elambda__1___spec__36___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l_Array_qsort_sort___at_Lean_Linter_unusedVariables___spec__36___lambda__1(x_1, x_2); +x_3 = l_Array_qsort_sort___at_Lean_Linter_unusedVariables___elambda__1___spec__36___lambda__1(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_qsort_sort___at_Lean_Linter_unusedVariables___spec__36___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Linter_unusedVariables___elambda__1___spec__36___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_Array_qsort_sort___at_Lean_Linter_unusedVariables___spec__36(x_1, x_2, x_3); +x_4 = l_Array_qsort_sort___at_Lean_Linter_unusedVariables___elambda__1___spec__36(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___spec__37___boxed(lean_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_Linter_logLint___at_Lean_Linter_unusedVariables___elambda__1___spec__37___boxed(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_Linter_logLint___at_Lean_Linter_unusedVariables___spec__37(x_1, x_2, x_3, x_4, x_5, x_6); +x_7 = l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___elambda__1___spec__37(x_1, x_2, x_3, x_4, x_5, x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); return x_7; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___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_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___elambda__1___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) { _start: { size_t x_8; size_t x_9; lean_object* x_10; @@ -10921,24 +10963,24 @@ 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_forInUnsafe_loop___at_Lean_Linter_unusedVariables___spec__38(x_1, x_8, x_9, x_4, x_5, x_6, x_7); +x_10 = l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__38(x_1, x_8, x_9, x_4, x_5, x_6, x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_1); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___spec__39___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_AssocList_foldlM___at_Lean_Linter_unusedVariables___elambda__1___spec__39___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_AssocList_foldlM___at_Lean_Linter_unusedVariables___spec__39___lambda__1(x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___elambda__1___spec__39___lambda__1(x_1, x_2, x_3, x_4, x_5); lean_dec(x_4); lean_dec(x_3); return x_6; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Linter_unusedVariables___spec__40___boxed(lean_object* x_1, lean_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_foldlMUnsafe_fold___at_Lean_Linter_unusedVariables___elambda__1___spec__40___boxed(lean_object* x_1, 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; @@ -10946,12 +10988,12 @@ 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_foldlMUnsafe_fold___at_Lean_Linter_unusedVariables___spec__40(x_1, x_8, x_9, x_4, x_5, x_6, x_7); +x_10 = l_Array_foldlMUnsafe_fold___at_Lean_Linter_unusedVariables___elambda__1___spec__40(x_1, x_8, x_9, x_4, x_5, x_6, x_7); lean_dec(x_1); return x_10; } } -LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Linter_unusedVariables___spec__41___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Linter_unusedVariables___elambda__1___spec__41___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; @@ -10959,12 +11001,12 @@ 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_foldrMUnsafe_fold___at_Lean_Linter_unusedVariables___spec__41(x_1, x_5, x_6, x_4); +x_7 = l_Array_foldrMUnsafe_fold___at_Lean_Linter_unusedVariables___elambda__1___spec__41(x_1, x_5, x_6, x_4); lean_dec(x_1); return x_7; } } -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___spec__42___boxed(lean_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_anyMUnsafe_any___at_Lean_Linter_unusedVariables___elambda__1___spec__42___boxed(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; @@ -10972,38 +11014,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_anyMUnsafe_any___at_Lean_Linter_unusedVariables___spec__42(x_1, x_7, x_8, x_4, x_5, x_6); +x_9 = l_Array_anyMUnsafe_any___at_Lean_Linter_unusedVariables___elambda__1___spec__42(x_1, x_7, x_8, x_4, x_5, x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_1); return x_9; } } -LEAN_EXPORT lean_object* l_Lean_Linter_unusedVariables___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_Lean_Linter_unusedVariables___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, 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_6); lean_dec(x_6); -x_12 = l_Lean_Linter_unusedVariables___lambda__1(x_1, x_2, x_3, x_4, x_5, x_11, x_7, x_8, x_9, x_10); +x_12 = l_Lean_Linter_unusedVariables___elambda__1___lambda__1(x_1, x_2, x_3, x_4, x_5, x_11, x_7, x_8, x_9, x_10); lean_dec(x_7); lean_dec(x_5); return x_12; } } -static lean_object* _init_l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_5176____closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Linter_unusedVariables), 4, 0); -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_5176_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_5191_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_5176____closed__1; +x_2 = l_Lean_Linter_unusedVariables; x_3 = l_Lean_Elab_Command_addLinter(x_2, x_1); return x_3; } @@ -11535,65 +11569,71 @@ l_Lean_Linter_unusedVariables_skipDeclIdIfPresent___closed__2 = _init_l_Lean_Lin lean_mark_persistent(l_Lean_Linter_unusedVariables_skipDeclIdIfPresent___closed__2); l_Lean_Linter_unusedVariables_isTopLevelDecl___closed__1 = _init_l_Lean_Linter_unusedVariables_isTopLevelDecl___closed__1(); lean_mark_persistent(l_Lean_Linter_unusedVariables_isTopLevelDecl___closed__1); -l_panic___at_Lean_Linter_unusedVariables___spec__28___closed__1 = _init_l_panic___at_Lean_Linter_unusedVariables___spec__28___closed__1(); -lean_mark_persistent(l_panic___at_Lean_Linter_unusedVariables___spec__28___closed__1); -l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___spec__27___closed__1 = _init_l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___spec__27___closed__1(); -lean_mark_persistent(l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___spec__27___closed__1); -l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___spec__27___closed__2 = _init_l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___spec__27___closed__2(); -lean_mark_persistent(l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___spec__27___closed__2); -l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___spec__27___closed__3 = _init_l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___spec__27___closed__3(); -lean_mark_persistent(l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___spec__27___closed__3); -l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___spec__27___closed__4 = _init_l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___spec__27___closed__4(); -lean_mark_persistent(l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___spec__27___closed__4); -l_Lean_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___spec__25___lambda__2___closed__1 = _init_l_Lean_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___spec__25___lambda__2___closed__1(); -lean_mark_persistent(l_Lean_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___spec__25___lambda__2___closed__1); -l_Lean_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___spec__25___closed__1 = _init_l_Lean_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___spec__25___closed__1(); -lean_mark_persistent(l_Lean_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___spec__25___closed__1); -l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35___lambda__4___closed__1 = _init_l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35___lambda__4___closed__1(); -lean_mark_persistent(l_List_forIn_loop___at_Lean_Linter_unusedVariables___spec__35___lambda__4___closed__1); -l_Array_qsort_sort___at_Lean_Linter_unusedVariables___spec__36___lambda__1___closed__1 = _init_l_Array_qsort_sort___at_Lean_Linter_unusedVariables___spec__36___lambda__1___closed__1(); -lean_mark_persistent(l_Array_qsort_sort___at_Lean_Linter_unusedVariables___spec__36___lambda__1___closed__1); -l_Array_qsort_sort___at_Lean_Linter_unusedVariables___spec__36___lambda__1___closed__2 = _init_l_Array_qsort_sort___at_Lean_Linter_unusedVariables___spec__36___lambda__1___closed__2(); -lean_mark_persistent(l_Array_qsort_sort___at_Lean_Linter_unusedVariables___spec__36___lambda__1___closed__2); -l_Array_qsort_sort___at_Lean_Linter_unusedVariables___spec__36___lambda__1___closed__3 = _init_l_Array_qsort_sort___at_Lean_Linter_unusedVariables___spec__36___lambda__1___closed__3(); -lean_mark_persistent(l_Array_qsort_sort___at_Lean_Linter_unusedVariables___spec__36___lambda__1___closed__3); -l_Array_qsort_sort___at_Lean_Linter_unusedVariables___spec__36___lambda__1___closed__4 = _init_l_Array_qsort_sort___at_Lean_Linter_unusedVariables___spec__36___lambda__1___closed__4(); -lean_mark_persistent(l_Array_qsort_sort___at_Lean_Linter_unusedVariables___spec__36___lambda__1___closed__4); -l_Array_qsort_sort___at_Lean_Linter_unusedVariables___spec__36___closed__1 = _init_l_Array_qsort_sort___at_Lean_Linter_unusedVariables___spec__36___closed__1(); -lean_mark_persistent(l_Array_qsort_sort___at_Lean_Linter_unusedVariables___spec__36___closed__1); -l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___spec__37___closed__1 = _init_l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___spec__37___closed__1(); -lean_mark_persistent(l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___spec__37___closed__1); -l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___spec__37___closed__2 = _init_l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___spec__37___closed__2(); -lean_mark_persistent(l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___spec__37___closed__2); -l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___spec__37___closed__3 = _init_l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___spec__37___closed__3(); -lean_mark_persistent(l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___spec__37___closed__3); -l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___spec__37___closed__4 = _init_l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___spec__37___closed__4(); -lean_mark_persistent(l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___spec__37___closed__4); -l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___spec__37___closed__5 = _init_l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___spec__37___closed__5(); -lean_mark_persistent(l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___spec__37___closed__5); -l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___spec__38___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___spec__38___closed__1(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___spec__38___closed__1); -l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___spec__38___closed__2 = _init_l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___spec__38___closed__2(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___spec__38___closed__2); -l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___spec__38___closed__3 = _init_l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___spec__38___closed__3(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___spec__38___closed__3); -l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___spec__38___closed__4 = _init_l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___spec__38___closed__4(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___spec__38___closed__4); -l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___spec__39___closed__1 = _init_l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___spec__39___closed__1(); -lean_mark_persistent(l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___spec__39___closed__1); -l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___spec__39___closed__2 = _init_l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___spec__39___closed__2(); -lean_mark_persistent(l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___spec__39___closed__2); -l_Array_foldrMUnsafe_fold___at_Lean_Linter_unusedVariables___spec__41___lambda__2___closed__1 = _init_l_Array_foldrMUnsafe_fold___at_Lean_Linter_unusedVariables___spec__41___lambda__2___closed__1(); -lean_mark_persistent(l_Array_foldrMUnsafe_fold___at_Lean_Linter_unusedVariables___spec__41___lambda__2___closed__1); -l_Array_foldrMUnsafe_fold___at_Lean_Linter_unusedVariables___spec__41___closed__1 = _init_l_Array_foldrMUnsafe_fold___at_Lean_Linter_unusedVariables___spec__41___closed__1(); -lean_mark_persistent(l_Array_foldrMUnsafe_fold___at_Lean_Linter_unusedVariables___spec__41___closed__1); -l_Lean_Linter_unusedVariables___lambda__1___closed__1 = _init_l_Lean_Linter_unusedVariables___lambda__1___closed__1(); -lean_mark_persistent(l_Lean_Linter_unusedVariables___lambda__1___closed__1); -l_Lean_Linter_unusedVariables___lambda__1___closed__2 = _init_l_Lean_Linter_unusedVariables___lambda__1___closed__2(); -lean_mark_persistent(l_Lean_Linter_unusedVariables___lambda__1___closed__2); -l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_5176____closed__1 = _init_l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_5176____closed__1(); -lean_mark_persistent(l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_5176____closed__1); -res = l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_5176_(lean_io_mk_world()); +l_panic___at_Lean_Linter_unusedVariables___elambda__1___spec__28___closed__1 = _init_l_panic___at_Lean_Linter_unusedVariables___elambda__1___spec__28___closed__1(); +lean_mark_persistent(l_panic___at_Lean_Linter_unusedVariables___elambda__1___spec__28___closed__1); +l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___elambda__1___spec__27___closed__1 = _init_l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___elambda__1___spec__27___closed__1(); +lean_mark_persistent(l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___elambda__1___spec__27___closed__1); +l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___elambda__1___spec__27___closed__2 = _init_l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___elambda__1___spec__27___closed__2(); +lean_mark_persistent(l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___elambda__1___spec__27___closed__2); +l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___elambda__1___spec__27___closed__3 = _init_l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___elambda__1___spec__27___closed__3(); +lean_mark_persistent(l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___elambda__1___spec__27___closed__3); +l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___elambda__1___spec__27___closed__4 = _init_l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___elambda__1___spec__27___closed__4(); +lean_mark_persistent(l_Lean_Elab_InfoTree_visitM_go___at_Lean_Linter_unusedVariables___elambda__1___spec__27___closed__4); +l_Lean_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___elambda__1___spec__25___lambda__2___closed__1 = _init_l_Lean_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___elambda__1___spec__25___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___elambda__1___spec__25___lambda__2___closed__1); +l_Lean_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___elambda__1___spec__25___closed__1 = _init_l_Lean_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___elambda__1___spec__25___closed__1(); +lean_mark_persistent(l_Lean_Linter_collectMacroExpansions_x3f_go___at_Lean_Linter_unusedVariables___elambda__1___spec__25___closed__1); +l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35___lambda__4___closed__1 = _init_l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35___lambda__4___closed__1(); +lean_mark_persistent(l_List_forIn_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__35___lambda__4___closed__1); +l_Array_qsort_sort___at_Lean_Linter_unusedVariables___elambda__1___spec__36___lambda__1___closed__1 = _init_l_Array_qsort_sort___at_Lean_Linter_unusedVariables___elambda__1___spec__36___lambda__1___closed__1(); +lean_mark_persistent(l_Array_qsort_sort___at_Lean_Linter_unusedVariables___elambda__1___spec__36___lambda__1___closed__1); +l_Array_qsort_sort___at_Lean_Linter_unusedVariables___elambda__1___spec__36___lambda__1___closed__2 = _init_l_Array_qsort_sort___at_Lean_Linter_unusedVariables___elambda__1___spec__36___lambda__1___closed__2(); +lean_mark_persistent(l_Array_qsort_sort___at_Lean_Linter_unusedVariables___elambda__1___spec__36___lambda__1___closed__2); +l_Array_qsort_sort___at_Lean_Linter_unusedVariables___elambda__1___spec__36___lambda__1___closed__3 = _init_l_Array_qsort_sort___at_Lean_Linter_unusedVariables___elambda__1___spec__36___lambda__1___closed__3(); +lean_mark_persistent(l_Array_qsort_sort___at_Lean_Linter_unusedVariables___elambda__1___spec__36___lambda__1___closed__3); +l_Array_qsort_sort___at_Lean_Linter_unusedVariables___elambda__1___spec__36___lambda__1___closed__4 = _init_l_Array_qsort_sort___at_Lean_Linter_unusedVariables___elambda__1___spec__36___lambda__1___closed__4(); +lean_mark_persistent(l_Array_qsort_sort___at_Lean_Linter_unusedVariables___elambda__1___spec__36___lambda__1___closed__4); +l_Array_qsort_sort___at_Lean_Linter_unusedVariables___elambda__1___spec__36___closed__1 = _init_l_Array_qsort_sort___at_Lean_Linter_unusedVariables___elambda__1___spec__36___closed__1(); +lean_mark_persistent(l_Array_qsort_sort___at_Lean_Linter_unusedVariables___elambda__1___spec__36___closed__1); +l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___elambda__1___spec__37___closed__1 = _init_l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___elambda__1___spec__37___closed__1(); +lean_mark_persistent(l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___elambda__1___spec__37___closed__1); +l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___elambda__1___spec__37___closed__2 = _init_l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___elambda__1___spec__37___closed__2(); +lean_mark_persistent(l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___elambda__1___spec__37___closed__2); +l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___elambda__1___spec__37___closed__3 = _init_l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___elambda__1___spec__37___closed__3(); +lean_mark_persistent(l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___elambda__1___spec__37___closed__3); +l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___elambda__1___spec__37___closed__4 = _init_l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___elambda__1___spec__37___closed__4(); +lean_mark_persistent(l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___elambda__1___spec__37___closed__4); +l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___elambda__1___spec__37___closed__5 = _init_l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___elambda__1___spec__37___closed__5(); +lean_mark_persistent(l_Lean_Linter_logLint___at_Lean_Linter_unusedVariables___elambda__1___spec__37___closed__5); +l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__38___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__38___closed__1(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__38___closed__1); +l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__38___closed__2 = _init_l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__38___closed__2(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__38___closed__2); +l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__38___closed__3 = _init_l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__38___closed__3(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__38___closed__3); +l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__38___closed__4 = _init_l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__38___closed__4(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Linter_unusedVariables___elambda__1___spec__38___closed__4); +l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___elambda__1___spec__39___closed__1 = _init_l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___elambda__1___spec__39___closed__1(); +lean_mark_persistent(l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___elambda__1___spec__39___closed__1); +l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___elambda__1___spec__39___closed__2 = _init_l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___elambda__1___spec__39___closed__2(); +lean_mark_persistent(l_Lean_AssocList_foldlM___at_Lean_Linter_unusedVariables___elambda__1___spec__39___closed__2); +l_Array_foldrMUnsafe_fold___at_Lean_Linter_unusedVariables___elambda__1___spec__41___lambda__2___closed__1 = _init_l_Array_foldrMUnsafe_fold___at_Lean_Linter_unusedVariables___elambda__1___spec__41___lambda__2___closed__1(); +lean_mark_persistent(l_Array_foldrMUnsafe_fold___at_Lean_Linter_unusedVariables___elambda__1___spec__41___lambda__2___closed__1); +l_Array_foldrMUnsafe_fold___at_Lean_Linter_unusedVariables___elambda__1___spec__41___closed__1 = _init_l_Array_foldrMUnsafe_fold___at_Lean_Linter_unusedVariables___elambda__1___spec__41___closed__1(); +lean_mark_persistent(l_Array_foldrMUnsafe_fold___at_Lean_Linter_unusedVariables___elambda__1___spec__41___closed__1); +l_Lean_Linter_unusedVariables___elambda__1___lambda__1___closed__1 = _init_l_Lean_Linter_unusedVariables___elambda__1___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Linter_unusedVariables___elambda__1___lambda__1___closed__1); +l_Lean_Linter_unusedVariables___elambda__1___lambda__1___closed__2 = _init_l_Lean_Linter_unusedVariables___elambda__1___lambda__1___closed__2(); +lean_mark_persistent(l_Lean_Linter_unusedVariables___elambda__1___lambda__1___closed__2); +l_Lean_Linter_unusedVariables___closed__1 = _init_l_Lean_Linter_unusedVariables___closed__1(); +lean_mark_persistent(l_Lean_Linter_unusedVariables___closed__1); +l_Lean_Linter_unusedVariables___closed__2 = _init_l_Lean_Linter_unusedVariables___closed__2(); +lean_mark_persistent(l_Lean_Linter_unusedVariables___closed__2); +l_Lean_Linter_unusedVariables___closed__3 = _init_l_Lean_Linter_unusedVariables___closed__3(); +lean_mark_persistent(l_Lean_Linter_unusedVariables___closed__3); +l_Lean_Linter_unusedVariables = _init_l_Lean_Linter_unusedVariables(); +lean_mark_persistent(l_Lean_Linter_unusedVariables); +res = l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_5191_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); l_Lean_MessageData_isUnusedVariableWarning___closed__1 = _init_l_Lean_MessageData_isUnusedVariableWarning___closed__1(); diff --git a/stage0/stdlib/Lean/Log.c b/stage0/stdlib/Lean/Log.c index df0ed49fce2d..301f495186a7 100644 --- a/stage0/stdlib/Lean/Log.c +++ b/stage0/stdlib/Lean/Log.c @@ -13,7 +13,6 @@ #ifdef __cplusplus extern "C" { #endif -uint8_t l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_log(lean_object*); LEAN_EXPORT lean_object* l_Lean_logAt___rarg___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_logAt___rarg___lambda__14___closed__1; @@ -61,6 +60,7 @@ lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_logAt___rarg___lambda__14(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_logAt___rarg___lambda__10(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_warningAsError; +uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_logErrorAt___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_logUnknownDecl___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_getRefPos(lean_object*); @@ -68,7 +68,6 @@ static lean_object* l_Lean_initFn____x40_Lean_Log___hyg_214____closed__4; LEAN_EXPORT lean_object* l_Lean_logWarning___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_logAt___rarg___lambda__11(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); static lean_object* l_Lean_logUnknownDecl___rarg___closed__3; -lean_object* l_Lean_Option_register___at_Std_Format_initFn____x40_Lean_Data_Format___hyg_95____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_logUnknownDecl(lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Log___hyg_214____closed__5; LEAN_EXPORT lean_object* l_Lean_logAt___rarg___lambda__4(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); @@ -81,6 +80,7 @@ static lean_object* l_Lean_logUnknownDecl___rarg___closed__4; LEAN_EXPORT lean_object* l_Lean_logAt___rarg___lambda__3(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_logAt___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_logAt___rarg___lambda__7(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); +lean_object* l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_logAt___rarg___lambda__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_logInfo(lean_object*); LEAN_EXPORT lean_object* l_Lean_logInfoAt___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -354,7 +354,7 @@ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_2 = l_Lean_initFn____x40_Lean_Log___hyg_214____closed__2; x_3 = l_Lean_initFn____x40_Lean_Log___hyg_214____closed__5; x_4 = l_Lean_initFn____x40_Lean_Log___hyg_214____closed__7; -x_5 = l_Lean_Option_register___at_Std_Format_initFn____x40_Lean_Data_Format___hyg_95____spec__1(x_2, x_3, x_4, x_1); +x_5 = l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__1(x_2, x_3, x_4, x_1); return x_5; } } @@ -721,7 +721,7 @@ lean_object* x_14; lean_object* x_15; uint8_t x_16; x_14 = lean_ctor_get(x_2, 1); lean_inc(x_14); x_15 = l_Lean_logAt___rarg___lambda__14___closed__1; -x_16 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_7, x_15); +x_16 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_7, x_15); if (x_16 == 0) { lean_object* x_17; lean_object* x_18; lean_object* x_19; @@ -1121,7 +1121,7 @@ LEAN_EXPORT lean_object* l_Lean_logWarning___rarg___lambda__1(lean_object* x_1, { lean_object* x_7; uint8_t x_8; x_7 = l_Lean_logAt___rarg___lambda__14___closed__1; -x_8 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_6, x_7); +x_8 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_6, x_7); if (x_8 == 0) { uint8_t x_9; lean_object* x_10; diff --git a/stage0/stdlib/Lean/Message.c b/stage0/stdlib/Lean/Message.c index 4ac7f7624738..0616defbf8d3 100644 --- a/stage0/stdlib/Lean/Message.c +++ b/stage0/stdlib/Lean/Message.c @@ -135,6 +135,7 @@ static lean_object* l_Lean_MessageData_ofList___closed__6; LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at_Lean_MessageLog_toList___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_KernelException_toMessageData___closed__1; static lean_object* l_Lean_KernelException_toMessageData___closed__30; +lean_object* l_Lean_Syntax_copyHeadTailInfoFrom(lean_object*, lean_object*); static lean_object* l_Lean_MessageData_formatAux___closed__8; static lean_object* l_Lean_MessageData_formatAux___closed__4; LEAN_EXPORT lean_object* l_String_split___at_Lean_stringToMessageData___spec__1(lean_object*); @@ -1156,16 +1157,18 @@ return x_1; LEAN_EXPORT lean_object* l_Lean_MessageData_ofSyntax(lean_object* x_1) { _start: { -lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_2 = lean_alloc_closure((void*)(l_Lean_MessageData_ofSyntax___lambda__1), 3, 1); -lean_closure_set(x_2, 0, x_1); -x_3 = l_Lean_MessageData_ofSyntax___closed__1; -x_4 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_4, 0, x_2); -lean_ctor_set(x_4, 1, x_3); -x_5 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_5, 0, x_4); -return x_5; +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_Syntax_copyHeadTailInfoFrom(x_1, x_2); +x_4 = lean_alloc_closure((void*)(l_Lean_MessageData_ofSyntax___lambda__1), 3, 1); +lean_closure_set(x_4, 0, x_3); +x_5 = l_Lean_MessageData_ofSyntax___closed__1; +x_6 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_6, 0, x_4); +lean_ctor_set(x_6, 1, x_5); +x_7 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_7, 0, x_6); +return x_7; } } LEAN_EXPORT lean_object* l_Lean_MessageData_ofSyntax___lambda__2___boxed(lean_object* x_1) { diff --git a/stage0/stdlib/Lean/Meta/AbstractMVars.c b/stage0/stdlib/Lean/Meta/AbstractMVars.c index 32e3e74f1fc0..c385dbd25d32 100644 --- a/stage0/stdlib/Lean/Meta/AbstractMVars.c +++ b/stage0/stdlib/Lean/Meta/AbstractMVars.c @@ -66,7 +66,6 @@ lean_object* l_Lean_Expr_forallE___override(lean_object*, lean_object*, lean_obj LEAN_EXPORT lean_object* l_Lean_Meta_AbstractMVars_State_fvars___default; lean_object* lean_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_AbstractMVars_instMonadMCtxM___lambda__2(lean_object*, lean_object*); -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_abstractMVars(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_AssocList_contains___at___private_Lean_Meta_AbstractMVars_0__Lean_Meta_AbstractMVars_abstractLevelMVars___spec__4(lean_object*, lean_object*); lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); @@ -98,6 +97,7 @@ lean_object* l_Lean_Expr_app___override(lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_AbstractMVars_mkFreshFVarId(lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_AbstractMVars_instMonadMCtxM___lambda__1___boxed(lean_object*, lean_object*); uint8_t l___private_Lean_Expr_0__Lean_beqBinderInfo____x40_Lean_Expr___hyg_473_(uint8_t, uint8_t); static lean_object* l___private_Lean_Meta_AbstractMVars_0__Lean_Meta_AbstractMVars_abstractLevelMVars___closed__2; @@ -3352,7 +3352,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_abstractMVars(lean_object* x_1, lean_object _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; 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_7 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_1, x_2, x_3, x_4, x_5, x_6); +x_7 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_1, x_2, x_3, x_4, x_5, x_6); x_8 = lean_ctor_get(x_7, 0); lean_inc(x_8); x_9 = lean_ctor_get(x_7, 1); diff --git a/stage0/stdlib/Lean/Meta/AppBuilder.c b/stage0/stdlib/Lean/Meta/AppBuilder.c index 6d5f14d5abf8..3359431f3920 100644 --- a/stage0/stdlib/Lean/Meta/AppBuilder.c +++ b/stage0/stdlib/Lean/Meta/AppBuilder.c @@ -29,17 +29,20 @@ LEAN_EXPORT lean_object* l_Lean_Meta_mkEqOfHEq___lambda__1(lean_object*, lean_ob static lean_object* l_Lean_Meta_mkProjection___lambda__1___closed__5; static lean_object* l_Lean_Meta_mkDefault___closed__2; static lean_object* l___private_Lean_Meta_AppBuilder_0__Lean_Meta_mkAppMFinal___closed__2; -static lean_object* l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___closed__3; static lean_object* l_panic___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_mkAppMFinal___spec__5___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___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_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mkAppOptM___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_mkLetValCongr___closed__2; static lean_object* l_Lean_Meta_mkAdd___closed__1; +static double l_Lean_withSeconds___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_mkSub(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_mkEqRefl___closed__1; static lean_object* l_Lean_Meta_mkSorry___closed__1; +extern lean_object* l_Lean_profiler; LEAN_EXPORT lean_object* l_Lean_Meta_mkIffOfEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___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_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__3___closed__5; 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_mkAppOptM(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_mkFunExt___closed__2; @@ -56,6 +59,7 @@ lean_object* l_Lean_Meta_mkFreshLevelMVar(lean_object*, lean_object*, lean_objec lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mkEqRec(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mkEqTrue(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__3___closed__3; static lean_object* l_Lean_Meta_mkImpDepCongrCtx___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_mkNoConfusion(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_mkEqSymm___closed__5; @@ -80,7 +84,6 @@ static lean_object* l___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilder static lean_object* l_Lean_Meta_mkMul___closed__3; static lean_object* l_Lean_Meta_mkAdd___closed__4; 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_mkProjection___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Meta_AppBuilder_0__Lean_Meta_mkAppMArgs_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_EXPORT lean_object* l_Lean_Meta_mkSyntheticSorry(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_mkMul___closed__2; @@ -89,6 +92,7 @@ uint8_t l_Lean_Expr_isAppOfArity(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_mkOfEqTrue___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_AppBuilder_0__Lean_Meta_throwAppBuilderException(lean_object*); lean_object* l_Lean_Name_toString(lean_object*, uint8_t); +double l_Lean_trace_profiler_threshold_getSecs(lean_object*); static lean_object* l_Lean_Meta_mkEqRec___closed__2; static lean_object* l_Lean_Meta_mkEqNDRec___closed__4; lean_object* l_Lean_Expr_sort___override(lean_object*); @@ -108,19 +112,22 @@ static lean_object* l_Lean_Meta_mkNumeral___closed__3; static lean_object* l_Lean_Meta_mkIffOfEq___closed__1; static lean_object* l___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___rarg___lambda__1___closed__8; static lean_object* l_Lean_Meta_mkEqMPR___closed__2; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__4(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_EXPORT lean_object* l___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace(lean_object*, lean_object*); -static lean_object* l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___closed__1; static lean_object* l_Lean_Meta_mkCongr___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_AppBuilder_0__Lean_Meta_infer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mkExpectedTypeHint(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_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___rarg___lambda__3___closed__6; static lean_object* l_Lean_Meta_mkNoConfusion___closed__5; +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_mkForallCongr___closed__1; +static lean_object* l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__4___closed__1; static lean_object* l_Lean_Meta_mkImpCongr___closed__1; static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_AppBuilder___hyg_7226____closed__7; LEAN_EXPORT lean_object* l___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___at_Lean_Meta_mkAppOptM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_environment_find(lean_object*, lean_object*); +uint8_t lean_float_decLt(double, double); static lean_object* l_Lean_Meta_mkSorry___closed__7; static lean_object* l___private_Lean_Meta_AppBuilder_0__Lean_Meta_mkAppOptMAux___closed__4; static lean_object* l___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___rarg___lambda__1___closed__7; @@ -143,8 +150,8 @@ static lean_object* l___private_Lean_Meta_AppBuilder_0__Lean_Meta_throwAppBuilde LEAN_EXPORT lean_object* l___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkApp4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_mkHEqSymm___closed__2; -static double l_Lean_withSeconds___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__3___closed__1; lean_object* l_Lean_Exception_toMessageData(lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__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_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___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* l_Lean_Meta_getDecLevel(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_AppBuilder_0__Lean_Meta_mkAppMArgs_loop___closed__6; @@ -159,6 +166,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_mkAdd(lean_object*, lean_object*, lean_obje static lean_object* l_Lean_Meta_mkCongrArg___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_AppBuilder___hyg_7226_(lean_object*); lean_object* l_Lean_getStructureFields(lean_object*, lean_object*); +static lean_object* l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__3___closed__2; static lean_object* l_Lean_Meta_mkNoConfusion___closed__7; uint8_t l_Lean_Expr_hasMVar(lean_object*); static lean_object* l_Lean_Meta_mkDecideProof___closed__2; @@ -186,6 +194,7 @@ static lean_object* l_Lean_Meta_mkPure___closed__3; static lean_object* l_Lean_Meta_mkLt___closed__1; static lean_object* l_Lean_Meta_mkFunExt___closed__1; static lean_object* l_Lean_Meta_mkEqFalse_x27___closed__1; +lean_object* lean_st_ref_take(lean_object*, lean_object*); lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_processPostponed___spec__2___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Meta_mkDecide___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___at_Lean_Meta_mkAppM_x27___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -195,6 +204,7 @@ static lean_object* l___private_Lean_Meta_AppBuilder_0__Lean_Meta_mkAppOptMAux__ LEAN_EXPORT lean_object* l_Lean_Meta_mkImpDepCongrCtx(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_mkEqTrans___closed__1; static lean_object* l_Lean_Meta_mkSub___closed__2; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___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 lean_object* l___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___rarg___lambda__3___closed__1; static lean_object* l___private_Lean_Meta_AppBuilder_0__Lean_Meta_mkAppOptMAux___closed__10; LEAN_EXPORT lean_object* l___private_Lean_Meta_AppBuilder_0__Lean_Meta_mkAppMArgs(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -219,6 +229,7 @@ static lean_object* l_Lean_Meta_mkAbsurd___closed__2; static lean_object* l_Lean_Meta_mkOfNonempty___closed__1; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_mkAppMFinal___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_Lean_MVarId_isAssignable___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_mkAppMFinal___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_PersistentArray_append___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Meta_mkSorry___closed__2; static lean_object* l___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___rarg___lambda__1___closed__5; LEAN_EXPORT lean_object* l_Lean_Meta_mkLetCongr(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -234,14 +245,12 @@ static lean_object* l_Lean_Meta_isMonad_x3f___closed__2; static lean_object* l_Lean_Meta_mkAbsurd___closed__1; extern lean_object* l_Lean_checkEmoji; lean_object* lean_st_ref_get(lean_object*, lean_object*); -static lean_object* l_Lean_withOptProfile___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_mkEqTrans(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mkAppM(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mkAppM___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_mkEqFalse_x27___closed__2; static lean_object* l_Lean_Meta_mkProjection___lambda__1___closed__1; -static lean_object* l_Lean_withOptProfile___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__2___closed__2; static lean_object* l_Lean_Meta_mkFalseElim___closed__2; static lean_object* l_Lean_Meta_mkHEq___closed__1; static lean_object* l_Lean_Meta_mkPure___closed__2; @@ -251,7 +260,6 @@ static lean_object* l_Lean_Meta_mkNoConfusion___closed__8; lean_object* lean_array_to_list(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mkEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_mkIffOfEq___closed__3; -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_isLevelMVarAssignable___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_mkAppMFinal___spec__4___closed__3; uint8_t l_Lean_Expr_hasLooseBVars(lean_object*); lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); @@ -265,6 +273,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_mkId(lean_object*, lean_object*, lean_objec extern lean_object* l_Lean_levelZero; lean_object* lean_io_mono_nanos_now(lean_object*); static lean_object* l___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___rarg___lambda__3___closed__5; +static lean_object* l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__3___closed__4; uint8_t l_Lean_Level_hasMVar(lean_object*); static lean_object* l_Lean_Meta_mkEq___closed__1; extern lean_object* l_Lean_instInhabitedExpr; @@ -276,12 +285,13 @@ LEAN_EXPORT lean_object* l_Lean_Meta_mkDecide(lean_object*, lean_object*, lean_o static lean_object* l_Lean_Meta_mkCongr___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_mkForallCongr(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_mkEqNDRec___closed__6; -lean_object* l_Lean_getBoolOption___at_Lean_Meta_processPostponed___spec__4(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_mkSorry___closed__5; static lean_object* l___private_Lean_Meta_AppBuilder_0__Lean_Meta_mkAppMArgs_loop___closed__9; +static lean_object* l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__3___closed__1; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_AppBuilder___hyg_7226____closed__11; static lean_object* l___private_Lean_Meta_AppBuilder_0__Lean_Meta_throwAppBuilderException___rarg___closed__1; +uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mkEqOfHEq___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_Meta_mkLt___closed__2; static lean_object* l_Lean_Meta_mkHEqSymm___closed__1; @@ -301,7 +311,6 @@ lean_object* l_Lean_Expr_appFn_x21(lean_object*); extern lean_object* l_Lean_Meta_instMonadMetaM; LEAN_EXPORT lean_object* l___private_Lean_Meta_AppBuilder_0__Lean_Meta_throwAppBuilderException___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_mkSub___closed__4; -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_mkAdd___closed__2; static lean_object* l_Lean_Meta_mkNoConfusion___closed__1; static lean_object* l_Lean_Meta_mkPure___closed__1; @@ -319,6 +328,7 @@ lean_object* l_Array_append___rarg(lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofExpr(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mkAppOptM___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_mkEqFalse___closed__2; +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_mkListLit___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_mkProjection___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_Lean_Meta_mkDecideProof___closed__3; @@ -330,7 +340,7 @@ static lean_object* l_Lean_Meta_mkLe___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_AppBuilder_0__Lean_Meta_mkAppMArgs___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mkCongr(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mkImpCongr(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_mkProjection___closed__5; +lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_processPostponed___spec__4(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___private_Lean_Meta_AppBuilder_0__Lean_Meta_mkAppMArgs_loop___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_throwError___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_throwAppBuilderException___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___rarg___lambda__3___closed__2; @@ -344,7 +354,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_mkLT(lean_object*, lean_object*, lean_objec static lean_object* l_Lean_Meta_mkPropExt___closed__1; lean_object* l_Lean_Expr_app___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_AppBuilder_0__Lean_Meta_mkAppOptMAux___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_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___closed__4; LEAN_EXPORT lean_object* l_List_mapM_loop___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_mkFun___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_mkDefault___closed__3; lean_object* l_Lean_MVarId_getDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -359,6 +368,7 @@ static lean_object* l___private_Lean_Meta_AppBuilder_0__Lean_Meta_mkAppMArgs_loo static lean_object* l_Lean_Meta_mkFalseElim___closed__1; static lean_object* l_Lean_Meta_mkImpDepCongrCtx___closed__1; lean_object* lean_float_to_string(double); +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkRawNatLit(lean_object*); static lean_object* l_List_mapTR_loop___at_Lean_Meta_mkAppOptM___spec__2___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Meta_AppBuilder_0__Lean_Meta_mkAppMFinal___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -371,8 +381,8 @@ LEAN_EXPORT lean_object* l_Lean_Meta_mkSorry___boxed(lean_object*, lean_object*, lean_object* l_Lean_indentExpr(lean_object*); static lean_object* l_Lean_Meta_mkEqOfHEq___closed__1; static lean_object* l___private_Lean_Meta_AppBuilder_0__Lean_Meta_hasTypeMsg___closed__2; +LEAN_EXPORT lean_object* l_Lean_withSeconds___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_mkLetCongr___closed__1; -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_mkNoConfusion___closed__10; static lean_object* l___private_Lean_Meta_AppBuilder_0__Lean_Meta_mkAppMArgs_loop___closed__3; static lean_object* l_Lean_Meta_mkEqTrue___closed__1; @@ -384,7 +394,6 @@ static lean_object* l___private_Lean_Meta_AppBuilder_0__Lean_Meta_hasTypeMsg___c static lean_object* l_Lean_Meta_mkSorry___closed__3; static lean_object* l_Lean_Meta_mkProjection___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_withSeconds___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_AppBuilder_0__Lean_Meta_hasTypeMsg(lean_object*, lean_object*); static lean_object* l_Lean_Meta_mkLe___closed__3; lean_object* lean_panic_fn(lean_object*, lean_object*); @@ -426,10 +435,11 @@ LEAN_EXPORT lean_object* l_Lean_Meta_mkEqOfHEq(lean_object*, lean_object*, lean_ static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_AppBuilder___hyg_7226____closed__12; lean_object* lean_array_uget(lean_object*, size_t); static lean_object* l___private_Lean_Meta_AppBuilder_0__Lean_Meta_mkAppOptMAux___closed__6; +extern lean_object* l_Lean_trace_profiler; +lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_mkCongrFun___closed__4; lean_object* l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_AppBuilder_0__Lean_Meta_mkAppMFinal___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withOptProfile___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_mkEqMP___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___at_Lean_Meta_mkAppOptM_x27___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___rarg___lambda__1___closed__9; @@ -491,13 +501,11 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_mkProjection___ static lean_object* l_Lean_Meta_mkLT___closed__1; static lean_object* l_Lean_Meta_mkSub___closed__3; static lean_object* l_Lean_Meta_mkEqMP___closed__3; -static lean_object* l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___closed__2; static lean_object* l_Lean_Meta_mkEqOfHEq___closed__2; static lean_object* l___private_Lean_Meta_AppBuilder_0__Lean_Meta_throwAppBuilderException___rarg___closed__2; static lean_object* l___private_Lean_Meta_AppBuilder_0__Lean_Meta_mkAppMArgs_loop___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_mkArrayLit(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mkProjection___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* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_processPostponed___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_AppBuilder___hyg_7226____closed__1; static lean_object* l_Lean_Meta_mkEqOfHEq___lambda__1___closed__2; static lean_object* l_Lean_Meta_mkEqFalse___closed__1; @@ -6757,7 +6765,7 @@ return x_48; { lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; x_11 = l_Lean_mkAppN(x_2, x_3); -x_12 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_11, x_5, x_6, x_7, x_8, x_10); +x_12 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_11, x_5, x_6, x_7, x_8, x_10); x_13 = lean_ctor_get(x_12, 0); lean_inc(x_13); x_14 = lean_ctor_get(x_12, 1); @@ -7647,7 +7655,7 @@ lean_dec(x_2); return x_7; } } -static double _init_l_Lean_withSeconds___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__3___closed__1() { +static double _init_l_Lean_withSeconds___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__2___closed__1() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; double x_4; @@ -7658,7 +7666,7 @@ x_4 = l_Float_ofScientific(x_1, x_2, x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_withSeconds___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___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_EXPORT lean_object* l_Lean_withSeconds___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___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; lean_object* x_8; lean_object* x_9; lean_object* x_10; @@ -7690,7 +7698,7 @@ x_17 = 0; x_18 = lean_unsigned_to_nat(0u); x_19 = l_Float_ofScientific(x_16, x_17, x_18); lean_dec(x_16); -x_20 = l_Lean_withSeconds___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__3___closed__1; +x_20 = l_Lean_withSeconds___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__2___closed__1; x_21 = lean_float_div(x_19, x_20); x_22 = lean_box_float(x_21); x_23 = lean_alloc_ctor(0, 2, 0); @@ -7714,7 +7722,7 @@ x_27 = 0; x_28 = lean_unsigned_to_nat(0u); x_29 = l_Float_ofScientific(x_26, x_27, x_28); lean_dec(x_26); -x_30 = l_Lean_withSeconds___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__3___closed__1; +x_30 = l_Lean_withSeconds___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__2___closed__1; x_31 = lean_float_div(x_29, x_30); x_32 = lean_box_float(x_31); x_33 = lean_alloc_ctor(0, 2, 0); @@ -7751,201 +7759,7 @@ return x_38; } } } -static lean_object* _init_l_Lean_withOptProfile___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__2___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("profiler", 8); -return x_1; -} -} -static lean_object* _init_l_Lean_withOptProfile___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___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_Lean_withOptProfile___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__2___closed__1; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Lean_withOptProfile___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___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; uint8_t x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_7 = l_Lean_withOptProfile___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__2___closed__2; -x_8 = 0; -x_9 = l_Lean_getBoolOption___at_Lean_Meta_processPostponed___spec__4(x_7, x_8, x_2, x_3, x_4, x_5, x_6); -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -x_11 = lean_unbox(x_10); -lean_dec(x_10); -if (x_11 == 0) -{ -lean_object* x_12; lean_object* x_13; -x_12 = lean_ctor_get(x_9, 1); -lean_inc(x_12); -lean_dec(x_9); -x_13 = lean_apply_5(x_1, x_2, x_3, x_4, x_5, x_12); -if (lean_obj_tag(x_13) == 0) -{ -uint8_t x_14; -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_13, 0); -x_16 = lean_box(0); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_15); -lean_ctor_set(x_17, 1, x_16); -lean_ctor_set(x_13, 0, x_17); -return x_13; -} -else -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_18 = lean_ctor_get(x_13, 0); -x_19 = lean_ctor_get(x_13, 1); -lean_inc(x_19); -lean_inc(x_18); -lean_dec(x_13); -x_20 = lean_box(0); -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(0, 2, 0); -lean_ctor_set(x_22, 0, x_21); -lean_ctor_set(x_22, 1, x_19); -return x_22; -} -} -else -{ -uint8_t x_23; -x_23 = !lean_is_exclusive(x_13); -if (x_23 == 0) -{ -return x_13; -} -else -{ -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_13, 0); -x_25 = lean_ctor_get(x_13, 1); -lean_inc(x_25); -lean_inc(x_24); -lean_dec(x_13); -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_27; lean_object* x_28; -x_27 = lean_ctor_get(x_9, 1); -lean_inc(x_27); -lean_dec(x_9); -x_28 = l_Lean_withSeconds___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__3(x_1, x_2, x_3, x_4, x_5, x_27); -if (lean_obj_tag(x_28) == 0) -{ -uint8_t x_29; -x_29 = !lean_is_exclusive(x_28); -if (x_29 == 0) -{ -lean_object* x_30; uint8_t x_31; -x_30 = lean_ctor_get(x_28, 0); -x_31 = !lean_is_exclusive(x_30); -if (x_31 == 0) -{ -lean_object* x_32; lean_object* x_33; -x_32 = lean_ctor_get(x_30, 1); -x_33 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_33, 0, x_32); -lean_ctor_set(x_30, 1, x_33); -return x_28; -} -else -{ -lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_34 = lean_ctor_get(x_30, 0); -x_35 = lean_ctor_get(x_30, 1); -lean_inc(x_35); -lean_inc(x_34); -lean_dec(x_30); -x_36 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_36, 0, x_35); -x_37 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_37, 0, x_34); -lean_ctor_set(x_37, 1, x_36); -lean_ctor_set(x_28, 0, x_37); -return x_28; -} -} -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; -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_ctor_get(x_38, 0); -lean_inc(x_40); -x_41 = lean_ctor_get(x_38, 1); -lean_inc(x_41); -if (lean_is_exclusive(x_38)) { - lean_ctor_release(x_38, 0); - lean_ctor_release(x_38, 1); - x_42 = x_38; -} else { - lean_dec_ref(x_38); - x_42 = lean_box(0); -} -x_43 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_43, 0, x_41); -if (lean_is_scalar(x_42)) { - x_44 = lean_alloc_ctor(0, 2, 0); -} else { - x_44 = x_42; -} -lean_ctor_set(x_44, 0, x_40); -lean_ctor_set(x_44, 1, 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_39); -return x_45; -} -} -else -{ -uint8_t x_46; -x_46 = !lean_is_exclusive(x_28); -if (x_46 == 0) -{ -return x_28; -} -else -{ -lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_47 = lean_ctor_get(x_28, 0); -x_48 = lean_ctor_get(x_28, 1); -lean_inc(x_48); -lean_inc(x_47); -lean_dec(x_28); -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_EXPORT lean_object* l_MonadExcept_ofExcept___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___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_EXPORT lean_object* l_MonadExcept_ofExcept___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___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: { if (lean_obj_tag(x_1) == 0) @@ -8036,237 +7850,565 @@ return x_21; } } } -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___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) { -_start: -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; -lean_inc(x_10); -x_13 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_processPostponed___spec__6(x_1, x_2, x_3, x_6, x_4, x_8, x_9, x_10, x_11, x_12); -x_14 = lean_ctor_get(x_13, 1); -lean_inc(x_14); -lean_dec(x_13); -x_15 = l_MonadExcept_ofExcept___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__4(x_5, x_8, x_9, x_10, x_11, x_14); -lean_dec(x_10); -return x_15; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___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) { +_start: +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; +lean_inc(x_10); +x_13 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_processPostponed___spec__4(x_1, x_2, x_3, x_6, x_4, x_8, x_9, x_10, x_11, x_12); +x_14 = lean_ctor_get(x_13, 1); +lean_inc(x_14); +lean_dec(x_13); +x_15 = l_MonadExcept_ofExcept___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__3(x_5, x_8, x_9, x_10, x_11, x_14); +lean_dec(x_10); +return x_15; +} +} +static lean_object* _init_l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_profiler; +return x_1; +} +} +static lean_object* _init_l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__3___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("[", 1); +return x_1; +} +} +static lean_object* _init_l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__3___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__3___closed__2; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__3___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("s] ", 3); +return x_1; +} +} +static lean_object* _init_l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__3___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__3___closed__4; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, uint8_t x_7, double 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; lean_object* x_16; +lean_dec(x_9); +x_15 = lean_ctor_get(x_12, 5); +lean_inc(x_15); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_2); +x_16 = lean_apply_6(x_1, x_2, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_16) == 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); +x_19 = l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__3___closed__1; +x_20 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_6, x_19); +lean_dec(x_6); +if (x_20 == 0) +{ +if (x_7 == 0) +{ +lean_object* x_21; lean_object* x_22; +x_21 = lean_box(0); +x_22 = l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__2(x_3, x_4, x_15, x_5, x_2, x_17, x_21, x_10, x_11, x_12, x_13, x_18); +lean_dec(x_13); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_2); +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; +x_23 = lean_float_to_string(x_8); +x_24 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_24, 0, x_23); +x_25 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_25, 0, x_24); +x_26 = l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__3___closed__3; +x_27 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_25); +x_28 = l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__3___closed__5; +x_29 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +x_30 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_17); +x_31 = l___private_Lean_Meta_AppBuilder_0__Lean_Meta_hasTypeMsg___closed__2; +x_32 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +x_33 = lean_box(0); +x_34 = l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__2(x_3, x_4, x_15, x_5, x_2, x_32, x_33, x_10, x_11, x_12, x_13, x_18); +lean_dec(x_13); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_2); +return x_34; +} +} +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; +x_35 = lean_float_to_string(x_8); +x_36 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_36, 0, x_35); +x_37 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_37, 0, x_36); +x_38 = l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__3___closed__3; +x_39 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_37); +x_40 = l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__3___closed__5; +x_41 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_40); +x_42 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_17); +x_43 = l___private_Lean_Meta_AppBuilder_0__Lean_Meta_hasTypeMsg___closed__2; +x_44 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_44, 0, x_42); +lean_ctor_set(x_44, 1, x_43); +x_45 = lean_box(0); +x_46 = l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__2(x_3, x_4, x_15, x_5, x_2, x_44, x_45, x_10, x_11, x_12, x_13, x_18); +lean_dec(x_13); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_2); +return x_46; +} +} +else +{ +uint8_t x_47; +lean_dec(x_15); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_47 = !lean_is_exclusive(x_16); +if (x_47 == 0) +{ +return x_16; +} +else +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_16, 0); +x_49 = lean_ctor_get(x_16, 1); +lean_inc(x_49); +lean_inc(x_48); +lean_dec(x_16); +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; +} +} +} +} +static lean_object* _init_l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__4___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_trace_profiler; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t 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, 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_dec(x_7); +x_13 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_processPostponed___spec__2___rarg(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_alloc_closure((void*)(l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__1), 6, 1); +lean_closure_set(x_16, 0, x_1); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_17 = l_Lean_withSeconds___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__2(x_16, x_8, x_9, x_10, x_11, x_15); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; lean_object* x_67; uint8_t x_68; +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_ctor_get(x_18, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_18, 1); +lean_inc(x_21); +lean_dec(x_18); +x_67 = l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__4___closed__1; +x_68 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_5, x_67); +if (x_68 == 0) +{ +uint8_t x_69; +x_69 = 0; +x_22 = x_69; +goto block_66; +} +else +{ +double x_70; double x_71; uint8_t x_72; +x_70 = l_Lean_trace_profiler_threshold_getSecs(x_5); +x_71 = lean_unbox_float(x_21); +x_72 = lean_float_decLt(x_70, x_71); +x_22 = x_72; +goto block_66; +} +block_66: +{ +if (x_6 == 0) +{ +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; +lean_dec(x_21); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +x_23 = lean_st_ref_take(x_11, 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_is_exclusive(x_24); +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_24, 3); +x_28 = l_Lean_PersistentArray_append___rarg(x_14, x_27); +lean_ctor_set(x_24, 3, x_28); +x_29 = lean_st_ref_set(x_11, x_24, x_25); +x_30 = lean_ctor_get(x_29, 1); +lean_inc(x_30); +lean_dec(x_29); +x_31 = l_MonadExcept_ofExcept___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__3(x_20, x_8, x_9, x_10, x_11, x_30); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_20); +if (lean_obj_tag(x_31) == 0) +{ +uint8_t x_32; +x_32 = !lean_is_exclusive(x_31); +if (x_32 == 0) +{ +return x_31; +} +else +{ +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(0, 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_36; +x_36 = !lean_is_exclusive(x_31); +if (x_36 == 0) +{ +return x_31; +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_31, 0); +x_38 = lean_ctor_get(x_31, 1); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_31); +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_40; lean_object* x_41; 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_40 = lean_ctor_get(x_24, 0); +x_41 = lean_ctor_get(x_24, 1); +x_42 = lean_ctor_get(x_24, 2); +x_43 = lean_ctor_get(x_24, 3); +x_44 = lean_ctor_get(x_24, 4); +x_45 = lean_ctor_get(x_24, 5); +x_46 = lean_ctor_get(x_24, 6); +lean_inc(x_46); +lean_inc(x_45); +lean_inc(x_44); +lean_inc(x_43); +lean_inc(x_42); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_24); +x_47 = l_Lean_PersistentArray_append___rarg(x_14, x_43); +x_48 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_48, 0, x_40); +lean_ctor_set(x_48, 1, x_41); +lean_ctor_set(x_48, 2, x_42); +lean_ctor_set(x_48, 3, x_47); +lean_ctor_set(x_48, 4, x_44); +lean_ctor_set(x_48, 5, x_45); +lean_ctor_set(x_48, 6, x_46); +x_49 = lean_st_ref_set(x_11, x_48, x_25); +x_50 = lean_ctor_get(x_49, 1); +lean_inc(x_50); +lean_dec(x_49); +x_51 = l_MonadExcept_ofExcept___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__3(x_20, x_8, x_9, x_10, x_11, x_50); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_20); +if (lean_obj_tag(x_51) == 0) +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_52 = lean_ctor_get(x_51, 0); +lean_inc(x_52); +x_53 = lean_ctor_get(x_51, 1); +lean_inc(x_53); +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_51); + x_54 = lean_box(0); +} +if (lean_is_scalar(x_54)) { + x_55 = lean_alloc_ctor(0, 2, 0); +} else { + x_55 = x_54; +} +lean_ctor_set(x_55, 0, x_52); +lean_ctor_set(x_55, 1, x_53); +return x_55; +} +else +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_56 = lean_ctor_get(x_51, 0); +lean_inc(x_56); +x_57 = lean_ctor_get(x_51, 1); +lean_inc(x_57); +if (lean_is_exclusive(x_51)) { + lean_ctor_release(x_51, 0); + lean_ctor_release(x_51, 1); + x_58 = x_51; +} else { + lean_dec_ref(x_51); + x_58 = lean_box(0); +} +if (lean_is_scalar(x_58)) { + x_59 = lean_alloc_ctor(1, 2, 0); +} else { + x_59 = x_58; +} +lean_ctor_set(x_59, 0, x_56); +lean_ctor_set(x_59, 1, x_57); +return x_59; } } -static lean_object* _init_l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___closed__1() { -_start: +} +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("[", 1); -return x_1; +lean_object* x_60; double x_61; lean_object* x_62; +x_60 = lean_box(0); +x_61 = lean_unbox_float(x_21); +lean_dec(x_21); +x_62 = l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__3(x_2, x_20, x_14, x_3, x_4, x_5, x_22, x_61, x_60, x_8, x_9, x_10, x_11, x_19); +return x_62; } } -static lean_object* _init_l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___closed__2() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___closed__1; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +lean_object* x_63; double x_64; lean_object* x_65; +x_63 = lean_box(0); +x_64 = lean_unbox_float(x_21); +lean_dec(x_21); +x_65 = l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__3(x_2, x_20, x_14, x_3, x_4, x_5, x_22, x_64, x_63, x_8, x_9, x_10, x_11, x_19); +return x_65; } } -static lean_object* _init_l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("s] ", 3); -return x_1; } +else +{ +uint8_t x_73; +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +x_73 = !lean_is_exclusive(x_17); +if (x_73 == 0) +{ +return x_17; } -static lean_object* _init_l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___closed__4() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___closed__3; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_74 = lean_ctor_get(x_17, 0); +x_75 = lean_ctor_get(x_17, 1); +lean_inc(x_75); +lean_inc(x_74); +lean_dec(x_17); +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_EXPORT lean_object* l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___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) { _start: { -lean_object* x_10; lean_object* x_11; uint8_t x_12; +lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_10 = lean_ctor_get(x_7, 2); +lean_inc(x_10); lean_inc(x_1); -x_10 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_1, 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_unbox(x_11); +x_11 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___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_unbox(x_12); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); lean_dec(x_11); -if (x_12 == 0) +x_15 = l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__4___closed__1; +x_16 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_10, x_15); +if (x_16 == 0) { -lean_object* x_13; lean_object* x_14; +lean_object* x_17; +lean_dec(x_12); +lean_dec(x_10); lean_dec(x_2); lean_dec(x_1); -x_13 = lean_ctor_get(x_10, 1); -lean_inc(x_13); -lean_dec(x_10); -x_14 = lean_apply_5(x_3, x_5, x_6, x_7, x_8, x_13); -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; -x_15 = lean_ctor_get(x_10, 1); -lean_inc(x_15); -lean_dec(x_10); -x_16 = lean_ctor_get(x_7, 5); -lean_inc(x_16); -x_17 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_processPostponed___spec__2___rarg(x_8, 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_alloc_closure((void*)(l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__1), 6, 1); -lean_closure_set(x_20, 0, x_3); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_21 = l_Lean_withOptProfile___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__2(x_20, x_5, x_6, x_7, x_8, 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; -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_22, 0); -lean_inc(x_24); -x_25 = lean_ctor_get(x_22, 1); -lean_inc(x_25); -lean_dec(x_22); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_24); -x_26 = lean_apply_6(x_2, x_24, x_5, x_6, x_7, x_8, x_23); -if (lean_obj_tag(x_26) == 0) +x_17 = lean_apply_5(x_3, x_5, x_6, x_7, x_8, x_14); +if (lean_obj_tag(x_17) == 0) { -if (lean_obj_tag(x_25) == 0) +uint8_t x_18; +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; -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_box(0); -x_30 = l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__2(x_18, x_1, x_16, x_4, x_24, x_27, x_29, x_5, x_6, x_7, x_8, x_28); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_24); -return x_30; +return x_17; } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; double x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; 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_31 = lean_ctor_get(x_26, 0); -lean_inc(x_31); -x_32 = lean_ctor_get(x_26, 1); -lean_inc(x_32); -lean_dec(x_26); -x_33 = lean_ctor_get(x_25, 0); -lean_inc(x_33); -lean_dec(x_25); -x_34 = lean_unbox_float(x_33); -lean_dec(x_33); -x_35 = lean_float_to_string(x_34); -x_36 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_36, 0, x_35); -x_37 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_37, 0, x_36); -x_38 = l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___closed__2; -x_39 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_39, 1, x_37); -x_40 = l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___closed__4; -x_41 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_41, 0, x_39); -lean_ctor_set(x_41, 1, x_40); -x_42 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_31); -x_43 = l___private_Lean_Meta_AppBuilder_0__Lean_Meta_hasTypeMsg___closed__2; -x_44 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_44, 0, x_42); -lean_ctor_set(x_44, 1, x_43); -x_45 = lean_box(0); -x_46 = l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__2(x_18, x_1, x_16, x_4, x_24, x_44, x_45, x_5, x_6, x_7, x_8, x_32); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_24); -return x_46; +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_17, 0); +x_20 = lean_ctor_get(x_17, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_17); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); +return x_21; } } else { -uint8_t x_47; -lean_dec(x_25); -lean_dec(x_24); -lean_dec(x_18); -lean_dec(x_16); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_1); -x_47 = !lean_is_exclusive(x_26); -if (x_47 == 0) +uint8_t x_22; +x_22 = !lean_is_exclusive(x_17); +if (x_22 == 0) { -return x_26; +return x_17; } else { -lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_48 = lean_ctor_get(x_26, 0); -x_49 = lean_ctor_get(x_26, 1); -lean_inc(x_49); -lean_inc(x_48); -lean_dec(x_26); -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; +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_17, 0); +x_24 = lean_ctor_get(x_17, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_17); +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 { -uint8_t x_51; -lean_dec(x_18); -lean_dec(x_16); -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_51 = !lean_is_exclusive(x_21); -if (x_51 == 0) -{ -return x_21; +lean_object* x_26; uint8_t x_27; lean_object* x_28; +x_26 = lean_box(0); +x_27 = lean_unbox(x_12); +lean_dec(x_12); +x_28 = l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__4(x_3, x_2, x_1, x_4, x_10, x_27, x_26, x_5, x_6, x_7, x_8, x_14); +return x_28; +} } else { -lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_52 = lean_ctor_get(x_21, 0); -x_53 = lean_ctor_get(x_21, 1); -lean_inc(x_53); -lean_inc(x_52); -lean_dec(x_21); -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_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; +x_29 = lean_ctor_get(x_11, 1); +lean_inc(x_29); +lean_dec(x_11); +x_30 = lean_box(0); +x_31 = lean_unbox(x_12); +lean_dec(x_12); +x_32 = l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__4(x_3, x_2, x_1, x_4, x_10, x_31, x_30, x_5, x_6, x_7, x_8, x_29); +return x_32; } } } @@ -8697,11 +8839,11 @@ x_3 = lean_alloc_closure((void*)(l___private_Lean_Meta_AppBuilder_0__Lean_Meta_w return x_3; } } -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___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_MonadExcept_ofExcept___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___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_MonadExcept_ofExcept___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__4(x_1, x_2, x_3, x_4, x_5, x_6); +x_7 = l_MonadExcept_ofExcept___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__3(x_1, x_2, x_3, x_4, x_5, x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); @@ -8725,6 +8867,32 @@ lean_dec(x_5); return x_14; } } +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___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) { +_start: +{ +uint8_t x_15; uint8_t x_16; double x_17; lean_object* x_18; +x_15 = lean_unbox(x_5); +lean_dec(x_5); +x_16 = lean_unbox(x_7); +lean_dec(x_7); +x_17 = lean_unbox_float(x_8); +lean_dec(x_8); +x_18 = l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__3(x_1, x_2, x_3, x_4, x_15, x_6, x_16, x_17, x_9, x_10, x_11, x_12, x_13, x_14); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___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) { +_start: +{ +uint8_t x_13; uint8_t x_14; lean_object* x_15; +x_13 = lean_unbox(x_4); +lean_dec(x_4); +x_14 = lean_unbox(x_6); +lean_dec(x_6); +x_15 = l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__4(x_1, x_2, x_3, x_13, x_5, x_14, x_7, x_8, x_9, x_10, x_11, x_12); +return x_15; +} +} LEAN_EXPORT lean_object* l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___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) { _start: { @@ -11657,7 +11825,7 @@ static lean_object* _init_l_Lean_Meta_mkProjection___lambda__1___closed__2() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("mkProjectionn", 13); +x_1 = lean_mk_string_from_bytes("mkProjection", 12); return x_1; } } @@ -11924,24 +12092,6 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Meta_mkProjection___closed__4() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("mkProjection", 12); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_mkProjection___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_mkProjection___closed__4; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} LEAN_EXPORT lean_object* l_Lean_Meta_mkProjection(lean_object* x_1, 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: { @@ -12006,7 +12156,7 @@ x_23 = l_Lean_Meta_mkProjection___closed__3; x_24 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_24, 0, x_23); lean_ctor_set(x_24, 1, x_22); -x_25 = l_Lean_Meta_mkProjection___closed__5; +x_25 = l_Lean_Meta_mkProjection___lambda__1___closed__3; x_26 = l___private_Lean_Meta_AppBuilder_0__Lean_Meta_throwAppBuilderException___rarg(x_25, x_24, x_3, x_4, x_5, x_6, x_19); x_27 = !lean_is_exclusive(x_26); if (x_27 == 0) @@ -14772,19 +14922,19 @@ l___private_Lean_Meta_AppBuilder_0__Lean_Meta_mkAppMArgs_loop___closed__9 = _ini lean_mark_persistent(l___private_Lean_Meta_AppBuilder_0__Lean_Meta_mkAppMArgs_loop___closed__9); l___private_Lean_Meta_AppBuilder_0__Lean_Meta_mkAppMArgs___closed__1 = _init_l___private_Lean_Meta_AppBuilder_0__Lean_Meta_mkAppMArgs___closed__1(); lean_mark_persistent(l___private_Lean_Meta_AppBuilder_0__Lean_Meta_mkAppMArgs___closed__1); -l_Lean_withSeconds___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__3___closed__1 = _init_l_Lean_withSeconds___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__3___closed__1(); -l_Lean_withOptProfile___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__2___closed__1 = _init_l_Lean_withOptProfile___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__2___closed__1(); -lean_mark_persistent(l_Lean_withOptProfile___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__2___closed__1); -l_Lean_withOptProfile___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__2___closed__2 = _init_l_Lean_withOptProfile___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__2___closed__2(); -lean_mark_persistent(l_Lean_withOptProfile___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__2___closed__2); -l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___closed__1 = _init_l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___closed__1(); -lean_mark_persistent(l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___closed__1); -l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___closed__2 = _init_l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___closed__2(); -lean_mark_persistent(l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___closed__2); -l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___closed__3 = _init_l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___closed__3(); -lean_mark_persistent(l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___closed__3); -l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___closed__4 = _init_l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___closed__4(); -lean_mark_persistent(l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___closed__4); +l_Lean_withSeconds___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__2___closed__1 = _init_l_Lean_withSeconds___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__2___closed__1(); +l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__3___closed__1 = _init_l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__3___closed__1); +l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__3___closed__2 = _init_l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__3___closed__2); +l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__3___closed__3 = _init_l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__3___closed__3(); +lean_mark_persistent(l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__3___closed__3); +l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__3___closed__4 = _init_l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__3___closed__4(); +lean_mark_persistent(l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__3___closed__4); +l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__3___closed__5 = _init_l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__3___closed__5(); +lean_mark_persistent(l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__3___closed__5); +l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__4___closed__1 = _init_l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_withTraceNode___at___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___spec__1___lambda__4___closed__1); l___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___rarg___lambda__1___closed__1 = _init_l___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___rarg___lambda__1___closed__1(); lean_mark_persistent(l___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___rarg___lambda__1___closed__1); l___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___rarg___lambda__1___closed__2 = _init_l___private_Lean_Meta_AppBuilder_0__Lean_Meta_withAppBuilderTrace___rarg___lambda__1___closed__2(); @@ -14921,10 +15071,6 @@ l_Lean_Meta_mkProjection___closed__2 = _init_l_Lean_Meta_mkProjection___closed__ lean_mark_persistent(l_Lean_Meta_mkProjection___closed__2); l_Lean_Meta_mkProjection___closed__3 = _init_l_Lean_Meta_mkProjection___closed__3(); lean_mark_persistent(l_Lean_Meta_mkProjection___closed__3); -l_Lean_Meta_mkProjection___closed__4 = _init_l_Lean_Meta_mkProjection___closed__4(); -lean_mark_persistent(l_Lean_Meta_mkProjection___closed__4); -l_Lean_Meta_mkProjection___closed__5 = _init_l_Lean_Meta_mkProjection___closed__5(); -lean_mark_persistent(l_Lean_Meta_mkProjection___closed__5); l_Lean_Meta_mkListLit___closed__1 = _init_l_Lean_Meta_mkListLit___closed__1(); lean_mark_persistent(l_Lean_Meta_mkListLit___closed__1); l_Lean_Meta_mkListLit___closed__2 = _init_l_Lean_Meta_mkListLit___closed__2(); diff --git a/stage0/stdlib/Lean/Meta/Basic.c b/stage0/stdlib/Lean/Meta/Basic.c index b103192c384b..a509f8e7ec63 100644 --- a/stage0/stdlib/Lean/Meta/Basic.c +++ b/stage0/stdlib/Lean/Meta/Basic.c @@ -50,6 +50,7 @@ lean_object* l_Lean_MetavarContext_findDecl_x3f(lean_object*, lean_object*); static lean_object* l_Lean_Meta_instAlternativeMetaM___closed__3; LEAN_EXPORT lean_object* l_Lean_MVarId_withContext(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_savingCache___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_profiler; lean_object* l_Lean_Core_getMaxHeartbeats(lean_object*); static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Meta_withIncRecDepth___spec__1___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_Meta_mkConstWithFreshMVarLevels___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -90,11 +91,11 @@ static lean_object* l_Lean_Meta_ParamInfo_backDeps___default___closed__1; static lean_object* l_Lean_Meta_processPostponed___lambda__1___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLetDeclImp(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls_loop(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_mkFreshLMVarId___at_Lean_Meta_mkFreshLevelMVar___spec__1___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_whnfI(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_2812____closed__16; LEAN_EXPORT lean_object* l_Lean_Meta_abstract(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_processPostponed___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT 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_Meta_instMetaEvalMetaM___rarg___closed__6; LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Meta_sortFVarIds___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); @@ -107,7 +108,6 @@ lean_object* l_Lean_ConstantInfo_type(lean_object*); LEAN_EXPORT lean_object* l_Lean_getExprMVarAssignment_x3f___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassQuick_x3f___spec__1(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_EXPORT uint8_t l_Lean_Meta_Config_assignSyntheticOpaque___default; -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_processPostponed___spec__6___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_maxRecDepthErrorMessage; LEAN_EXPORT lean_object* l_Lean_Meta_withConfig(lean_object*); lean_object* l_Lean_MessageData_toString(lean_object*, lean_object*); @@ -131,16 +131,16 @@ static lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageC LEAN_EXPORT lean_object* l_Lean_Meta_withAtLeastTransparency___rarg___lambda__1(uint8_t, 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_withLocalDecls_loop___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_Meta_forallMetaTelescope(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_getPostponed___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withNewMCtxDepth___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_Meta_processPostponed___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_getLocalDeclFromUserName___closed__2; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18865____closed__7; LEAN_EXPORT lean_object* l_Lean_Meta_approxDefEq(lean_object*); -LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_instantiateForallAux___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getZetaFVarIds(lean_object*); lean_object* l_Lean_Level_succ___override(lean_object*); +double l_Lean_trace_profiler_threshold_getSecs(lean_object*); LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Meta_withLocalInstancesImp___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_Meta_normalizeLevel___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_ParamInfo_isProp___default; @@ -170,6 +170,7 @@ LEAN_EXPORT lean_object* l_Lean_mkFreshId___at_Lean_Meta_mkFreshExprMVarAt___spe static lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_instantiateLambdaAux___closed__2; static lean_object* l_Lean_Meta_liftMkBindingM___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_FVarId_throwUnknown(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___boxed(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_Meta_instMonadMetaM___closed__5; LEAN_EXPORT lean_object* l_Lean_FVarId_isLetVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -187,7 +188,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_mkFreshExprMVa LEAN_EXPORT lean_object* l_Lean_Meta_getMVarDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_mkFunUnit___closed__3; 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*); -lean_object* l_Lean_Option_get___at_Std_Format_pretty_x27___spec__1(lean_object*, lean_object*); lean_object* l_Lean_replaceRef(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at___private_Lean_Meta_Basic_0__Lean_Meta_processPostponedStep___spec__2(uint8_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___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -209,11 +209,11 @@ static lean_object* l_Lean_Meta_instMonadEnvMetaM___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_withNewBinderInfos___at_Lean_Meta_withInstImplicitAsImplict___spec__3(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_resetZetaFVarIds___boxed(lean_object*); static lean_object* l_Lean_Meta_State_postponed___default___closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_processPostponed___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_isDefEqNoConstantApprox(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ConstantInfo_name(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withAssignableSyntheticOpaque(lean_object*); LEAN_EXPORT uint8_t l_Array_qsort_sort___at_Lean_Meta_sortFVarIds___spec__1___lambda__1(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___closed__2; lean_object* lean_synth_pending(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Meta_sortFVarIds___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_FVarId_isLetVar___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -227,6 +227,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_isLevelDefEq(lean_object*, lean_object*, le LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___rarg___lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_lambdaMetaTelescope(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_environment_find(lean_object*, lean_object*); +uint8_t lean_float_decLt(double, double); static lean_object* l_Lean_Meta_instInhabitedPostponedEntry___closed__1; static lean_object* l_Lean_Meta_Cache_defEq___default___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withNewLocalInstanceImp(lean_object*); @@ -251,7 +252,6 @@ static lean_object* l_Lean_Meta_instMetaEvalMetaM___rarg___closed__2; LEAN_EXPORT uint8_t l_Lean_Meta_Config_proofIrrelevance___default; LEAN_EXPORT uint8_t l_Lean_Meta_Config_ctxApprox___default; static lean_object* l_Lean_Meta_State_mctx___default___closed__4; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18710____closed__5; LEAN_EXPORT lean_object* l_Lean_getExprMVarAssignment_x3f___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassQuick_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_instMonadBacktrackSavedStateMetaM___closed__2; lean_object* l_Lean_Expr_setPPUniverses(lean_object*, uint8_t); @@ -260,6 +260,7 @@ uint8_t l_Lean_Name_quickLt(lean_object*, lean_object*); lean_object* lean_io_get_num_heartbeats(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_mkFreshExprMVarAtCore___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*); size_t lean_usize_of_nat(lean_object*); +static lean_object* l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__3___closed__2; lean_object* l_Lean_EnumAttributes_setValue___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope(lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_ParamInfo_dependsOnHigherOrderOutParam___default; @@ -310,12 +311,12 @@ lean_object* l_Lean_LocalDecl_index(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_processPostponed___spec__2(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_2812____closed__15; LEAN_EXPORT lean_object* l_Lean_Meta_State_postponed___default; -LEAN_EXPORT lean_object* l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_findLocalDecl_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_hasMVar(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_MVarId_getDecl___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_5____closed__1; static lean_object* l_Lean_Meta_mkLevelStuckErrorMessage___closed__1; +static lean_object* l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__4___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_getLocalInstances___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_instHashableArray___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mkFreshExprMVarWithId___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -355,10 +356,11 @@ LEAN_EXPORT lean_object* l_Lean_FVarId_getDecl___boxed(lean_object*, lean_object LEAN_EXPORT lean_object* l_Lean_Meta_withReducible___rarg___lambda__1(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_mkLeveErrorMessageCore___lambda__2(lean_object*, lean_object*); LEAN_EXPORT 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*); -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18710_(lean_object*); +static lean_object* l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__3___closed__5; LEAN_EXPORT lean_object* l_Lean_Meta_getFVarFromUserName___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_findCore___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mkFreshLevelMVar___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static double l_Lean_withSeconds___at_Lean_Meta_processPostponed___spec__3___closed__1; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Expr_abstractRangeM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_processPostponedStep(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_local_ctx_find(lean_object*, lean_object*); @@ -379,10 +381,10 @@ LEAN_EXPORT lean_object* l_Lean_Meta_FunInfo_resultDeps___default; LEAN_EXPORT lean_object* l_Lean_Meta_fullApproxDefEq___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_Config_constApprox___default; uint8_t lean_expr_eqv(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__3(lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_2812____closed__13; LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDeclImp___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_Expr_abstractRangeM___boxed(lean_object*, lean_object*, 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_12510_(uint8_t, uint8_t); static lean_object* l_Lean_Meta_Cache_funInfo___default___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_withNewBinderInfos___rarg___lambda__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_Meta_Basic_0__Lean_Meta_getConstTemp_x3f___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -390,6 +392,7 @@ uint8_t lean_get_reducibility_status(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_isClassQuickConst_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Cache_defEq___default___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Cache_whnfAll___default; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18865____closed__3; LEAN_EXPORT uint8_t l_Lean_Meta_Config_unificationHints___default; LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_mkFreshExprMVarWithIdCore___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_instMetaEvalMetaM(lean_object*); @@ -407,13 +410,13 @@ LEAN_EXPORT lean_object* l_Lean_Meta_map2MetaM___rarg(lean_object*, lean_object* LEAN_EXPORT lean_object* l_Lean_Meta_withNewLocalInstances___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Cache_synthInstance___default___closed__2; LEAN_EXPORT 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* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(lean_object*, lean_object*); lean_object* l_Lean_ConstantInfo_numLevelParams(lean_object*); static lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_mkLevelStuckErrorMessage(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_setPostponed___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_2812____closed__11; LEAN_EXPORT lean_object* l_Lean_Meta_orElse___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18710____closed__4; lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean_object*); lean_object* l_Lean_registerInternalExceptionId(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_withContext___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -440,8 +443,8 @@ LEAN_EXPORT lean_object* l_Lean_Meta_ParamInfo_isExplicit___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_isReadOnlyOrSyntheticOpaqueExprMVar___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_isDefEqGuarded(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_liftMkBindingM___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_processPostponed___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getTransparency___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_getBoolOption___at_Lean_Meta_processPostponed___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_instMetaEvalMetaM___rarg___closed__8; LEAN_EXPORT lean_object* l_Lean_Meta_withNewMCtxDepth___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mkFreshExprMVar___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -460,7 +463,7 @@ lean_object* l_Lean_Expr_forallE___override(lean_object*, lean_object*, lean_obj LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Meta_withIncRecDepth___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_setMVarUserName___boxed(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_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___lambda__1(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___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getZetaFVarIds___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_liftMetaM(lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Meta_Basic_0__Lean_Meta_processPostponedStep___spec__5___closed__2; @@ -487,7 +490,6 @@ lean_object* lean_st_mk_ref(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_elimMVarDeps___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_setMVarKind(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withNewBinderInfos___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_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__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_withMVarContextImp___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_getConstInfo___at_Lean_Meta_mkConstWithFreshMVarLevels___spec__1___closed__4; static lean_object* l_Lean_FVarId_throwUnknown___rarg___closed__1; @@ -510,6 +512,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls_loop___rarg___lambda__1(lean lean_object* l_Lean_printTraces___at_Lean_Core_instMetaEvalCoreM___spec__1(lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Expr_instHashableExpr; LEAN_EXPORT lean_object* l_Lean_Meta_inferType___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_levelZero; LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withNewMCtxDepthImp___rarg(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_lambdaMetaTelescope_process(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -528,11 +531,11 @@ static lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageC static lean_object* l_Lean_Meta_instInhabitedCache___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withExistingLocalDeclsImp(lean_object*); static lean_object* l_Lean_Meta_liftMkBindingM___rarg___closed__3; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18710____closed__3; uint64_t l_Lean_Expr_hash(lean_object*); uint8_t l_Lean_BinderInfo_isInstImplicit(uint8_t); 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_FVarId_getType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18865____closed__4; LEAN_EXPORT lean_object* l_Lean_LMVarId_isReadOnly(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLCtx(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_MVarId_getDecl___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -541,14 +544,12 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_liftMkBindingM___spec_ LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_lambdaTelescopeImp(lean_object*); LEAN_EXPORT lean_object* l_Lean_FVarId_getValue_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_shouldReduceAll(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_getBoolOption___at_Lean_Meta_processPostponed___spec__4(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mapError___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_instMonadBacktrackSavedStateMetaM___closed__3; lean_object* l_Lean_instHashableLocalInstance___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_instMonadEnvMetaM___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_letE___override(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); lean_object* l_instHashableProd___rarg___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_mkLetFVars(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_getParamNames___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_Meta_mkFunUnit(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -566,8 +567,10 @@ LEAN_EXPORT lean_object* l_Lean_Meta_mapMetaM(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Meta_Basic_0__Lean_Meta_processPostponedStep___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_Lean_MetavarContext_MkBinding_collectForwardDeps(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_approxDefEq___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_mkConstWithFreshMVarLevels___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Meta_Basic_0__Lean_Meta_processPostponedStep___spec__4(uint8_t, lean_object*, lean_object*, 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_initFn____x40_Lean_Meta_Basic___hyg_18865____closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_withoutProofIrrelevance___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withoutProofIrrelevance___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getParamNames___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -601,9 +604,8 @@ lean_object* l_Lean_MetavarContext_MkBinding_elimMVarDeps(lean_object*, lean_obj static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Meta_Basic_0__Lean_Meta_processPostponedStep___spec__5___closed__3; static lean_object* l_Lean_Meta_processPostponed_loop___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_ppExpr___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static double l_Lean_withSeconds___at_Lean_Meta_processPostponed___spec__5___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclsD(lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18710____closed__2; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18865____closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_isListLevelDefEqAux(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_isReadOnlyOrSyntheticOpaque(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at___private_Lean_Meta_Basic_0__Lean_Meta_processPostponedStep___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -617,6 +619,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_instMonadMetaM; LEAN_EXPORT lean_object* l_Lean_Meta_liftMkBindingM(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mapError(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLCtx___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_isClassExpensive_x3f___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_instMonadBacktrackSavedStateMetaM; static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_2812____closed__7; LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withNewFVar___boxed(lean_object*, lean_object*); @@ -624,15 +627,16 @@ LEAN_EXPORT lean_object* l_Lean_Meta_approxDefEq___rarg(lean_object*, lean_objec LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_withInstImplicitAsImplict___spec__2(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getParamNames(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_mod(size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___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_throwError___at_Lean_LMVarId_getLevel___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_instInhabitedMetaM___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getZetaFVarIds___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getMVarDecl___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_instInhabitedMetaM___rarg___closed__2; lean_object* l_Lean_LocalContext_mkLocalDecl(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t); -static lean_object* l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18710____closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_withLCtx___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_Lean_FVarId_throwUnknown___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_processPostponed___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_Config_offsetCnstrs___default; LEAN_EXPORT lean_object* l_Lean_Meta_isSyntheticMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_LMVarId_getLevel___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -644,6 +648,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_getConfig(lean_object*, lean_object*, lean_ LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_forallTelescopeReducingAuxAux___rarg(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Level_collectMVars(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_filterMapM___at_Lean_Meta_withInstImplicitAsImplict___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_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_LocalDecl_isImplementationDetail(lean_object*); static lean_object* l_Lean_Meta_Cache_funInfo___default___closed__1; static lean_object* l_Lean_Meta_instMetaEvalMetaM___rarg___closed__10; @@ -668,21 +673,20 @@ LEAN_EXPORT lean_object* l_Lean_Meta_mapMetaM___rarg(lean_object*, lean_object*, LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_modifyInferTypeCache___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallMetaTelescope___boxed(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_addTraceNodeCore___at_Lean_Meta_processPostponed___spec__7(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___private_Lean_Meta_Basic_0__Lean_Meta_mkFreshExprMVarAtCore(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_Lean_Meta_withLocalDecl(lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18710____closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Meta_withLocalDecls_loop___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withReducibleAndInstances___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_processPostponed___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_processPostponedStep___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_instInhabitedCache; lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_getLevelMVarAssignment_x3f___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_2812____closed__2; -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_processPostponed___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_throwIsDefEqStuck___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkFreshId___at_Lean_Meta_mkFreshExprMVarAt___spec__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_instantiateLambdaAux(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_instMonadMetaM___closed__2; +LEAN_EXPORT lean_object* l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2___rarg(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_lambdaTelescopeImp_process___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_Meta_Cache_inferType___default___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -695,16 +699,17 @@ LEAN_EXPORT lean_object* l_Lean_Meta_shouldReduceReducibleOnly___boxed(lean_obje LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_forallTelescopeReducingAuxAux_process___rarg(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_Meta_throwUnknownFVar___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_Config_etaStruct___default; +static lean_object* l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__3___closed__1; 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_instMonadEnvMetaM; LEAN_EXPORT lean_object* l_Lean_Meta_withMCtx(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getPostponed___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_FunInfo_paramInfo___default; +LEAN_EXPORT lean_object* l_Lean_withSeconds___at_Lean_Meta_processPostponed___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_instAlternativeMetaM___lambda__1___closed__1; LEAN_EXPORT uint8_t l_Lean_Meta_ParamInfo_isStrictImplicit(lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Meta_Basic_0__Lean_Meta_fvarsSizeLtMaxFVars(lean_object*, lean_object*); lean_object* l_Lean_Expr_ReplaceImpl_Cache_new(lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18710____closed__7; LEAN_EXPORT uint8_t l_Lean_Meta_ParamInfo_isDecInst___default; lean_object* l_Lean_Expr_app___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkFreshFVarId___at___private_Lean_Meta_Basic_0__Lean_Meta_forallTelescopeReducingAuxAux_process___spec__1(lean_object*, lean_object*, lean_object*); @@ -734,11 +739,11 @@ LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Meta_sortFVarIds___spec__1 LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_lambdaTelescopeImp___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_EXPORT lean_object* l_Lean_Meta_mkFreshExprMVar(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withSeconds___at_Lean_Meta_processPostponed___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Meta_withIncRecDepth___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_setInlineAttribute___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_instOrElseMetaM(lean_object*); lean_object* lean_float_to_string(double); +LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofLevel(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Context_localInstances___default; LEAN_EXPORT lean_object* l_Lean_Expr_abstractRangeM(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -764,12 +769,12 @@ LEAN_EXPORT lean_object* l_Lean_Meta_liftMetaM___rarg(lean_object*, lean_object* LEAN_EXPORT lean_object* l_Lean_Meta_map1MetaM___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_instMonadEnvMetaM___lambda__1(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_1116____at___private_Lean_Meta_Basic_0__Lean_Meta_beqInfoCacheKey____x40_Lean_Meta_Basic___hyg_368____spec__1___boxed(lean_object*, lean_object*); -uint8_t l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_12336_(uint8_t, uint8_t); LEAN_EXPORT lean_object* l_Lean_MVarId_setType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_instAlternativeMetaM___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withExistingLocalDecls___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_processPostponed_loop___closed__5; LEAN_EXPORT lean_object* l_Lean_FVarId_getUserName(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__3___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_collectForwardDeps___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_setMVarType___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT 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*); @@ -779,7 +784,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_MetaM_run(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_savingCache___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_processPostponed___lambda__1___closed__2; static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Meta_withIncRecDepth___spec__1___rarg___closed__2; -uint8_t l_Lean_KVMap_getBool(lean_object*, lean_object*, uint8_t); static lean_object* l_Lean_Meta_instMonadMetaM___closed__1; static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_2812____closed__3; LEAN_EXPORT lean_object* l_Lean_mkFreshFVarId___at___private_Lean_Meta_Basic_0__Lean_Meta_forallTelescopeReducingAuxAux_process___spec__1___rarg(lean_object*, lean_object*); @@ -795,9 +799,9 @@ static lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageC LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclsD___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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_withLetDeclImp___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_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Meta_Basic_0__Lean_Meta_processPostponedStep___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_Lean_throwError___at___private_Lean_Meta_Basic_0__Lean_Meta_getConstTemp_x3f___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Meta_processPostponed___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_Lean_Meta_withNewLocalInstance(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_findLocalDecl_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Meta_Basic_0__Lean_Meta_beqInfoCacheKey____x40_Lean_Meta_Basic___hyg_368_(lean_object*, lean_object*); @@ -836,8 +840,10 @@ LEAN_EXPORT lean_object* l_Lean_Meta_withTransparency___rarg___boxed(lean_object lean_object* l_Lean_mkLevelIMax_x27(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withNewLocalInstancesImpAux___at_Lean_Meta_withNewLocalInstances___spec__1(lean_object*); static lean_object* l_Lean_Meta_instInhabitedParamInfo___closed__1; +static lean_object* l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__3___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_isReadOnlyExprMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_instAlternativeMetaM; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18865____closed__2; extern lean_object* l_Lean_firstFrontendMacroScope; LEAN_EXPORT lean_object* l_Lean_Meta_ParamInfo_isInstImplicit___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalInstancesImp(lean_object*); @@ -851,8 +857,10 @@ lean_object* l_Lean_instantiateMVarsCore(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_FVarId_findDecl_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_instInhabitedMetaM___rarg(lean_object*); static lean_object* l_Lean_Meta_instMonadMetaM___closed__6; +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18865_(lean_object*); uint8_t l_Lean_Name_quickCmp(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___rarg___lambda__1(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_addTraceNodeCore___at_Lean_Meta_processPostponed___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_Meta_getParamNames___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_instInhabitedParamInfo; uint8_t l_Lean_Meta_TransparencyMode_lt(uint8_t, uint8_t); @@ -866,6 +874,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_etaExpand___lambda__1___boxed(lean_object*, LEAN_EXPORT lean_object* l_Lean_Meta_instMonadMCtxMetaM___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withExistingLocalDecls(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getLevelMVarDepth___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18865____closed__5; LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_approxDefEqImp___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Cache_inferType___default___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_withNewMCtxDepth(lean_object*); @@ -884,16 +893,17 @@ static lean_object* l_Lean_Meta_throwIsDefEqStuck___rarg___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withNewLocalInstancesImpAux___rarg___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_Meta_instInhabitedState___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_instMonadMCtxMetaM___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_trace_profiler; static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_2812____closed__5; LEAN_EXPORT lean_object* l_Lean_mkFreshLMVarId___at_Lean_Meta_mkFreshLevelMVar___spec__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1___closed__1; lean_object* lean_io_error_to_string(lean_object*); LEAN_EXPORT lean_object* l_Lean_mkFreshLMVarId___at_Lean_Meta_mkFreshLevelMVar___spec__1___rarg(lean_object*, lean_object*); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withOptProfile___at_Lean_Meta_processPostponed___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(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_forallMetaTelescopeReducingAux_process(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 uint8_t l_Lean_Meta_ParamInfo_binderInfo___default; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__4(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_EXPORT lean_object* l_Lean_Meta_MetaM_toIO(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withIncRecDepth___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT 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*); @@ -902,7 +912,7 @@ lean_object* l_Lean_MetavarContext_setMVarKind(lean_object*, lean_object*, uint8 LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at___private_Lean_Meta_Basic_0__Lean_Meta_processPostponedStep___spec__3___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_sortFVarIds(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_useEtaStruct(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___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_Meta_Basic_0__Lean_Meta_isClassExpensive_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_map2MetaM___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_List_foldr___at___private_Lean_Meta_Basic_0__Lean_Meta_exposeRelevantUniverses___spec__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_instMonadMCtxMetaM___closed__1; @@ -914,6 +924,7 @@ static lean_object* l_Lean_Meta_State_mctx___default___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_MetaM_run_x27(lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_map2MetaM(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Meta_processPostponed___spec__5(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_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withMCtx___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_filterMapM___at_Lean_Meta_withInstImplicitAsImplict___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -947,7 +958,6 @@ uint8_t lean_nat_dec_le(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_instAddMessageContextMetaM; LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_Meta_normalizeLevel___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withAtLeastTransparency___rarg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); -static lean_object* l_Lean_withOptProfile___at_Lean_Meta_processPostponed___spec__3___closed__1; static lean_object* l_Lean_Meta_instInhabitedState___closed__1; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Basic_0__Lean_Meta_processPostponedStep___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); @@ -964,12 +974,12 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_instantiateLam static lean_object* l_Lean_Meta_mkFunUnit___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_processPostponed___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT 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___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withNewLocalInstance___at___private_Lean_Meta_Basic_0__Lean_Meta_withNewLocalInstancesImp___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getFVarLocalDecl___boxed(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_Lean_Meta_instantiateLambda(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_modifyDefEqCache(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_FVarId_throwUnknown___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_lam___override(lean_object*, lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_processPostponed___spec__2___rarg___boxed(lean_object*, lean_object*); @@ -988,6 +998,7 @@ static lean_object* l_Lean_Meta_State_postponed___default___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_forallMetaTelescopeReducingAux___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_MVarId_setKind___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withMVarContext(lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__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_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_instantiateForall(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Meta_withIncRecDepth___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withoutProofIrrelevance(lean_object*); @@ -1034,7 +1045,6 @@ static lean_object* l_Lean_Meta_instBEqInfoCacheKey___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_withDefault___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_FVarId_getValue_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getLevelMVarDepth(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_withOptProfile___at_Lean_Meta_processPostponed___spec__3___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_beqInfoCacheKey____x40_Lean_Meta_Basic___hyg_368____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_whnfR(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getLocalDecl___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1049,7 +1059,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMes LEAN_EXPORT lean_object* l_Lean_Expr_ReplaceImpl_replaceUnsafeM_visit___at___private_Lean_Meta_Basic_0__Lean_Meta_exposeRelevantUniverses___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withMCtxImp(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___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___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_processPostponed___spec__6(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_Lean_Meta_processPostponed(uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Array_isEmpty___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mkFreshExprMVarAt___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1449,7 +1458,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_12336_(x_3, x_6); +x_9 = l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_12510_(x_3, x_6); if (x_9 == 0) { uint8_t x_10; @@ -3840,7 +3849,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_instMetaEvalMetaM___rarg(lean_object* x_1, { lean_object* x_7; lean_object* x_8; lean_object* x_29; lean_object* x_30; lean_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_29 = l_Lean_Meta_instMetaEvalMetaM___rarg___closed__5; -x_30 = l_Lean_Option_get___at_Std_Format_pretty_x27___spec__1(x_3, x_29); +x_30 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_3, x_29); x_31 = lean_box(0); x_32 = l_Lean_Core_getMaxHeartbeats(x_3); x_33 = l_Lean_Meta_instMetaEvalMetaM___rarg___closed__6; @@ -7043,7 +7052,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_12336_(x_10, x_9); +x_11 = l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_12510_(x_10, x_9); x_12 = lean_box(x_11); lean_ctor_set(x_6, 0, x_12); return x_6; @@ -7059,7 +7068,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_12336_(x_16, x_15); +x_17 = l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_12510_(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); @@ -7093,7 +7102,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_12336_(x_10, x_9); +x_11 = l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_12510_(x_10, x_9); x_12 = lean_box(x_11); lean_ctor_set(x_6, 0, x_12); return x_6; @@ -7109,7 +7118,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_12336_(x_16, x_15); +x_17 = l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_12510_(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); @@ -15688,84 +15697,186 @@ lean_dec(x_1); x_22 = l_Lean_Expr_getAppFn(x_21); lean_dec(x_21); switch (lean_obj_tag(x_22)) { -case 4: +case 2: { -lean_object* x_23; lean_object* x_24; +lean_object* x_23; lean_object* x_24; lean_object* x_25; x_23 = lean_ctor_get(x_22, 0); lean_inc(x_23); lean_dec(x_22); -x_24 = l___private_Lean_Meta_Basic_0__Lean_Meta_isClassQuickConst_x3f(x_23, x_2, x_3, x_4, x_5, x_6); +x_24 = l_Lean_getExprMVarAssignment_x3f___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassQuick_x3f___spec__1(x_23, x_2, x_3, x_4, x_5, x_6); +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +if (lean_obj_tag(x_25) == 0) +{ +uint8_t x_26; +x_26 = !lean_is_exclusive(x_24); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; +x_27 = lean_ctor_get(x_24, 0); +lean_dec(x_27); +x_28 = lean_box(0); +lean_ctor_set(x_24, 0, x_28); +return x_24; +} +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_24, 1); +lean_inc(x_29); +lean_dec(x_24); +x_30 = lean_box(0); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_29); +return x_31; +} +} +else +{ +uint8_t x_32; +x_32 = !lean_is_exclusive(x_24); +if (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_24, 1); +x_34 = lean_ctor_get(x_24, 0); +lean_dec(x_34); +x_35 = lean_ctor_get(x_25, 0); +lean_inc(x_35); +lean_dec(x_25); +x_36 = l_Lean_Expr_getAppFn(x_35); +lean_dec(x_35); +if (lean_obj_tag(x_36) == 4) +{ +lean_object* x_37; lean_object* x_38; +lean_free_object(x_24); +x_37 = lean_ctor_get(x_36, 0); +lean_inc(x_37); +lean_dec(x_36); +x_38 = l___private_Lean_Meta_Basic_0__Lean_Meta_isClassQuickConst_x3f(x_37, x_2, x_3, x_4, x_5, x_33); +return x_38; +} +else +{ +lean_object* x_39; +lean_dec(x_36); +x_39 = lean_box(2); +lean_ctor_set(x_24, 0, x_39); return x_24; } +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_24, 1); +lean_inc(x_40); +lean_dec(x_24); +x_41 = lean_ctor_get(x_25, 0); +lean_inc(x_41); +lean_dec(x_25); +x_42 = l_Lean_Expr_getAppFn(x_41); +lean_dec(x_41); +if (lean_obj_tag(x_42) == 4) +{ +lean_object* x_43; lean_object* x_44; +x_43 = lean_ctor_get(x_42, 0); +lean_inc(x_43); +lean_dec(x_42); +x_44 = l___private_Lean_Meta_Basic_0__Lean_Meta_isClassQuickConst_x3f(x_43, x_2, x_3, x_4, x_5, x_40); +return x_44; +} +else +{ +lean_object* x_45; lean_object* x_46; +lean_dec(x_42); +x_45 = lean_box(2); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_40); +return x_46; +} +} +} +} +case 4: +{ +lean_object* x_47; lean_object* x_48; +x_47 = lean_ctor_get(x_22, 0); +lean_inc(x_47); +lean_dec(x_22); +x_48 = l___private_Lean_Meta_Basic_0__Lean_Meta_isClassQuickConst_x3f(x_47, x_2, x_3, x_4, x_5, x_6); +return x_48; +} case 6: { -lean_object* x_25; lean_object* x_26; +lean_object* x_49; lean_object* x_50; lean_dec(x_22); -x_25 = lean_box(2); -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_6); -return x_26; +x_49 = lean_box(2); +x_50 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_50, 0, x_49); +lean_ctor_set(x_50, 1, x_6); +return x_50; } default: { -lean_object* x_27; lean_object* x_28; +lean_object* x_51; lean_object* x_52; lean_dec(x_22); -x_27 = lean_box(0); -x_28 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_28, 0, x_27); -lean_ctor_set(x_28, 1, x_6); -return x_28; +x_51 = lean_box(0); +x_52 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_52, 0, x_51); +lean_ctor_set(x_52, 1, x_6); +return x_52; } } } case 7: { -lean_object* x_29; -x_29 = lean_ctor_get(x_1, 2); -lean_inc(x_29); +lean_object* x_53; +x_53 = lean_ctor_get(x_1, 2); +lean_inc(x_53); lean_dec(x_1); -x_1 = x_29; +x_1 = x_53; goto _start; } case 8: { -lean_object* x_31; lean_object* x_32; +lean_object* x_55; lean_object* x_56; lean_dec(x_1); -x_31 = lean_box(2); -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_32, 1, x_6); -return x_32; +x_55 = lean_box(2); +x_56 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_56, 0, x_55); +lean_ctor_set(x_56, 1, x_6); +return x_56; } case 10: { -lean_object* x_33; -x_33 = lean_ctor_get(x_1, 1); -lean_inc(x_33); +lean_object* x_57; +x_57 = lean_ctor_get(x_1, 1); +lean_inc(x_57); lean_dec(x_1); -x_1 = x_33; +x_1 = x_57; goto _start; } case 11: { -lean_object* x_35; lean_object* x_36; +lean_object* x_59; lean_object* x_60; lean_dec(x_1); -x_35 = lean_box(2); -x_36 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_36, 0, x_35); -lean_ctor_set(x_36, 1, x_6); -return x_36; +x_59 = lean_box(2); +x_60 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_60, 0, x_59); +lean_ctor_set(x_60, 1, x_6); +return x_60; } default: { -lean_object* x_37; lean_object* x_38; +lean_object* x_61; lean_object* x_62; lean_dec(x_1); -x_37 = lean_box(0); -x_38 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_38, 0, x_37); -lean_ctor_set(x_38, 1, x_6); -return x_38; +x_61 = lean_box(0); +x_62 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_62, 0, x_61); +lean_ctor_set(x_62, 1, x_6); +return x_62; } } } @@ -15976,6 +16087,491 @@ x_4 = lean_box(x_3); return x_4; } } +LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___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: +{ +uint8_t x_7; +x_7 = l_Lean_Expr_hasMVar(x_1); +if (x_7 == 0) +{ +lean_object* x_8; +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_1); +lean_ctor_set(x_8, 1, x_6); +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; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_9 = lean_st_ref_get(x_3, 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 = lean_ctor_get(x_10, 0); +lean_inc(x_12); +lean_dec(x_10); +x_13 = l_Lean_instantiateMVarsCore(x_12, x_1); +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_st_ref_take(x_3, x_11); +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_17); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_20 = lean_ctor_get(x_17, 0); +lean_dec(x_20); +lean_ctor_set(x_17, 0, x_15); +x_21 = lean_st_ref_set(x_3, x_17, x_18); +x_22 = !lean_is_exclusive(x_21); +if (x_22 == 0) +{ +lean_object* x_23; +x_23 = lean_ctor_get(x_21, 0); +lean_dec(x_23); +lean_ctor_set(x_21, 0, x_14); +return x_21; +} +else +{ +lean_object* x_24; lean_object* x_25; +x_24 = lean_ctor_get(x_21, 1); +lean_inc(x_24); +lean_dec(x_21); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_14); +lean_ctor_set(x_25, 1, x_24); +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; +x_26 = lean_ctor_get(x_17, 1); +x_27 = lean_ctor_get(x_17, 2); +x_28 = lean_ctor_get(x_17, 3); +lean_inc(x_28); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_17); +x_29 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_29, 0, x_15); +lean_ctor_set(x_29, 1, x_26); +lean_ctor_set(x_29, 2, x_27); +lean_ctor_set(x_29, 3, x_28); +x_30 = lean_st_ref_set(x_3, x_29, x_18); +x_31 = lean_ctor_get(x_30, 1); +lean_inc(x_31); +if (lean_is_exclusive(x_30)) { + lean_ctor_release(x_30, 0); + lean_ctor_release(x_30, 1); + x_32 = x_30; +} else { + lean_dec_ref(x_30); + x_32 = lean_box(0); +} +if (lean_is_scalar(x_32)) { + x_33 = lean_alloc_ctor(0, 2, 0); +} else { + x_33 = x_32; +} +lean_ctor_set(x_33, 0, x_14); +lean_ctor_set(x_33, 1, x_31); +return x_33; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___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; lean_object* x_9; lean_object* x_10; uint8_t x_11; lean_object* x_12; +lean_dec(x_2); +x_8 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_1, x_3, x_4, x_5, x_6, 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_dec(x_8); +x_11 = 1; +x_12 = l___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f(x_9, x_11, x_3, x_4, x_5, x_6, x_10); +return x_12; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f(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) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_Expr_getAppFn(x_1); +switch (lean_obj_tag(x_8)) { +case 2: +{ +lean_dec(x_8); +if (x_2 == 0) +{ +lean_object* x_9; lean_object* x_10; +x_9 = lean_box(0); +x_10 = l___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___lambda__1(x_1, x_9, x_3, x_4, x_5, x_6, x_7); +return x_10; +} +else +{ +lean_object* x_11; lean_object* x_12; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +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_7); +return x_12; +} +} +case 4: +{ +lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_13 = lean_ctor_get(x_8, 0); +lean_inc(x_13); +lean_dec(x_8); +x_14 = lean_st_ref_get(x_6, x_7); +x_15 = !lean_is_exclusive(x_14); +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 = lean_ctor_get(x_16, 0); +lean_inc(x_18); +lean_dec(x_16); +lean_inc(x_13); +lean_inc(x_18); +x_19 = lean_is_class(x_18, x_13); +if (x_19 == 0) +{ +lean_object* x_20; +lean_free_object(x_14); +lean_dec(x_13); +x_20 = lean_whnf(x_1, x_3, x_4, x_5, x_6, x_17); +if (lean_obj_tag(x_20) == 0) +{ +uint8_t x_21; +x_21 = !lean_is_exclusive(x_20); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; +x_22 = lean_ctor_get(x_20, 0); +x_23 = l_Lean_Expr_getAppFn(x_22); +lean_dec(x_22); +if (lean_obj_tag(x_23) == 4) +{ +lean_object* x_24; uint8_t x_25; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +lean_dec(x_23); +lean_inc(x_24); +x_25 = lean_is_class(x_18, x_24); +if (x_25 == 0) +{ +lean_object* x_26; +lean_dec(x_24); +x_26 = lean_box(0); +lean_ctor_set(x_20, 0, x_26); +return x_20; +} +else +{ +lean_object* x_27; +x_27 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_27, 0, x_24); +lean_ctor_set(x_20, 0, x_27); +return x_20; +} +} +else +{ +lean_object* x_28; +lean_dec(x_23); +lean_dec(x_18); +x_28 = lean_box(0); +lean_ctor_set(x_20, 0, x_28); +return x_20; +} +} +else +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_20, 0); +x_30 = lean_ctor_get(x_20, 1); +lean_inc(x_30); +lean_inc(x_29); +lean_dec(x_20); +x_31 = l_Lean_Expr_getAppFn(x_29); +lean_dec(x_29); +if (lean_obj_tag(x_31) == 4) +{ +lean_object* x_32; uint8_t x_33; +x_32 = lean_ctor_get(x_31, 0); +lean_inc(x_32); +lean_dec(x_31); +lean_inc(x_32); +x_33 = lean_is_class(x_18, x_32); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; +lean_dec(x_32); +x_34 = lean_box(0); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_30); +return x_35; +} +else +{ +lean_object* x_36; lean_object* x_37; +x_36 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_36, 0, x_32); +x_37 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set(x_37, 1, x_30); +return x_37; +} +} +else +{ +lean_object* x_38; lean_object* x_39; +lean_dec(x_31); +lean_dec(x_18); +x_38 = lean_box(0); +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_30); +return x_39; +} +} +} +else +{ +uint8_t x_40; +lean_dec(x_18); +x_40 = !lean_is_exclusive(x_20); +if (x_40 == 0) +{ +return x_20; +} +else +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = lean_ctor_get(x_20, 0); +x_42 = lean_ctor_get(x_20, 1); +lean_inc(x_42); +lean_inc(x_41); +lean_dec(x_20); +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_dec(x_18); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_44 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_44, 0, x_13); +lean_ctor_set(x_14, 0, x_44); +return x_14; +} +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; +x_45 = lean_ctor_get(x_14, 0); +x_46 = lean_ctor_get(x_14, 1); +lean_inc(x_46); +lean_inc(x_45); +lean_dec(x_14); +x_47 = lean_ctor_get(x_45, 0); +lean_inc(x_47); +lean_dec(x_45); +lean_inc(x_13); +lean_inc(x_47); +x_48 = lean_is_class(x_47, x_13); +if (x_48 == 0) +{ +lean_object* x_49; +lean_dec(x_13); +x_49 = lean_whnf(x_1, x_3, x_4, x_5, x_6, x_46); +if (lean_obj_tag(x_49) == 0) +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_50 = lean_ctor_get(x_49, 0); +lean_inc(x_50); +x_51 = lean_ctor_get(x_49, 1); +lean_inc(x_51); +if (lean_is_exclusive(x_49)) { + lean_ctor_release(x_49, 0); + lean_ctor_release(x_49, 1); + x_52 = x_49; +} else { + lean_dec_ref(x_49); + x_52 = lean_box(0); +} +x_53 = l_Lean_Expr_getAppFn(x_50); +lean_dec(x_50); +if (lean_obj_tag(x_53) == 4) +{ +lean_object* x_54; uint8_t x_55; +x_54 = lean_ctor_get(x_53, 0); +lean_inc(x_54); +lean_dec(x_53); +lean_inc(x_54); +x_55 = lean_is_class(x_47, x_54); +if (x_55 == 0) +{ +lean_object* x_56; lean_object* x_57; +lean_dec(x_54); +x_56 = lean_box(0); +if (lean_is_scalar(x_52)) { + x_57 = lean_alloc_ctor(0, 2, 0); +} else { + x_57 = x_52; +} +lean_ctor_set(x_57, 0, x_56); +lean_ctor_set(x_57, 1, x_51); +return x_57; +} +else +{ +lean_object* x_58; lean_object* x_59; +x_58 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_58, 0, x_54); +if (lean_is_scalar(x_52)) { + x_59 = lean_alloc_ctor(0, 2, 0); +} else { + x_59 = x_52; +} +lean_ctor_set(x_59, 0, x_58); +lean_ctor_set(x_59, 1, x_51); +return x_59; +} +} +else +{ +lean_object* x_60; lean_object* x_61; +lean_dec(x_53); +lean_dec(x_47); +x_60 = lean_box(0); +if (lean_is_scalar(x_52)) { + x_61 = lean_alloc_ctor(0, 2, 0); +} else { + x_61 = x_52; +} +lean_ctor_set(x_61, 0, x_60); +lean_ctor_set(x_61, 1, x_51); +return x_61; +} +} +else +{ +lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; +lean_dec(x_47); +x_62 = lean_ctor_get(x_49, 0); +lean_inc(x_62); +x_63 = lean_ctor_get(x_49, 1); +lean_inc(x_63); +if (lean_is_exclusive(x_49)) { + lean_ctor_release(x_49, 0); + lean_ctor_release(x_49, 1); + x_64 = x_49; +} else { + lean_dec_ref(x_49); + x_64 = lean_box(0); +} +if (lean_is_scalar(x_64)) { + x_65 = lean_alloc_ctor(1, 2, 0); +} else { + x_65 = x_64; +} +lean_ctor_set(x_65, 0, x_62); +lean_ctor_set(x_65, 1, x_63); +return x_65; +} +} +else +{ +lean_object* x_66; lean_object* x_67; +lean_dec(x_47); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_66 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_66, 0, x_13); +x_67 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_67, 0, x_66); +lean_ctor_set(x_67, 1, x_46); +return x_67; +} +} +} +default: +{ +lean_object* x_68; lean_object* x_69; +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +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_7); +return x_69; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___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) { +_start: +{ +lean_object* x_7; +x_7 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__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); +lean_dec(x_2); +return x_7; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___boxed(lean_object* x_1, 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_9; +x_8 = lean_unbox(x_2); +lean_dec(x_2); +x_9 = l___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f(x_1, x_8, x_3, x_4, x_5, x_6, x_7); +return x_9; +} +} LEAN_EXPORT lean_object* l_Lean_mkFreshFVarId___at___private_Lean_Meta_Basic_0__Lean_Meta_forallTelescopeReducingAuxAux_process___spec__1___rarg(lean_object* x_1, lean_object* x_2) { _start: { @@ -16557,360 +17153,41 @@ x_2 = lean_alloc_closure((void*)(l___private_Lean_Meta_Basic_0__Lean_Meta_withNe return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_isClassExpensive_x3f___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___private_Lean_Meta_Basic_0__Lean_Meta_isClassExpensive_x3f___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_9; uint8_t x_10; -x_9 = lean_st_ref_get(x_7, x_8); -x_10 = !lean_is_exclusive(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_ctor_get(x_9, 0); -x_12 = lean_ctor_get(x_9, 1); -x_13 = lean_ctor_get(x_11, 0); -lean_inc(x_13); -lean_dec(x_11); -x_14 = l_Lean_Expr_getAppFn(x_3); -if (lean_obj_tag(x_14) == 4) -{ -lean_object* x_15; uint8_t x_16; -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -lean_dec(x_14); -lean_inc(x_15); -lean_inc(x_13); -x_16 = lean_is_class(x_13, x_15); -if (x_16 == 0) -{ -lean_object* x_17; -lean_dec(x_15); -lean_free_object(x_9); -x_17 = lean_whnf(x_3, x_4, x_5, x_6, x_7, x_12); -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; -x_19 = lean_ctor_get(x_17, 0); -x_20 = l_Lean_Expr_getAppFn(x_19); -lean_dec(x_19); -if (lean_obj_tag(x_20) == 4) -{ -lean_object* x_21; uint8_t x_22; -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); -lean_dec(x_20); -lean_inc(x_21); -x_22 = lean_is_class(x_13, x_21); -if (x_22 == 0) -{ -lean_dec(x_21); -lean_ctor_set(x_17, 0, x_1); -return x_17; -} -else -{ -lean_object* x_23; -lean_dec(x_1); -x_23 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_23, 0, x_21); -lean_ctor_set(x_17, 0, x_23); -return x_17; -} -} -else -{ -lean_dec(x_20); -lean_dec(x_13); -lean_ctor_set(x_17, 0, x_1); -return x_17; -} -} -else -{ -lean_object* x_24; lean_object* x_25; lean_object* x_26; -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 = l_Lean_Expr_getAppFn(x_24); -lean_dec(x_24); -if (lean_obj_tag(x_26) == 4) -{ -lean_object* x_27; uint8_t x_28; -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -lean_dec(x_26); -lean_inc(x_27); -x_28 = lean_is_class(x_13, x_27); -if (x_28 == 0) -{ -lean_object* x_29; -lean_dec(x_27); -x_29 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_29, 0, x_1); -lean_ctor_set(x_29, 1, x_25); -return x_29; -} -else -{ -lean_object* x_30; lean_object* x_31; -lean_dec(x_1); -x_30 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_30, 0, x_27); -x_31 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_25); -return x_31; -} -} -else -{ -lean_object* x_32; -lean_dec(x_26); -lean_dec(x_13); -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_1); -lean_ctor_set(x_32, 1, x_25); -return x_32; -} -} -} -else -{ -uint8_t x_33; -lean_dec(x_13); -lean_dec(x_1); -x_33 = !lean_is_exclusive(x_17); -if (x_33 == 0) -{ -return x_17; -} -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_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; -} -} -} -else -{ -lean_object* x_37; -lean_dec(x_13); -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_37 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_37, 0, x_15); -lean_ctor_set(x_9, 0, x_37); -return x_9; -} -} -else -{ -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_ctor_set(x_9, 0, x_1); +uint8_t x_8; lean_object* x_9; +x_8 = 0; +x_9 = l___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f(x_2, x_8, x_3, x_4, x_5, x_6, x_7); return x_9; } } -else -{ -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_38 = lean_ctor_get(x_9, 0); -x_39 = lean_ctor_get(x_9, 1); -lean_inc(x_39); -lean_inc(x_38); -lean_dec(x_9); -x_40 = lean_ctor_get(x_38, 0); -lean_inc(x_40); -lean_dec(x_38); -x_41 = l_Lean_Expr_getAppFn(x_3); -if (lean_obj_tag(x_41) == 4) -{ -lean_object* x_42; uint8_t x_43; -x_42 = lean_ctor_get(x_41, 0); -lean_inc(x_42); -lean_dec(x_41); -lean_inc(x_42); -lean_inc(x_40); -x_43 = lean_is_class(x_40, x_42); -if (x_43 == 0) -{ -lean_object* x_44; -lean_dec(x_42); -x_44 = lean_whnf(x_3, x_4, x_5, x_6, x_7, x_39); -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, 0); -lean_inc(x_45); -x_46 = lean_ctor_get(x_44, 1); -lean_inc(x_46); -if (lean_is_exclusive(x_44)) { - lean_ctor_release(x_44, 0); - lean_ctor_release(x_44, 1); - x_47 = x_44; -} else { - lean_dec_ref(x_44); - x_47 = lean_box(0); -} -x_48 = l_Lean_Expr_getAppFn(x_45); -lean_dec(x_45); -if (lean_obj_tag(x_48) == 4) -{ -lean_object* x_49; uint8_t x_50; -x_49 = lean_ctor_get(x_48, 0); -lean_inc(x_49); -lean_dec(x_48); -lean_inc(x_49); -x_50 = lean_is_class(x_40, x_49); -if (x_50 == 0) -{ -lean_object* x_51; -lean_dec(x_49); -if (lean_is_scalar(x_47)) { - x_51 = lean_alloc_ctor(0, 2, 0); -} else { - x_51 = x_47; -} -lean_ctor_set(x_51, 0, x_1); -lean_ctor_set(x_51, 1, x_46); -return x_51; -} -else -{ -lean_object* x_52; lean_object* x_53; -lean_dec(x_1); -x_52 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_52, 0, x_49); -if (lean_is_scalar(x_47)) { - x_53 = lean_alloc_ctor(0, 2, 0); -} else { - x_53 = x_47; -} -lean_ctor_set(x_53, 0, x_52); -lean_ctor_set(x_53, 1, x_46); -return x_53; -} -} -else -{ -lean_object* x_54; -lean_dec(x_48); -lean_dec(x_40); -if (lean_is_scalar(x_47)) { - x_54 = lean_alloc_ctor(0, 2, 0); -} else { - x_54 = x_47; -} -lean_ctor_set(x_54, 0, x_1); -lean_ctor_set(x_54, 1, x_46); -return x_54; -} -} -else -{ -lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; -lean_dec(x_40); -lean_dec(x_1); -x_55 = lean_ctor_get(x_44, 0); -lean_inc(x_55); -x_56 = lean_ctor_get(x_44, 1); -lean_inc(x_56); -if (lean_is_exclusive(x_44)) { - lean_ctor_release(x_44, 0); - lean_ctor_release(x_44, 1); - x_57 = x_44; -} else { - lean_dec_ref(x_44); - x_57 = lean_box(0); -} -if (lean_is_scalar(x_57)) { - x_58 = lean_alloc_ctor(1, 2, 0); -} else { - x_58 = x_57; -} -lean_ctor_set(x_58, 0, x_55); -lean_ctor_set(x_58, 1, x_56); -return x_58; -} -} -else -{ -lean_object* x_59; lean_object* x_60; -lean_dec(x_40); -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_59 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_59, 0, x_42); -x_60 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_60, 0, x_59); -lean_ctor_set(x_60, 1, x_39); -return x_60; -} -} -else +static lean_object* _init_l___private_Lean_Meta_Basic_0__Lean_Meta_isClassExpensive_x3f___closed__1() { +_start: { -lean_object* x_61; -lean_dec(x_41); -lean_dec(x_40); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -x_61 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_61, 0, x_1); -lean_ctor_set(x_61, 1, x_39); -return x_61; -} -} +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_Basic_0__Lean_Meta_isClassExpensive_x3f___lambda__1___boxed), 7, 0); +return x_1; } } LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_isClassExpensive_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) { _start: { -lean_object* x_7; lean_object* x_8; uint8_t x_9; +lean_object* x_7; uint8_t x_8; x_7 = lean_box(0); -x_8 = lean_alloc_closure((void*)(l___private_Lean_Meta_Basic_0__Lean_Meta_isClassExpensive_x3f___lambda__1___boxed), 8, 1); -lean_closure_set(x_8, 0, x_7); -x_9 = !lean_is_exclusive(x_2); -if (x_9 == 0) +x_8 = !lean_is_exclusive(x_2); +if (x_8 == 0) { -lean_object* x_10; uint8_t x_11; -x_10 = lean_ctor_get(x_2, 0); -x_11 = !lean_is_exclusive(x_10); -if (x_11 == 0) +lean_object* x_9; uint8_t x_10; +x_9 = lean_ctor_get(x_2, 0); +x_10 = !lean_is_exclusive(x_9); +if (x_10 == 0) { -uint8_t x_12; lean_object* x_13; -x_12 = 2; -lean_ctor_set_uint8(x_10, 5, x_12); -x_13 = l___private_Lean_Meta_Basic_0__Lean_Meta_forallTelescopeReducingAux___rarg(x_1, x_7, x_8, x_2, x_3, x_4, x_5, x_6); +uint8_t x_11; lean_object* x_12; lean_object* x_13; +x_11 = 2; +lean_ctor_set_uint8(x_9, 5, x_11); +x_12 = l___private_Lean_Meta_Basic_0__Lean_Meta_isClassExpensive_x3f___closed__1; +x_13 = l___private_Lean_Meta_Basic_0__Lean_Meta_forallTelescopeReducingAux___rarg(x_1, x_7, x_12, x_2, x_3, x_4, x_5, x_6); if (lean_obj_tag(x_13) == 0) { uint8_t x_14; @@ -16958,20 +17235,20 @@ return x_21; } else { -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; uint8_t x_31; uint8_t x_32; uint8_t x_33; uint8_t x_34; lean_object* x_35; lean_object* x_36; -x_22 = lean_ctor_get_uint8(x_10, 0); -x_23 = lean_ctor_get_uint8(x_10, 1); -x_24 = lean_ctor_get_uint8(x_10, 2); -x_25 = lean_ctor_get_uint8(x_10, 3); -x_26 = lean_ctor_get_uint8(x_10, 4); -x_27 = lean_ctor_get_uint8(x_10, 6); -x_28 = lean_ctor_get_uint8(x_10, 7); -x_29 = lean_ctor_get_uint8(x_10, 8); -x_30 = lean_ctor_get_uint8(x_10, 9); -x_31 = lean_ctor_get_uint8(x_10, 10); -x_32 = lean_ctor_get_uint8(x_10, 11); -x_33 = lean_ctor_get_uint8(x_10, 12); -lean_dec(x_10); +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; uint8_t x_31; uint8_t x_32; uint8_t x_33; uint8_t x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_22 = lean_ctor_get_uint8(x_9, 0); +x_23 = lean_ctor_get_uint8(x_9, 1); +x_24 = lean_ctor_get_uint8(x_9, 2); +x_25 = lean_ctor_get_uint8(x_9, 3); +x_26 = lean_ctor_get_uint8(x_9, 4); +x_27 = lean_ctor_get_uint8(x_9, 6); +x_28 = lean_ctor_get_uint8(x_9, 7); +x_29 = lean_ctor_get_uint8(x_9, 8); +x_30 = lean_ctor_get_uint8(x_9, 9); +x_31 = lean_ctor_get_uint8(x_9, 10); +x_32 = lean_ctor_get_uint8(x_9, 11); +x_33 = lean_ctor_get_uint8(x_9, 12); +lean_dec(x_9); x_34 = 2; x_35 = lean_alloc_ctor(0, 0, 13); lean_ctor_set_uint8(x_35, 0, x_22); @@ -16988,165 +17265,167 @@ lean_ctor_set_uint8(x_35, 10, x_31); lean_ctor_set_uint8(x_35, 11, x_32); lean_ctor_set_uint8(x_35, 12, x_33); lean_ctor_set(x_2, 0, x_35); -x_36 = l___private_Lean_Meta_Basic_0__Lean_Meta_forallTelescopeReducingAux___rarg(x_1, x_7, x_8, x_2, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_36) == 0) +x_36 = l___private_Lean_Meta_Basic_0__Lean_Meta_isClassExpensive_x3f___closed__1; +x_37 = l___private_Lean_Meta_Basic_0__Lean_Meta_forallTelescopeReducingAux___rarg(x_1, x_7, x_36, x_2, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_37) == 0) { -lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -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; lean_object* x_41; +x_38 = lean_ctor_get(x_37, 0); 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; +x_39 = lean_ctor_get(x_37, 1); +lean_inc(x_39); +if (lean_is_exclusive(x_37)) { + lean_ctor_release(x_37, 0); + lean_ctor_release(x_37, 1); + x_40 = x_37; } else { - lean_dec_ref(x_36); - x_39 = lean_box(0); + lean_dec_ref(x_37); + x_40 = lean_box(0); } -if (lean_is_scalar(x_39)) { - x_40 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_40)) { + x_41 = lean_alloc_ctor(0, 2, 0); } else { - x_40 = x_39; + x_41 = x_40; } -lean_ctor_set(x_40, 0, x_37); -lean_ctor_set(x_40, 1, x_38); -return x_40; +lean_ctor_set(x_41, 0, x_38); +lean_ctor_set(x_41, 1, x_39); +return x_41; } else { -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_41 = lean_ctor_get(x_36, 0); -lean_inc(x_41); -x_42 = lean_ctor_get(x_36, 1); +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_42 = lean_ctor_get(x_37, 0); lean_inc(x_42); -if (lean_is_exclusive(x_36)) { - lean_ctor_release(x_36, 0); - lean_ctor_release(x_36, 1); - x_43 = x_36; +x_43 = lean_ctor_get(x_37, 1); +lean_inc(x_43); +if (lean_is_exclusive(x_37)) { + lean_ctor_release(x_37, 0); + lean_ctor_release(x_37, 1); + x_44 = x_37; } else { - lean_dec_ref(x_36); - x_43 = lean_box(0); + lean_dec_ref(x_37); + x_44 = lean_box(0); } -if (lean_is_scalar(x_43)) { - x_44 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_44)) { + x_45 = lean_alloc_ctor(1, 2, 0); } else { - x_44 = x_43; + x_45 = x_44; } -lean_ctor_set(x_44, 0, x_41); -lean_ctor_set(x_44, 1, x_42); -return x_44; +lean_ctor_set(x_45, 0, x_42); +lean_ctor_set(x_45, 1, x_43); +return x_45; } } } 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; uint8_t x_52; uint8_t x_53; uint8_t x_54; uint8_t x_55; uint8_t x_56; uint8_t x_57; uint8_t x_58; uint8_t x_59; uint8_t x_60; uint8_t x_61; uint8_t x_62; lean_object* x_63; uint8_t x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; -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_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; uint8_t x_53; uint8_t x_54; uint8_t x_55; uint8_t x_56; uint8_t x_57; uint8_t x_58; uint8_t x_59; uint8_t x_60; uint8_t x_61; uint8_t x_62; uint8_t x_63; lean_object* x_64; uint8_t x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_46 = lean_ctor_get(x_2, 0); +x_47 = lean_ctor_get(x_2, 1); +x_48 = lean_ctor_get(x_2, 2); +x_49 = lean_ctor_get(x_2, 3); +x_50 = lean_ctor_get(x_2, 4); +x_51 = lean_ctor_get(x_2, 5); +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_dec(x_2); -x_51 = lean_ctor_get_uint8(x_45, 0); -x_52 = lean_ctor_get_uint8(x_45, 1); -x_53 = lean_ctor_get_uint8(x_45, 2); -x_54 = lean_ctor_get_uint8(x_45, 3); -x_55 = lean_ctor_get_uint8(x_45, 4); -x_56 = lean_ctor_get_uint8(x_45, 6); -x_57 = lean_ctor_get_uint8(x_45, 7); -x_58 = lean_ctor_get_uint8(x_45, 8); -x_59 = lean_ctor_get_uint8(x_45, 9); -x_60 = lean_ctor_get_uint8(x_45, 10); -x_61 = lean_ctor_get_uint8(x_45, 11); -x_62 = lean_ctor_get_uint8(x_45, 12); -if (lean_is_exclusive(x_45)) { - x_63 = x_45; -} else { - lean_dec_ref(x_45); - x_63 = lean_box(0); -} -x_64 = 2; -if (lean_is_scalar(x_63)) { - x_65 = lean_alloc_ctor(0, 0, 13); +x_52 = lean_ctor_get_uint8(x_46, 0); +x_53 = lean_ctor_get_uint8(x_46, 1); +x_54 = lean_ctor_get_uint8(x_46, 2); +x_55 = lean_ctor_get_uint8(x_46, 3); +x_56 = lean_ctor_get_uint8(x_46, 4); +x_57 = lean_ctor_get_uint8(x_46, 6); +x_58 = lean_ctor_get_uint8(x_46, 7); +x_59 = lean_ctor_get_uint8(x_46, 8); +x_60 = lean_ctor_get_uint8(x_46, 9); +x_61 = lean_ctor_get_uint8(x_46, 10); +x_62 = lean_ctor_get_uint8(x_46, 11); +x_63 = lean_ctor_get_uint8(x_46, 12); +if (lean_is_exclusive(x_46)) { + x_64 = x_46; } else { - x_65 = x_63; + lean_dec_ref(x_46); + x_64 = lean_box(0); } -lean_ctor_set_uint8(x_65, 0, x_51); -lean_ctor_set_uint8(x_65, 1, x_52); -lean_ctor_set_uint8(x_65, 2, x_53); -lean_ctor_set_uint8(x_65, 3, x_54); -lean_ctor_set_uint8(x_65, 4, x_55); -lean_ctor_set_uint8(x_65, 5, x_64); -lean_ctor_set_uint8(x_65, 6, x_56); -lean_ctor_set_uint8(x_65, 7, x_57); -lean_ctor_set_uint8(x_65, 8, x_58); -lean_ctor_set_uint8(x_65, 9, x_59); -lean_ctor_set_uint8(x_65, 10, x_60); -lean_ctor_set_uint8(x_65, 11, x_61); -lean_ctor_set_uint8(x_65, 12, x_62); -x_66 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_66, 0, x_65); -lean_ctor_set(x_66, 1, x_46); -lean_ctor_set(x_66, 2, x_47); -lean_ctor_set(x_66, 3, x_48); -lean_ctor_set(x_66, 4, x_49); -lean_ctor_set(x_66, 5, x_50); -x_67 = l___private_Lean_Meta_Basic_0__Lean_Meta_forallTelescopeReducingAux___rarg(x_1, x_7, x_8, x_66, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_67) == 0) +x_65 = 2; +if (lean_is_scalar(x_64)) { + x_66 = lean_alloc_ctor(0, 0, 13); +} else { + x_66 = x_64; +} +lean_ctor_set_uint8(x_66, 0, x_52); +lean_ctor_set_uint8(x_66, 1, x_53); +lean_ctor_set_uint8(x_66, 2, x_54); +lean_ctor_set_uint8(x_66, 3, x_55); +lean_ctor_set_uint8(x_66, 4, x_56); +lean_ctor_set_uint8(x_66, 5, x_65); +lean_ctor_set_uint8(x_66, 6, x_57); +lean_ctor_set_uint8(x_66, 7, x_58); +lean_ctor_set_uint8(x_66, 8, x_59); +lean_ctor_set_uint8(x_66, 9, x_60); +lean_ctor_set_uint8(x_66, 10, x_61); +lean_ctor_set_uint8(x_66, 11, x_62); +lean_ctor_set_uint8(x_66, 12, x_63); +x_67 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_67, 0, x_66); +lean_ctor_set(x_67, 1, x_47); +lean_ctor_set(x_67, 2, x_48); +lean_ctor_set(x_67, 3, x_49); +lean_ctor_set(x_67, 4, x_50); +lean_ctor_set(x_67, 5, x_51); +x_68 = l___private_Lean_Meta_Basic_0__Lean_Meta_isClassExpensive_x3f___closed__1; +x_69 = l___private_Lean_Meta_Basic_0__Lean_Meta_forallTelescopeReducingAux___rarg(x_1, x_7, x_68, x_67, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_69) == 0) { -lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; -x_68 = lean_ctor_get(x_67, 0); -lean_inc(x_68); -x_69 = lean_ctor_get(x_67, 1); -lean_inc(x_69); -if (lean_is_exclusive(x_67)) { - lean_ctor_release(x_67, 0); - lean_ctor_release(x_67, 1); - x_70 = x_67; +lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; +x_70 = lean_ctor_get(x_69, 0); +lean_inc(x_70); +x_71 = lean_ctor_get(x_69, 1); +lean_inc(x_71); +if (lean_is_exclusive(x_69)) { + lean_ctor_release(x_69, 0); + lean_ctor_release(x_69, 1); + x_72 = x_69; } else { - lean_dec_ref(x_67); - x_70 = lean_box(0); + lean_dec_ref(x_69); + x_72 = lean_box(0); } -if (lean_is_scalar(x_70)) { - x_71 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_72)) { + x_73 = lean_alloc_ctor(0, 2, 0); } else { - x_71 = x_70; + x_73 = x_72; } -lean_ctor_set(x_71, 0, x_68); -lean_ctor_set(x_71, 1, x_69); -return x_71; +lean_ctor_set(x_73, 0, x_70); +lean_ctor_set(x_73, 1, x_71); +return x_73; } else { -lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; -x_72 = lean_ctor_get(x_67, 0); -lean_inc(x_72); -x_73 = lean_ctor_get(x_67, 1); -lean_inc(x_73); -if (lean_is_exclusive(x_67)) { - lean_ctor_release(x_67, 0); - lean_ctor_release(x_67, 1); - x_74 = x_67; +lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; +x_74 = lean_ctor_get(x_69, 0); +lean_inc(x_74); +x_75 = lean_ctor_get(x_69, 1); +lean_inc(x_75); +if (lean_is_exclusive(x_69)) { + lean_ctor_release(x_69, 0); + lean_ctor_release(x_69, 1); + x_76 = x_69; } else { - lean_dec_ref(x_67); - x_74 = lean_box(0); + lean_dec_ref(x_69); + x_76 = lean_box(0); } -if (lean_is_scalar(x_74)) { - x_75 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_76)) { + x_77 = lean_alloc_ctor(1, 2, 0); } else { - x_75 = x_74; + x_77 = x_76; } -lean_ctor_set(x_75, 0, x_72); -lean_ctor_set(x_75, 1, x_73); -return x_75; +lean_ctor_set(x_77, 0, x_74); +lean_ctor_set(x_77, 1, x_75); +return x_77; } } } @@ -17371,13 +17650,13 @@ x_14 = l___private_Lean_Meta_Basic_0__Lean_Meta_forallTelescopeReducingAuxAux_pr return x_14; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_isClassExpensive_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) { +LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_isClassExpensive_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) { _start: { -lean_object* x_9; -x_9 = l___private_Lean_Meta_Basic_0__Lean_Meta_isClassExpensive_x3f___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_2); -return x_9; +lean_object* x_8; +x_8 = l___private_Lean_Meta_Basic_0__Lean_Meta_isClassExpensive_x3f___lambda__1(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___private_Lean_Meta_Basic_0__Lean_Meta_forallTelescopeReducingAuxAux___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { @@ -27628,111 +27907,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___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: -{ -uint8_t x_7; -x_7 = l_Lean_Expr_hasMVar(x_1); -if (x_7 == 0) -{ -lean_object* x_8; -x_8 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_8, 0, x_1); -lean_ctor_set(x_8, 1, x_6); -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; lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_9 = lean_st_ref_get(x_3, 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 = lean_ctor_get(x_10, 0); -lean_inc(x_12); -lean_dec(x_10); -x_13 = l_Lean_instantiateMVarsCore(x_12, x_1); -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_st_ref_take(x_3, x_11); -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_17); -if (x_19 == 0) -{ -lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_20 = lean_ctor_get(x_17, 0); -lean_dec(x_20); -lean_ctor_set(x_17, 0, x_15); -x_21 = lean_st_ref_set(x_3, x_17, x_18); -x_22 = !lean_is_exclusive(x_21); -if (x_22 == 0) -{ -lean_object* x_23; -x_23 = lean_ctor_get(x_21, 0); -lean_dec(x_23); -lean_ctor_set(x_21, 0, x_14); -return x_21; -} -else -{ -lean_object* x_24; lean_object* x_25; -x_24 = lean_ctor_get(x_21, 1); -lean_inc(x_24); -lean_dec(x_21); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_14); -lean_ctor_set(x_25, 1, x_24); -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; -x_26 = lean_ctor_get(x_17, 1); -x_27 = lean_ctor_get(x_17, 2); -x_28 = lean_ctor_get(x_17, 3); -lean_inc(x_28); -lean_inc(x_27); -lean_inc(x_26); -lean_dec(x_17); -x_29 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_29, 0, x_15); -lean_ctor_set(x_29, 1, x_26); -lean_ctor_set(x_29, 2, x_27); -lean_ctor_set(x_29, 3, x_28); -x_30 = lean_st_ref_set(x_3, x_29, x_18); -x_31 = lean_ctor_get(x_30, 1); -lean_inc(x_31); -if (lean_is_exclusive(x_30)) { - lean_ctor_release(x_30, 0); - lean_ctor_release(x_30, 1); - x_32 = x_30; -} else { - lean_dec_ref(x_30); - x_32 = lean_box(0); -} -if (lean_is_scalar(x_32)) { - x_33 = lean_alloc_ctor(0, 2, 0); -} else { - x_33 = x_32; -} -lean_ctor_set(x_33, 0, x_14); -lean_ctor_set(x_33, 1, x_31); -return x_33; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__3___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_EXPORT lean_object* l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_9; @@ -27783,11 +27958,11 @@ return x_17; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__3(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__3___rarg), 8, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2___rarg), 8, 0); return x_2; } } @@ -27912,7 +28087,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMes _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; lean_object* x_21; lean_object* x_48; -x_12 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_1, x_7, x_8, x_9, x_10, x_11); +x_12 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_1, x_7, x_8, 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); @@ -27923,7 +28098,7 @@ x_15 = l___private_Lean_Meta_Basic_0__Lean_Meta_exposeRelevantUniverses(x_13, x_ x_16 = lean_ctor_get(x_3, 1); lean_inc(x_16); lean_dec(x_3); -x_17 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_16, x_7, x_8, x_9, x_10, x_14); +x_17 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_16, x_7, x_8, x_9, x_10, x_14); x_18 = lean_ctor_get(x_17, 0); lean_inc(x_18); x_19 = lean_ctor_get(x_17, 1); @@ -28205,7 +28380,7 @@ lean_closure_set(x_36, 2, x_26); lean_closure_set(x_36, 3, x_1); lean_closure_set(x_36, 4, x_29); lean_closure_set(x_36, 5, x_30); -x_37 = l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__3___rarg(x_27, x_28, x_36, x_3, x_4, x_5, x_6, x_7); +x_37 = l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2___rarg(x_27, x_28, x_36, x_3, x_4, x_5, x_6, x_7); return x_37; } } @@ -28220,18 +28395,6 @@ lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___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_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(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); -return x_7; -} -} LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { @@ -30999,20 +31162,7 @@ x_4 = lean_alloc_closure((void*)(l___private_Lean_Util_Trace_0__Lean_getResetTra return x_4; } } -LEAN_EXPORT lean_object* l_Lean_getBoolOption___at_Lean_Meta_processPostponed___spec__4(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) { -_start: -{ -lean_object* x_8; uint8_t x_9; lean_object* x_10; lean_object* x_11; -x_8 = lean_ctor_get(x_5, 2); -x_9 = l_Lean_KVMap_getBool(x_8, x_1, x_2); -x_10 = lean_box(x_9); -x_11 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_11, 0, x_10); -lean_ctor_set(x_11, 1, x_7); -return x_11; -} -} -static double _init_l_Lean_withSeconds___at_Lean_Meta_processPostponed___spec__5___closed__1() { +static double _init_l_Lean_withSeconds___at_Lean_Meta_processPostponed___spec__3___closed__1() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; double x_4; @@ -31023,7 +31173,7 @@ x_4 = l_Float_ofScientific(x_1, x_2, x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_withSeconds___at_Lean_Meta_processPostponed___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_EXPORT lean_object* l_Lean_withSeconds___at_Lean_Meta_processPostponed___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; lean_object* x_8; lean_object* x_9; lean_object* x_10; @@ -31055,262 +31205,68 @@ x_17 = 0; x_18 = lean_unsigned_to_nat(0u); x_19 = l_Float_ofScientific(x_16, x_17, x_18); lean_dec(x_16); -x_20 = l_Lean_withSeconds___at_Lean_Meta_processPostponed___spec__5___closed__1; +x_20 = l_Lean_withSeconds___at_Lean_Meta_processPostponed___spec__3___closed__1; x_21 = lean_float_div(x_19, x_20); x_22 = lean_box_float(x_21); x_23 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_23, 0, x_11); lean_ctor_set(x_23, 1, x_22); -lean_ctor_set(x_13, 0, x_23); -return x_13; -} -else -{ -lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; lean_object* x_28; double x_29; double x_30; double x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_24 = lean_ctor_get(x_13, 0); -x_25 = lean_ctor_get(x_13, 1); -lean_inc(x_25); -lean_inc(x_24); -lean_dec(x_13); -x_26 = lean_nat_sub(x_24, x_8); -lean_dec(x_8); -lean_dec(x_24); -x_27 = 0; -x_28 = lean_unsigned_to_nat(0u); -x_29 = l_Float_ofScientific(x_26, x_27, x_28); -lean_dec(x_26); -x_30 = l_Lean_withSeconds___at_Lean_Meta_processPostponed___spec__5___closed__1; -x_31 = lean_float_div(x_29, x_30); -x_32 = lean_box_float(x_31); -x_33 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_33, 0, x_11); -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_25); -return x_34; -} -} -else -{ -uint8_t x_35; -lean_dec(x_8); -x_35 = !lean_is_exclusive(x_10); -if (x_35 == 0) -{ -return x_10; -} -else -{ -lean_object* x_36; lean_object* x_37; lean_object* x_38; -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 = 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; -} -} -} -} -static lean_object* _init_l_Lean_withOptProfile___at_Lean_Meta_processPostponed___spec__3___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("profiler", 8); -return x_1; -} -} -static lean_object* _init_l_Lean_withOptProfile___at_Lean_Meta_processPostponed___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_Lean_withOptProfile___at_Lean_Meta_processPostponed___spec__3___closed__1; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Lean_withOptProfile___at_Lean_Meta_processPostponed___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; uint8_t x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_7 = l_Lean_withOptProfile___at_Lean_Meta_processPostponed___spec__3___closed__2; -x_8 = 0; -x_9 = l_Lean_getBoolOption___at_Lean_Meta_processPostponed___spec__4(x_7, x_8, x_2, x_3, x_4, x_5, x_6); -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -x_11 = lean_unbox(x_10); -lean_dec(x_10); -if (x_11 == 0) -{ -lean_object* x_12; lean_object* x_13; -x_12 = lean_ctor_get(x_9, 1); -lean_inc(x_12); -lean_dec(x_9); -x_13 = lean_apply_5(x_1, x_2, x_3, x_4, x_5, x_12); -if (lean_obj_tag(x_13) == 0) -{ -uint8_t x_14; -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_13, 0); -x_16 = lean_box(0); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_15); -lean_ctor_set(x_17, 1, x_16); -lean_ctor_set(x_13, 0, x_17); -return x_13; -} -else -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_18 = lean_ctor_get(x_13, 0); -x_19 = lean_ctor_get(x_13, 1); -lean_inc(x_19); -lean_inc(x_18); -lean_dec(x_13); -x_20 = lean_box(0); -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(0, 2, 0); -lean_ctor_set(x_22, 0, x_21); -lean_ctor_set(x_22, 1, x_19); -return x_22; -} -} -else -{ -uint8_t x_23; -x_23 = !lean_is_exclusive(x_13); -if (x_23 == 0) -{ -return x_13; -} -else -{ -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_13, 0); -x_25 = lean_ctor_get(x_13, 1); -lean_inc(x_25); -lean_inc(x_24); -lean_dec(x_13); -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_27; lean_object* x_28; -x_27 = lean_ctor_get(x_9, 1); -lean_inc(x_27); -lean_dec(x_9); -x_28 = l_Lean_withSeconds___at_Lean_Meta_processPostponed___spec__5(x_1, x_2, x_3, x_4, x_5, x_27); -if (lean_obj_tag(x_28) == 0) -{ -uint8_t x_29; -x_29 = !lean_is_exclusive(x_28); -if (x_29 == 0) -{ -lean_object* x_30; uint8_t x_31; -x_30 = lean_ctor_get(x_28, 0); -x_31 = !lean_is_exclusive(x_30); -if (x_31 == 0) -{ -lean_object* x_32; lean_object* x_33; -x_32 = lean_ctor_get(x_30, 1); -x_33 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_33, 0, x_32); -lean_ctor_set(x_30, 1, x_33); -return x_28; -} -else -{ -lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_34 = lean_ctor_get(x_30, 0); -x_35 = lean_ctor_get(x_30, 1); -lean_inc(x_35); -lean_inc(x_34); -lean_dec(x_30); -x_36 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_36, 0, x_35); -x_37 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_37, 0, x_34); -lean_ctor_set(x_37, 1, x_36); -lean_ctor_set(x_28, 0, x_37); -return x_28; -} -} -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; -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_ctor_get(x_38, 0); -lean_inc(x_40); -x_41 = lean_ctor_get(x_38, 1); -lean_inc(x_41); -if (lean_is_exclusive(x_38)) { - lean_ctor_release(x_38, 0); - lean_ctor_release(x_38, 1); - x_42 = x_38; -} else { - lean_dec_ref(x_38); - x_42 = lean_box(0); -} -x_43 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_43, 0, x_41); -if (lean_is_scalar(x_42)) { - x_44 = lean_alloc_ctor(0, 2, 0); -} else { - x_44 = x_42; +lean_ctor_set(x_13, 0, x_23); +return x_13; } -lean_ctor_set(x_44, 0, x_40); -lean_ctor_set(x_44, 1, 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_39); -return x_45; +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; lean_object* x_28; double x_29; double x_30; double x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_24 = lean_ctor_get(x_13, 0); +x_25 = lean_ctor_get(x_13, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_13); +x_26 = lean_nat_sub(x_24, x_8); +lean_dec(x_8); +lean_dec(x_24); +x_27 = 0; +x_28 = lean_unsigned_to_nat(0u); +x_29 = l_Float_ofScientific(x_26, x_27, x_28); +lean_dec(x_26); +x_30 = l_Lean_withSeconds___at_Lean_Meta_processPostponed___spec__3___closed__1; +x_31 = lean_float_div(x_29, x_30); +x_32 = lean_box_float(x_31); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_11); +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_25); +return x_34; } } else { -uint8_t x_46; -x_46 = !lean_is_exclusive(x_28); -if (x_46 == 0) +uint8_t x_35; +lean_dec(x_8); +x_35 = !lean_is_exclusive(x_10); +if (x_35 == 0) { -return x_28; +return x_10; } else { -lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_47 = lean_ctor_get(x_28, 0); -x_48 = lean_ctor_get(x_28, 1); -lean_inc(x_48); -lean_inc(x_47); -lean_dec(x_28); -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_object* x_36; lean_object* x_37; lean_object* x_38; +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 = 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; } } } } -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Meta_processPostponed___spec__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, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Meta_processPostponed___spec__5(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) { _start: { lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; @@ -31429,7 +31385,7 @@ return x_51; } } } -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_processPostponed___spec__6(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_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_processPostponed___spec__4(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) { _start: { uint8_t x_11; @@ -31447,7 +31403,7 @@ lean_inc(x_15); x_16 = lean_ctor_get(x_14, 1); lean_inc(x_16); lean_dec(x_14); -x_17 = l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Meta_processPostponed___spec__7(x_1, x_2, x_3, x_15, x_5, x_6, x_7, x_8, x_9, x_16); +x_17 = l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Meta_processPostponed___spec__5(x_1, x_2, x_3, x_15, x_5, x_6, x_7, x_8, x_9, x_16); lean_dec(x_8); return x_17; } @@ -31497,13 +31453,13 @@ 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_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Meta_processPostponed___spec__7(x_1, x_2, x_3, x_32, x_5, x_6, x_7, x_30, x_9, x_33); +x_34 = l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Meta_processPostponed___spec__5(x_1, x_2, x_3, x_32, x_5, x_6, x_7, x_30, x_9, x_33); lean_dec(x_30); return x_34; } } } -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_processPostponed___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_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_processPostponed___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: { if (lean_obj_tag(x_1) == 0) @@ -31599,16 +31555,24 @@ LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_processPostponed___ { lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_inc(x_10); -x_13 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_processPostponed___spec__6(x_1, x_2, x_3, x_6, x_4, x_8, x_9, x_10, x_11, x_12); +x_13 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_processPostponed___spec__4(x_1, x_2, x_3, x_6, x_4, x_8, x_9, x_10, x_11, x_12); x_14 = lean_ctor_get(x_13, 1); lean_inc(x_14); lean_dec(x_13); -x_15 = l_MonadExcept_ofExcept___at_Lean_Meta_processPostponed___spec__8(x_5, x_8, x_9, x_10, x_11, x_14); +x_15 = l_MonadExcept_ofExcept___at_Lean_Meta_processPostponed___spec__6(x_5, x_8, x_9, x_10, x_11, x_14); lean_dec(x_10); return x_15; } } -static lean_object* _init_l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___closed__1() { +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_profiler; +return x_1; +} +} +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__3___closed__2() { _start: { lean_object* x_1; @@ -31616,16 +31580,16 @@ x_1 = lean_mk_string_from_bytes("[", 1); return x_1; } } -static lean_object* _init_l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___closed__2() { +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__3___closed__3() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___closed__1; +x_1 = l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__3___closed__2; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___closed__3() { +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__3___closed__4() { _start: { lean_object* x_1; @@ -31633,161 +31597,141 @@ x_1 = lean_mk_string_from_bytes("s] ", 3); return x_1; } } -static lean_object* _init_l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___closed__4() { +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__3___closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___closed__3; +x_1 = l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__3___closed__4; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_processPostponed___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_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, uint8_t x_7, double 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_10; lean_object* x_11; uint8_t x_12; -lean_inc(x_1); -x_10 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_1, x_5, x_6, x_7, x_8, x_9); -x_11 = lean_ctor_get(x_10, 0); +lean_object* x_15; lean_object* x_16; +lean_dec(x_9); +x_15 = lean_ctor_get(x_12, 5); +lean_inc(x_15); +lean_inc(x_13); +lean_inc(x_12); lean_inc(x_11); -x_12 = lean_unbox(x_11); -lean_dec(x_11); -if (x_12 == 0) +lean_inc(x_10); +lean_inc(x_2); +x_16 = lean_apply_6(x_1, x_2, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_13; lean_object* x_14; -lean_dec(x_2); -lean_dec(x_1); -x_13 = lean_ctor_get(x_10, 1); -lean_inc(x_13); +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); +x_19 = l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__3___closed__1; +x_20 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_6, x_19); +lean_dec(x_6); +if (x_20 == 0) +{ +if (x_7 == 0) +{ +lean_object* x_21; lean_object* x_22; +x_21 = lean_box(0); +x_22 = l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__2(x_3, x_4, x_15, x_5, x_2, x_17, x_21, x_10, x_11, x_12, x_13, x_18); +lean_dec(x_13); +lean_dec(x_11); lean_dec(x_10); -x_14 = lean_apply_5(x_3, x_5, x_6, x_7, x_8, x_13); -return x_14; +lean_dec(x_2); +return x_22; } 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; -x_15 = lean_ctor_get(x_10, 1); -lean_inc(x_15); +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; 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_23 = lean_float_to_string(x_8); +x_24 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_24, 0, x_23); +x_25 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_25, 0, x_24); +x_26 = l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__3___closed__3; +x_27 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_25); +x_28 = l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__3___closed__5; +x_29 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +x_30 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_17); +x_31 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__1; +x_32 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +x_33 = lean_box(0); +x_34 = l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__2(x_3, x_4, x_15, x_5, x_2, x_32, x_33, x_10, x_11, x_12, x_13, x_18); +lean_dec(x_13); +lean_dec(x_11); lean_dec(x_10); -x_16 = lean_ctor_get(x_7, 5); -lean_inc(x_16); -x_17 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_processPostponed___spec__2___rarg(x_8, 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_alloc_closure((void*)(l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__1), 6, 1); -lean_closure_set(x_20, 0, x_3); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_21 = l_Lean_withOptProfile___at_Lean_Meta_processPostponed___spec__3(x_20, x_5, x_6, x_7, x_8, 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; -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_22, 0); -lean_inc(x_24); -x_25 = lean_ctor_get(x_22, 1); -lean_inc(x_25); -lean_dec(x_22); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_24); -x_26 = lean_apply_6(x_2, x_24, x_5, x_6, x_7, x_8, x_23); -if (lean_obj_tag(x_26) == 0) -{ -if (lean_obj_tag(x_25) == 0) -{ -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; -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_box(0); -x_30 = l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__2(x_18, x_1, x_16, x_4, x_24, x_27, x_29, x_5, x_6, x_7, x_8, x_28); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_24); -return x_30; +lean_dec(x_2); +return x_34; +} } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; double x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; 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_31 = lean_ctor_get(x_26, 0); -lean_inc(x_31); -x_32 = lean_ctor_get(x_26, 1); -lean_inc(x_32); -lean_dec(x_26); -x_33 = lean_ctor_get(x_25, 0); -lean_inc(x_33); -lean_dec(x_25); -x_34 = lean_unbox_float(x_33); -lean_dec(x_33); -x_35 = lean_float_to_string(x_34); +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; 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_float_to_string(x_8); x_36 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_36, 0, x_35); x_37 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_37, 0, x_36); -x_38 = l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___closed__2; +x_38 = l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__3___closed__3; x_39 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_39, 0, x_38); lean_ctor_set(x_39, 1, x_37); -x_40 = l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___closed__4; +x_40 = l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__3___closed__5; x_41 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_41, 0, x_39); lean_ctor_set(x_41, 1, x_40); x_42 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_31); +lean_ctor_set(x_42, 1, x_17); x_43 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__1; x_44 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_44, 0, x_42); lean_ctor_set(x_44, 1, x_43); x_45 = lean_box(0); -x_46 = l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__2(x_18, x_1, x_16, x_4, x_24, x_44, x_45, x_5, x_6, x_7, x_8, x_32); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_24); +x_46 = l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__2(x_3, x_4, x_15, x_5, x_2, x_44, x_45, x_10, x_11, x_12, x_13, x_18); +lean_dec(x_13); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_2); return x_46; } } else { uint8_t x_47; -lean_dec(x_25); -lean_dec(x_24); -lean_dec(x_18); -lean_dec(x_16); -lean_dec(x_8); -lean_dec(x_7); +lean_dec(x_15); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_1); -x_47 = !lean_is_exclusive(x_26); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_47 = !lean_is_exclusive(x_16); if (x_47 == 0) { -return x_26; +return x_16; } else { lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_48 = lean_ctor_get(x_26, 0); -x_49 = lean_ctor_get(x_26, 1); +x_48 = lean_ctor_get(x_16, 0); +x_49 = lean_ctor_get(x_16, 1); lean_inc(x_49); lean_inc(x_48); -lean_dec(x_26); +lean_dec(x_16); x_50 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_50, 0, x_48); lean_ctor_set(x_50, 1, x_49); @@ -31795,36 +31739,376 @@ return x_50; } } } -else +} +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__4___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_trace_profiler; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t 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, 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_dec(x_7); +x_13 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_processPostponed___spec__2___rarg(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_alloc_closure((void*)(l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__1), 6, 1); +lean_closure_set(x_16, 0, x_1); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_17 = l_Lean_withSeconds___at_Lean_Meta_processPostponed___spec__3(x_16, x_8, x_9, x_10, x_11, x_15); +if (lean_obj_tag(x_17) == 0) { -uint8_t x_51; +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; lean_object* x_67; uint8_t x_68; +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_ctor_get(x_18, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_18, 1); +lean_inc(x_21); lean_dec(x_18); -lean_dec(x_16); +x_67 = l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__4___closed__1; +x_68 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_5, x_67); +if (x_68 == 0) +{ +uint8_t x_69; +x_69 = 0; +x_22 = x_69; +goto block_66; +} +else +{ +double x_70; double x_71; uint8_t x_72; +x_70 = l_Lean_trace_profiler_threshold_getSecs(x_5); +x_71 = lean_unbox_float(x_21); +x_72 = lean_float_decLt(x_70, x_71); +x_22 = x_72; +goto block_66; +} +block_66: +{ +if (x_6 == 0) +{ +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; +lean_dec(x_21); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +x_23 = lean_st_ref_take(x_11, 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_is_exclusive(x_24); +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_24, 3); +x_28 = l_Lean_PersistentArray_append___rarg(x_14, x_27); +lean_ctor_set(x_24, 3, x_28); +x_29 = lean_st_ref_set(x_11, x_24, x_25); +x_30 = lean_ctor_get(x_29, 1); +lean_inc(x_30); +lean_dec(x_29); +x_31 = l_MonadExcept_ofExcept___at_Lean_Meta_processPostponed___spec__6(x_20, x_8, x_9, x_10, x_11, x_30); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_20); +if (lean_obj_tag(x_31) == 0) +{ +uint8_t x_32; +x_32 = !lean_is_exclusive(x_31); +if (x_32 == 0) +{ +return x_31; +} +else +{ +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(0, 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_36; +x_36 = !lean_is_exclusive(x_31); +if (x_36 == 0) +{ +return x_31; +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_31, 0); +x_38 = lean_ctor_get(x_31, 1); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_31); +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_40; lean_object* x_41; 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_40 = lean_ctor_get(x_24, 0); +x_41 = lean_ctor_get(x_24, 1); +x_42 = lean_ctor_get(x_24, 2); +x_43 = lean_ctor_get(x_24, 3); +x_44 = lean_ctor_get(x_24, 4); +x_45 = lean_ctor_get(x_24, 5); +x_46 = lean_ctor_get(x_24, 6); +lean_inc(x_46); +lean_inc(x_45); +lean_inc(x_44); +lean_inc(x_43); +lean_inc(x_42); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_24); +x_47 = l_Lean_PersistentArray_append___rarg(x_14, x_43); +x_48 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_48, 0, x_40); +lean_ctor_set(x_48, 1, x_41); +lean_ctor_set(x_48, 2, x_42); +lean_ctor_set(x_48, 3, x_47); +lean_ctor_set(x_48, 4, x_44); +lean_ctor_set(x_48, 5, x_45); +lean_ctor_set(x_48, 6, x_46); +x_49 = lean_st_ref_set(x_11, x_48, x_25); +x_50 = lean_ctor_get(x_49, 1); +lean_inc(x_50); +lean_dec(x_49); +x_51 = l_MonadExcept_ofExcept___at_Lean_Meta_processPostponed___spec__6(x_20, x_8, x_9, x_10, x_11, x_50); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_20); +if (lean_obj_tag(x_51) == 0) +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_52 = lean_ctor_get(x_51, 0); +lean_inc(x_52); +x_53 = lean_ctor_get(x_51, 1); +lean_inc(x_53); +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_51); + x_54 = lean_box(0); +} +if (lean_is_scalar(x_54)) { + x_55 = lean_alloc_ctor(0, 2, 0); +} else { + x_55 = x_54; +} +lean_ctor_set(x_55, 0, x_52); +lean_ctor_set(x_55, 1, x_53); +return x_55; +} +else +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_56 = lean_ctor_get(x_51, 0); +lean_inc(x_56); +x_57 = lean_ctor_get(x_51, 1); +lean_inc(x_57); +if (lean_is_exclusive(x_51)) { + lean_ctor_release(x_51, 0); + lean_ctor_release(x_51, 1); + x_58 = x_51; +} else { + lean_dec_ref(x_51); + x_58 = lean_box(0); +} +if (lean_is_scalar(x_58)) { + x_59 = lean_alloc_ctor(1, 2, 0); +} else { + x_59 = x_58; +} +lean_ctor_set(x_59, 0, x_56); +lean_ctor_set(x_59, 1, x_57); +return x_59; +} +} +} +else +{ +lean_object* x_60; double x_61; lean_object* x_62; +x_60 = lean_box(0); +x_61 = lean_unbox_float(x_21); +lean_dec(x_21); +x_62 = l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__3(x_2, x_20, x_14, x_3, x_4, x_5, x_22, x_61, x_60, x_8, x_9, x_10, x_11, x_19); +return x_62; +} +} +else +{ +lean_object* x_63; double x_64; lean_object* x_65; +x_63 = lean_box(0); +x_64 = lean_unbox_float(x_21); +lean_dec(x_21); +x_65 = l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__3(x_2, x_20, x_14, x_3, x_4, x_5, x_22, x_64, x_63, x_8, x_9, x_10, x_11, x_19); +return x_65; +} +} +} +else +{ +uint8_t x_73; +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_5); +lean_dec(x_3); +lean_dec(x_2); +x_73 = !lean_is_exclusive(x_17); +if (x_73 == 0) +{ +return x_17; +} +else +{ +lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_74 = lean_ctor_get(x_17, 0); +x_75 = lean_ctor_get(x_17, 1); +lean_inc(x_75); +lean_inc(x_74); +lean_dec(x_17); +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_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_processPostponed___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) { +_start: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_10 = lean_ctor_get(x_7, 2); +lean_inc(x_10); +lean_inc(x_1); +x_11 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___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_unbox(x_12); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); +lean_dec(x_11); +x_15 = l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__4___closed__1; +x_16 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_10, x_15); +if (x_16 == 0) +{ +lean_object* x_17; +lean_dec(x_12); +lean_dec(x_10); lean_dec(x_2); lean_dec(x_1); -x_51 = !lean_is_exclusive(x_21); -if (x_51 == 0) +x_17 = lean_apply_5(x_3, x_5, x_6, x_7, x_8, x_14); +if (lean_obj_tag(x_17) == 0) +{ +uint8_t x_18; +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) +{ +return x_17; +} +else { +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_17, 0); +x_20 = lean_ctor_get(x_17, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_17); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); return x_21; } +} else { -lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_52 = lean_ctor_get(x_21, 0); -x_53 = lean_ctor_get(x_21, 1); -lean_inc(x_53); -lean_inc(x_52); -lean_dec(x_21); -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; +uint8_t x_22; +x_22 = !lean_is_exclusive(x_17); +if (x_22 == 0) +{ +return x_17; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_17, 0); +x_24 = lean_ctor_get(x_17, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_17); +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_26; uint8_t x_27; lean_object* x_28; +x_26 = lean_box(0); +x_27 = lean_unbox(x_12); +lean_dec(x_12); +x_28 = l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__4(x_3, x_2, x_1, x_4, x_10, x_27, x_26, x_5, x_6, x_7, x_8, x_14); +return x_28; } } +else +{ +lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; +x_29 = lean_ctor_get(x_11, 1); +lean_inc(x_29); +lean_dec(x_11); +x_30 = lean_box(0); +x_31 = lean_unbox(x_12); +lean_dec(x_12); +x_32 = l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__4(x_3, x_2, x_1, x_4, x_10, x_31, x_30, x_5, x_6, x_7, x_8, x_29); +return x_32; } } } @@ -32003,28 +32287,13 @@ lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_getBoolOption___at_Lean_Meta_processPostponed___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) { -_start: -{ -uint8_t x_8; lean_object* x_9; -x_8 = lean_unbox(x_2); -lean_dec(x_2); -x_9 = l_Lean_getBoolOption___at_Lean_Meta_processPostponed___spec__4(x_1, x_8, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -return x_9; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Meta_processPostponed___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_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Meta_processPostponed___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: { uint8_t x_11; lean_object* x_12; x_11 = lean_unbox(x_5); lean_dec(x_5); -x_12 = l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Meta_processPostponed___spec__7(x_1, x_2, x_3, x_4, x_11, x_6, x_7, x_8, x_9, x_10); +x_12 = l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Meta_processPostponed___spec__5(x_1, x_2, x_3, x_4, x_11, x_6, x_7, x_8, x_9, x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -32032,24 +32301,24 @@ lean_dec(x_6); return x_12; } } -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_processPostponed___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_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_processPostponed___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: { uint8_t x_11; lean_object* x_12; x_11 = lean_unbox(x_5); lean_dec(x_5); -x_12 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_processPostponed___spec__6(x_1, x_2, x_3, x_4, x_11, x_6, x_7, x_8, x_9, x_10); +x_12 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_processPostponed___spec__4(x_1, x_2, x_3, x_4, x_11, x_6, x_7, x_8, x_9, x_10); lean_dec(x_9); lean_dec(x_7); lean_dec(x_6); return x_12; } } -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_processPostponed___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_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_processPostponed___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: { lean_object* x_7; -x_7 = l_MonadExcept_ofExcept___at_Lean_Meta_processPostponed___spec__8(x_1, x_2, x_3, x_4, x_5, x_6); +x_7 = l_MonadExcept_ofExcept___at_Lean_Meta_processPostponed___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); @@ -32073,6 +32342,32 @@ lean_dec(x_5); return x_14; } } +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___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) { +_start: +{ +uint8_t x_15; uint8_t x_16; double x_17; lean_object* x_18; +x_15 = lean_unbox(x_5); +lean_dec(x_5); +x_16 = lean_unbox(x_7); +lean_dec(x_7); +x_17 = lean_unbox_float(x_8); +lean_dec(x_8); +x_18 = l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__3(x_1, x_2, x_3, x_4, x_15, x_6, x_16, x_17, x_9, x_10, x_11, x_12, x_13, x_14); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___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) { +_start: +{ +uint8_t x_13; uint8_t x_14; lean_object* x_15; +x_13 = lean_unbox(x_4); +lean_dec(x_4); +x_14 = lean_unbox(x_6); +lean_dec(x_6); +x_15 = l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__4(x_1, x_2, x_3, x_13, x_5, x_14, x_7, x_8, x_9, x_10, x_11, x_12); +return x_15; +} +} LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_processPostponed___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) { _start: { @@ -35182,7 +35477,7 @@ lean_dec(x_3); return x_9; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18710____closed__1() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18865____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -35192,73 +35487,73 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18710____closed__2() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18865____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18710____closed__1; +x_1 = l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18865____closed__1; x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_2812____closed__8; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18710____closed__3() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18865____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18710____closed__2; +x_1 = l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18865____closed__2; x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_2812____closed__3; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18710____closed__4() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18865____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18710____closed__3; +x_1 = l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18865____closed__3; x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_2812____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18710____closed__5() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18865____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18710____closed__4; +x_1 = l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18865____closed__4; x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_2812____closed__12; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18710____closed__6() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18865____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18710____closed__5; +x_1 = l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18865____closed__5; x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_2812____closed__14; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18710____closed__7() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18865____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18710____closed__6; -x_2 = lean_unsigned_to_nat(18710u); +x_1 = l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18865____closed__6; +x_2 = lean_unsigned_to_nat(18865u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18710_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18865_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; x_2 = l_Lean_Meta_processPostponed_loop___closed__3; x_3 = 0; -x_4 = l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18710____closed__7; +x_4 = l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18865____closed__7; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } @@ -35616,6 +35911,8 @@ l_Lean_Meta_mkFunUnit___closed__4 = _init_l_Lean_Meta_mkFunUnit___closed__4(); lean_mark_persistent(l_Lean_Meta_mkFunUnit___closed__4); l_Lean_Meta_mkFunUnit___closed__5 = _init_l_Lean_Meta_mkFunUnit___closed__5(); lean_mark_persistent(l_Lean_Meta_mkFunUnit___closed__5); +l___private_Lean_Meta_Basic_0__Lean_Meta_isClassExpensive_x3f___closed__1 = _init_l___private_Lean_Meta_Basic_0__Lean_Meta_isClassExpensive_x3f___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_Basic_0__Lean_Meta_isClassExpensive_x3f___closed__1); l_Lean_Meta_getParamNames___closed__1 = _init_l_Lean_Meta_getParamNames___closed__1(); lean_mark_persistent(l_Lean_Meta_getParamNames___closed__1); l_Lean_Meta_setInlineAttribute___closed__1 = _init_l_Lean_Meta_setInlineAttribute___closed__1(); @@ -35688,19 +35985,19 @@ l_Lean_Meta_processPostponed_loop___closed__4 = _init_l_Lean_Meta_processPostpon lean_mark_persistent(l_Lean_Meta_processPostponed_loop___closed__4); l_Lean_Meta_processPostponed_loop___closed__5 = _init_l_Lean_Meta_processPostponed_loop___closed__5(); lean_mark_persistent(l_Lean_Meta_processPostponed_loop___closed__5); -l_Lean_withSeconds___at_Lean_Meta_processPostponed___spec__5___closed__1 = _init_l_Lean_withSeconds___at_Lean_Meta_processPostponed___spec__5___closed__1(); -l_Lean_withOptProfile___at_Lean_Meta_processPostponed___spec__3___closed__1 = _init_l_Lean_withOptProfile___at_Lean_Meta_processPostponed___spec__3___closed__1(); -lean_mark_persistent(l_Lean_withOptProfile___at_Lean_Meta_processPostponed___spec__3___closed__1); -l_Lean_withOptProfile___at_Lean_Meta_processPostponed___spec__3___closed__2 = _init_l_Lean_withOptProfile___at_Lean_Meta_processPostponed___spec__3___closed__2(); -lean_mark_persistent(l_Lean_withOptProfile___at_Lean_Meta_processPostponed___spec__3___closed__2); -l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___closed__1 = _init_l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___closed__1(); -lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___closed__1); -l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___closed__2 = _init_l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___closed__2(); -lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___closed__2); -l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___closed__3 = _init_l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___closed__3(); -lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___closed__3); -l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___closed__4 = _init_l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___closed__4(); -lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___closed__4); +l_Lean_withSeconds___at_Lean_Meta_processPostponed___spec__3___closed__1 = _init_l_Lean_withSeconds___at_Lean_Meta_processPostponed___spec__3___closed__1(); +l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__3___closed__1 = _init_l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__3___closed__1); +l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__3___closed__2 = _init_l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__3___closed__2); +l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__3___closed__3 = _init_l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__3___closed__3(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__3___closed__3); +l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__3___closed__4 = _init_l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__3___closed__4(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__3___closed__4); +l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__3___closed__5 = _init_l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__3___closed__5(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__3___closed__5); +l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__4___closed__1 = _init_l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__4___closed__1); l_Lean_Meta_processPostponed___lambda__1___closed__1 = _init_l_Lean_Meta_processPostponed___lambda__1___closed__1(); lean_mark_persistent(l_Lean_Meta_processPostponed___lambda__1___closed__1); l_Lean_Meta_processPostponed___lambda__1___closed__2 = _init_l_Lean_Meta_processPostponed___lambda__1___closed__2(); @@ -35709,21 +36006,21 @@ l_Lean_Meta_processPostponed___lambda__1___closed__3 = _init_l_Lean_Meta_process lean_mark_persistent(l_Lean_Meta_processPostponed___lambda__1___closed__3); l_Lean_Meta_processPostponed___lambda__1___closed__4 = _init_l_Lean_Meta_processPostponed___lambda__1___closed__4(); lean_mark_persistent(l_Lean_Meta_processPostponed___lambda__1___closed__4); -l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18710____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18710____closed__1(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18710____closed__1); -l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18710____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18710____closed__2(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18710____closed__2); -l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18710____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18710____closed__3(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18710____closed__3); -l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18710____closed__4 = _init_l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18710____closed__4(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18710____closed__4); -l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18710____closed__5 = _init_l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18710____closed__5(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18710____closed__5); -l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18710____closed__6 = _init_l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18710____closed__6(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18710____closed__6); -l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18710____closed__7 = _init_l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18710____closed__7(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18710____closed__7); -res = l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18710_(lean_io_mk_world()); +l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18865____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18865____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18865____closed__1); +l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18865____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18865____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18865____closed__2); +l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18865____closed__3 = _init_l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18865____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18865____closed__3); +l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18865____closed__4 = _init_l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18865____closed__4(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18865____closed__4); +l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18865____closed__5 = _init_l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18865____closed__5(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18865____closed__5); +l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18865____closed__6 = _init_l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18865____closed__6(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18865____closed__6); +l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18865____closed__7 = _init_l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18865____closed__7(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18865____closed__7); +res = l_Lean_initFn____x40_Lean_Meta_Basic___hyg_18865_(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/Meta/Check.c b/stage0/stdlib/Lean/Meta/Check.c index bc98ec0c37b0..988b5d472bc1 100644 --- a/stage0/stdlib/Lean/Meta/Check.c +++ b/stage0/stdlib/Lean/Meta/Check.c @@ -19,15 +19,18 @@ LEAN_EXPORT lean_object* l_Lean_Meta_lambdaLetTelescope___at___private_Lean_Meta lean_object* l___private_Init_Util_0__outOfBounds___rarg(lean_object*); static lean_object* l_Lean_Meta_check___lambda__1___closed__8; static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_3234____closed__14; +extern lean_object* l_Lean_profiler; 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_throwLetTypeMismatchMessage___rarg___closed__8; lean_object* l_Lean_mkAppN(lean_object*, lean_object*); static lean_object* l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__3; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__4(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* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); static lean_object* l_Lean_Meta_check___closed__1; static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_3234____closed__3; static lean_object* l_Lean_Meta_addPPExplicitToExposeDiff___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_check___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__3___closed__3; lean_object* lean_whnf(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_addPPExplicitToExposeDiff_visit___lambda__1___closed__1; lean_object* l_Lean_ConstantInfo_levelParams(lean_object*); @@ -40,12 +43,14 @@ static lean_object* l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__10; static lean_object* l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__3; LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_throwLetTypeMismatchMessage___spec__2(lean_object*); uint8_t l_Lean_Expr_isApp(lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_check___spec__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_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Check_0__Lean_Meta_checkConstant(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +double l_Lean_trace_profiler_threshold_getSecs(lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_3234____closed__13; LEAN_EXPORT lean_object* l_Lean_Meta_addPPExplicitToExposeDiff_visit(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__3___closed__4; lean_object* l_Lean_Expr_sort___override(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Check_0__Lean_Meta_checkAux_checkForall___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_withOptProfile___at_Lean_Meta_check___spec__2___closed__2; lean_object* lean_mk_array(lean_object*, lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); static lean_object* l_Lean_Meta_mkHasTypeButIsExpectedMsg___closed__1; @@ -56,11 +61,11 @@ lean_object* l_Lean_Expr_setAppPPExplicit(lean_object*); lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_addPPExplicitToExposeDiff_hasExplicitDiff_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMapImp_expand___at___private_Lean_Meta_Check_0__Lean_Meta_checkAux_check___spec__5(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withSeconds___at_Lean_Meta_check___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__7; lean_object* l_Lean_Expr_fvarId_x21(lean_object*); static lean_object* l_Lean_Meta_check___lambda__1___closed__2; static lean_object* l___private_Lean_Meta_Check_0__Lean_Meta_checkAux_checkLambdaLet___closed__1; +uint8_t lean_float_decLt(double, double); static lean_object* l_panic___at_Lean_Meta_throwLetTypeMismatchMessage___spec__1___rarg___closed__1; static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_3234____closed__5; static lean_object* l_Lean_Meta_check___closed__3; @@ -70,6 +75,7 @@ static lean_object* l_Lean_Meta_addPPExplicitToExposeDiff___closed__5; lean_object* l_Lean_stringToMessageData(lean_object*); static lean_object* l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__6; static lean_object* l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__1; +LEAN_EXPORT lean_object* l_Lean_withSeconds___at_Lean_Meta_check___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_check___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_lambdaLetTelescope___at___private_Lean_Meta_Check_0__Lean_Meta_checkAux_checkLambdaLet___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Exception_toMessageData(lean_object*); @@ -78,6 +84,9 @@ lean_object* l_Lean_Meta_throwFunctionExpected___rarg(lean_object*, lean_object* static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_3234____closed__10; LEAN_EXPORT lean_object* l_Lean_Meta_addPPExplicitToExposeDiff_hasExplicitDiff_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at___private_Lean_Meta_Check_0__Lean_Meta_checkAux_check___spec__7(lean_object*, lean_object*); +static lean_object* l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__3___closed__5; +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_check___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_check___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_isTypeCorrect(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_throwAppTypeMismatch___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_3234_(lean_object*); @@ -93,7 +102,6 @@ static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_3234____clos uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_replace___at___private_Lean_Meta_Check_0__Lean_Meta_checkAux_check___spec__8(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__9; -static lean_object* l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_throwLetTypeMismatchMessage(lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_3234____closed__11; lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean_object*); @@ -101,10 +109,11 @@ static lean_object* l_Lean_Meta_addPPExplicitToExposeDiff___closed__2; static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_3234____closed__9; LEAN_EXPORT lean_object* l_Lean_Meta_addPPExplicitToExposeDiff_hasExplicitDiff_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_check___lambda__1___closed__3; +lean_object* l_Lean_PersistentArray_append___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___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_EXPORT lean_object* l_Lean_HashMap_insert___at___private_Lean_Meta_Check_0__Lean_Meta_checkAux_check___spec__3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_check(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_withOptProfile___at_Lean_Meta_check___spec__2___closed__1; +static lean_object* l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__3___closed__2; LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_throwLetTypeMismatchMessage___spec__1(lean_object*); lean_object* l_Lean_Meta_getLevel(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_throwAppTypeMismatch___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -112,31 +121,27 @@ lean_object* l_Lean_FVarId_getDecl(lean_object*, lean_object*, lean_object*, lea LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_throwLetTypeMismatchMessage___spec__3___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_contains___at___private_Lean_Meta_Check_0__Lean_Meta_checkAux_check___spec__4___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Meta_addPPExplicitToExposeDiff_hasExplicitDiff_x3f___closed__1; -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_check___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_checkEmoji; lean_object* lean_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_throwAppTypeMismatch(lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); static lean_object* l_Lean_Meta_throwAppTypeMismatch___rarg___lambda__1___closed__4; -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_3234____closed__1; lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_throwLetTypeMismatchMessage___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_throwAppTypeMismatch___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_levelZero; lean_object* lean_io_mono_nanos_now(lean_object*); -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_check___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_instInhabitedExpr; lean_object* l_Lean_Meta_whnfD(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_addPPExplicitToExposeDiff_hasExplicitDiff_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint64_t l_Lean_Expr_hash(lean_object*); lean_object* l_List_lengthTRAux___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_throwLetTypeMismatchMessage___spec__3(lean_object*); -lean_object* l_Lean_getBoolOption___at_Lean_Meta_processPostponed___spec__4(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_lambdaLetTelescope___at___private_Lean_Meta_Check_0__Lean_Meta_checkAux_checkLambdaLet___spec__2(lean_object*); -static lean_object* l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___closed__2; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); uint8_t lean_expr_equal(lean_object*, lean_object*); +uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, 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_Meta_throwAppTypeMismatch___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_Lean_Meta_throwLetTypeMismatchMessage___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -159,8 +164,8 @@ extern lean_object* l_Lean_crossEmoji; LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at___private_Lean_Meta_Check_0__Lean_Meta_checkAux_check___spec__2___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Check_0__Lean_Meta_checkAux(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_addPPExplicitToExposeDiff(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_processPostponed___spec__4(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_foldlMUnsafe_fold___at___private_Lean_Meta_Check_0__Lean_Meta_checkAux_checkLambdaLet___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___closed__4; static lean_object* l_Lean_Meta_addPPExplicitToExposeDiff___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_check___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_check___lambda__1___closed__5; @@ -172,9 +177,9 @@ lean_object* l_Lean_Meta_getFVarLocalDecl(lean_object*, lean_object*, lean_objec static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_3234____closed__12; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); lean_object* lean_float_to_string(double); +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_check___lambda__1___closed__4; static lean_object* l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__5; -static lean_object* l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___closed__1; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Check_0__Lean_Meta_checkAux_checkLambdaLet___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_indentExpr(lean_object*); @@ -193,10 +198,12 @@ lean_object* lean_panic_fn(lean_object*, lean_object*); static lean_object* l_Lean_Meta_check___lambda__1___closed__6; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Meta_addPPExplicitToExposeDiff_hasExplicitDiff_x3f___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_nat_sub(lean_object*, lean_object*); +static lean_object* l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__3___closed__1; lean_object* l_Lean_Expr_getAppFn(lean_object*); LEAN_EXPORT lean_object* l_Lean_mkHashMap___at___private_Lean_Meta_Check_0__Lean_Meta_checkAux___spec__1(lean_object*); +static lean_object* l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__4___closed__1; lean_object* lean_nat_mul(lean_object*, lean_object*); -static double l_Lean_withSeconds___at_Lean_Meta_check___spec__3___closed__1; +static double l_Lean_withSeconds___at_Lean_Meta_check___spec__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_throwAppTypeMismatch___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_LocalDecl_binderInfo(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at___private_Lean_Meta_Check_0__Lean_Meta_checkAux_checkForall___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -208,12 +215,13 @@ lean_object* l_Lean_mkHashMapImp___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_mkHashMap___at___private_Lean_Meta_Check_0__Lean_Meta_checkAux___spec__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_checkApp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__12; -LEAN_EXPORT lean_object* l_Lean_withOptProfile___at_Lean_Meta_check___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Check_0__Lean_Meta_getFunctionDomain(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_Meta_mkHasTypeButIsExpectedMsg___closed__4; lean_object* lean_array_uget(lean_object*, size_t); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Meta_addPPExplicitToExposeDiff_visit___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_trace_profiler; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Check_0__Lean_Meta_checkAux_checkLambdaLet___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_st_ref_set(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -229,6 +237,7 @@ uint8_t lean_nat_dec_le(lean_object*, lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_3234____closed__4; lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Check_0__Lean_Meta_checkAux_checkForall___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_withTraceNode___at_Lean_Meta_check___spec__1___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 lean_object* l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_Check_0__Lean_Meta_checkAux_checkForall(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_addPPExplicitToExposeDiff_visit___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*); @@ -237,7 +246,6 @@ static lean_object* l_Lean_Meta_mkHasTypeButIsExpectedMsg___closed__3; LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_throwLetTypeMismatchMessage___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Data_HashMap_0__Lean_numBucketsForCapacity(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_check___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_addTraceNode___at_Lean_Meta_processPostponed___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Check___hyg_3234____closed__6; LEAN_EXPORT lean_object* l___private_Lean_Meta_Check_0__Lean_Meta_ensureType(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: @@ -3253,13 +3261,13 @@ lean_dec(x_8); 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; -x_14 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_1, x_3, x_4, x_5, x_6, x_7); +x_14 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_1, x_3, x_4, x_5, x_6, x_7); 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_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_2, x_3, x_4, x_5, x_6, x_16); +x_17 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_2, x_3, x_4, x_5, x_6, x_16); x_18 = lean_ctor_get(x_17, 0); lean_inc(x_18); x_19 = lean_ctor_get(x_17, 1); @@ -7065,7 +7073,7 @@ lean_dec(x_1); return x_2; } } -static double _init_l_Lean_withSeconds___at_Lean_Meta_check___spec__3___closed__1() { +static double _init_l_Lean_withSeconds___at_Lean_Meta_check___spec__2___closed__1() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; double x_4; @@ -7076,7 +7084,7 @@ x_4 = l_Float_ofScientific(x_1, x_2, x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_withSeconds___at_Lean_Meta_check___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_EXPORT lean_object* l_Lean_withSeconds___at_Lean_Meta_check___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; lean_object* x_8; lean_object* x_9; lean_object* x_10; @@ -7108,7 +7116,7 @@ x_17 = 0; x_18 = lean_unsigned_to_nat(0u); x_19 = l_Float_ofScientific(x_16, x_17, x_18); lean_dec(x_16); -x_20 = l_Lean_withSeconds___at_Lean_Meta_check___spec__3___closed__1; +x_20 = l_Lean_withSeconds___at_Lean_Meta_check___spec__2___closed__1; x_21 = lean_float_div(x_19, x_20); x_22 = lean_box_float(x_21); x_23 = lean_alloc_ctor(0, 2, 0); @@ -7132,7 +7140,7 @@ x_27 = 0; x_28 = lean_unsigned_to_nat(0u); x_29 = l_Float_ofScientific(x_26, x_27, x_28); lean_dec(x_26); -x_30 = l_Lean_withSeconds___at_Lean_Meta_check___spec__3___closed__1; +x_30 = l_Lean_withSeconds___at_Lean_Meta_check___spec__2___closed__1; x_31 = lean_float_div(x_29, x_30); x_32 = lean_box_float(x_31); x_33 = lean_alloc_ctor(0, 2, 0); @@ -7169,522 +7177,656 @@ return x_38; } } } -static lean_object* _init_l_Lean_withOptProfile___at_Lean_Meta_check___spec__2___closed__1() { +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_check___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_1; -x_1 = lean_mk_string_from_bytes("profiler", 8); -return x_1; -} +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_7; lean_object* x_8; +x_7 = lean_ctor_get(x_1, 0); +lean_inc(x_7); +x_8 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_8, 0, x_7); +lean_ctor_set(x_8, 1, x_6); +return x_8; } -static lean_object* _init_l_Lean_withOptProfile___at_Lean_Meta_check___spec__2___closed__2() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_withOptProfile___at_Lean_Meta_check___spec__2___closed__1; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; +lean_object* x_9; lean_object* x_10; +x_9 = lean_ctor_get(x_1, 0); +lean_inc(x_9); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_6); +return x_10; +} } } -LEAN_EXPORT lean_object* l_Lean_withOptProfile___at_Lean_Meta_check___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -lean_object* x_7; uint8_t x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_7 = l_Lean_withOptProfile___at_Lean_Meta_check___spec__2___closed__2; -x_8 = 0; -x_9 = l_Lean_getBoolOption___at_Lean_Meta_processPostponed___spec__4(x_7, x_8, x_2, x_3, x_4, x_5, x_6); -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -x_11 = lean_unbox(x_10); -lean_dec(x_10); -if (x_11 == 0) +lean_object* x_7; +x_7 = lean_apply_5(x_1, x_2, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_7) == 0) { -lean_object* x_12; lean_object* x_13; -x_12 = lean_ctor_get(x_9, 1); +uint8_t x_8; +x_8 = !lean_is_exclusive(x_7); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; +x_9 = lean_ctor_get(x_7, 0); +x_10 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_7, 0, x_10); +return x_7; +} +else +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_11 = lean_ctor_get(x_7, 0); +x_12 = lean_ctor_get(x_7, 1); lean_inc(x_12); -lean_dec(x_9); -x_13 = lean_apply_5(x_1, x_2, x_3, x_4, x_5, x_12); -if (lean_obj_tag(x_13) == 0) +lean_inc(x_11); +lean_dec(x_7); +x_13 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_13, 0, x_11); +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; +} +} +else { -uint8_t x_14; -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) +uint8_t x_15; +x_15 = !lean_is_exclusive(x_7); +if (x_15 == 0) { -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_13, 0); -x_16 = lean_box(0); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_15); -lean_ctor_set(x_17, 1, x_16); -lean_ctor_set(x_13, 0, x_17); -return x_13; +lean_object* x_16; lean_object* x_17; +x_16 = lean_ctor_get(x_7, 0); +x_17 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set_tag(x_7, 0); +lean_ctor_set(x_7, 0, x_17); +return x_7; } else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_18 = lean_ctor_get(x_13, 0); -x_19 = lean_ctor_get(x_13, 1); +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_18 = lean_ctor_get(x_7, 0); +x_19 = lean_ctor_get(x_7, 1); lean_inc(x_19); lean_inc(x_18); -lean_dec(x_13); -x_20 = lean_box(0); +lean_dec(x_7); +x_20 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_20, 0, x_18); 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(0, 2, 0); -lean_ctor_set(x_22, 0, x_21); -lean_ctor_set(x_22, 1, x_19); -return x_22; +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_19); +return x_21; } } -else +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___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) { +_start: { -uint8_t x_23; -x_23 = !lean_is_exclusive(x_13); -if (x_23 == 0) +lean_object* x_13; lean_object* x_14; lean_object* x_15; +lean_inc(x_10); +x_13 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_processPostponed___spec__4(x_1, x_2, x_3, x_6, x_4, x_8, x_9, x_10, x_11, x_12); +x_14 = lean_ctor_get(x_13, 1); +lean_inc(x_14); +lean_dec(x_13); +x_15 = l_MonadExcept_ofExcept___at_Lean_Meta_check___spec__3(x_5, x_8, x_9, x_10, x_11, x_14); +lean_dec(x_10); +return x_15; +} +} +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__3___closed__1() { +_start: { -return x_13; +lean_object* x_1; +x_1 = l_Lean_profiler; +return x_1; } -else +} +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__3___closed__2() { +_start: { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_13, 0); -x_25 = lean_ctor_get(x_13, 1); -lean_inc(x_25); -lean_inc(x_24); -lean_dec(x_13); -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; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("[", 1); +return x_1; } } +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__3___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__3___closed__2; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } -else +} +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__3___closed__4() { +_start: { -lean_object* x_27; lean_object* x_28; -x_27 = lean_ctor_get(x_9, 1); -lean_inc(x_27); +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("s] ", 3); +return x_1; +} +} +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__3___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__3___closed__4; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, uint8_t x_7, double 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; lean_object* x_16; lean_dec(x_9); -x_28 = l_Lean_withSeconds___at_Lean_Meta_check___spec__3(x_1, x_2, x_3, x_4, x_5, x_27); -if (lean_obj_tag(x_28) == 0) +x_15 = lean_ctor_get(x_12, 5); +lean_inc(x_15); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_2); +x_16 = lean_apply_6(x_1, x_2, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_16) == 0) { -uint8_t x_29; -x_29 = !lean_is_exclusive(x_28); -if (x_29 == 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); +x_19 = l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__3___closed__1; +x_20 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_6, x_19); +lean_dec(x_6); +if (x_20 == 0) { -lean_object* x_30; uint8_t x_31; -x_30 = lean_ctor_get(x_28, 0); -x_31 = !lean_is_exclusive(x_30); -if (x_31 == 0) +if (x_7 == 0) { -lean_object* x_32; lean_object* x_33; -x_32 = lean_ctor_get(x_30, 1); -x_33 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_33, 0, x_32); -lean_ctor_set(x_30, 1, x_33); -return x_28; +lean_object* x_21; lean_object* x_22; +x_21 = lean_box(0); +x_22 = l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__2(x_3, x_4, x_15, x_5, x_2, x_17, x_21, x_10, x_11, x_12, x_13, x_18); +lean_dec(x_13); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_2); +return x_22; } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_34 = lean_ctor_get(x_30, 0); -x_35 = lean_ctor_get(x_30, 1); -lean_inc(x_35); -lean_inc(x_34); -lean_dec(x_30); -x_36 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_36, 0, x_35); -x_37 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_37, 0, x_34); -lean_ctor_set(x_37, 1, x_36); -lean_ctor_set(x_28, 0, x_37); -return x_28; +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; 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_23 = lean_float_to_string(x_8); +x_24 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_24, 0, x_23); +x_25 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_25, 0, x_24); +x_26 = l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__3___closed__3; +x_27 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_25); +x_28 = l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__3___closed__5; +x_29 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +x_30 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_17); +x_31 = l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__12; +x_32 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +x_33 = lean_box(0); +x_34 = l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__2(x_3, x_4, x_15, x_5, x_2, x_32, x_33, x_10, x_11, x_12, x_13, x_18); +lean_dec(x_13); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_2); +return x_34; } } 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; -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_ctor_get(x_38, 0); -lean_inc(x_40); -x_41 = lean_ctor_get(x_38, 1); -lean_inc(x_41); -if (lean_is_exclusive(x_38)) { - lean_ctor_release(x_38, 0); - lean_ctor_release(x_38, 1); - x_42 = x_38; -} else { - lean_dec_ref(x_38); - x_42 = lean_box(0); -} -x_43 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_43, 0, x_41); -if (lean_is_scalar(x_42)) { - x_44 = lean_alloc_ctor(0, 2, 0); -} else { - x_44 = x_42; -} -lean_ctor_set(x_44, 0, x_40); +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; 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_float_to_string(x_8); +x_36 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_36, 0, x_35); +x_37 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_37, 0, x_36); +x_38 = l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__3___closed__3; +x_39 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_37); +x_40 = l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__3___closed__5; +x_41 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_40); +x_42 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_17); +x_43 = l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__12; +x_44 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_44, 0, x_42); lean_ctor_set(x_44, 1, 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_39); -return x_45; +x_45 = lean_box(0); +x_46 = l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__2(x_3, x_4, x_15, x_5, x_2, x_44, x_45, x_10, x_11, x_12, x_13, x_18); +lean_dec(x_13); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_2); +return x_46; } } else { -uint8_t x_46; -x_46 = !lean_is_exclusive(x_28); -if (x_46 == 0) +uint8_t x_47; +lean_dec(x_15); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_47 = !lean_is_exclusive(x_16); +if (x_47 == 0) { -return x_28; +return x_16; } else { -lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_47 = lean_ctor_get(x_28, 0); -x_48 = lean_ctor_get(x_28, 1); +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_16, 0); +x_49 = lean_ctor_get(x_16, 1); +lean_inc(x_49); lean_inc(x_48); -lean_inc(x_47); -lean_dec(x_28); -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_dec(x_16); +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; } } } } +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__4___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_trace_profiler; +return x_1; +} } -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_check___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_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t 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, lean_object* x_12) { _start: { -if (lean_obj_tag(x_1) == 0) +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +lean_dec(x_7); +x_13 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_processPostponed___spec__2___rarg(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_alloc_closure((void*)(l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__1), 6, 1); +lean_closure_set(x_16, 0, x_1); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_17 = l_Lean_withSeconds___at_Lean_Meta_check___spec__2(x_16, x_8, x_9, x_10, x_11, x_15); +if (lean_obj_tag(x_17) == 0) { -lean_object* x_7; lean_object* x_8; -x_7 = lean_ctor_get(x_1, 0); -lean_inc(x_7); -x_8 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_8, 0, x_7); -lean_ctor_set(x_8, 1, x_6); -return x_8; +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; lean_object* x_67; uint8_t x_68; +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_ctor_get(x_18, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_18, 1); +lean_inc(x_21); +lean_dec(x_18); +x_67 = l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__4___closed__1; +x_68 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_5, x_67); +if (x_68 == 0) +{ +uint8_t x_69; +x_69 = 0; +x_22 = x_69; +goto block_66; } else { -lean_object* x_9; lean_object* x_10; -x_9 = lean_ctor_get(x_1, 0); -lean_inc(x_9); -x_10 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_10, 0, x_9); -lean_ctor_set(x_10, 1, x_6); -return x_10; +double x_70; double x_71; uint8_t x_72; +x_70 = l_Lean_trace_profiler_threshold_getSecs(x_5); +x_71 = lean_unbox_float(x_21); +x_72 = lean_float_decLt(x_70, x_71); +x_22 = x_72; +goto block_66; +} +block_66: +{ +if (x_6 == 0) +{ +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; +lean_dec(x_21); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +x_23 = lean_st_ref_take(x_11, 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_is_exclusive(x_24); +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_24, 3); +x_28 = l_Lean_PersistentArray_append___rarg(x_14, x_27); +lean_ctor_set(x_24, 3, x_28); +x_29 = lean_st_ref_set(x_11, x_24, x_25); +x_30 = lean_ctor_get(x_29, 1); +lean_inc(x_30); +lean_dec(x_29); +x_31 = l_MonadExcept_ofExcept___at_Lean_Meta_check___spec__3(x_20, x_8, x_9, x_10, x_11, x_30); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_20); +if (lean_obj_tag(x_31) == 0) +{ +uint8_t x_32; +x_32 = !lean_is_exclusive(x_31); +if (x_32 == 0) +{ +return x_31; } +else +{ +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(0, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +return x_35; } } -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: +else { -lean_object* x_7; -x_7 = lean_apply_5(x_1, x_2, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_7) == 0) +uint8_t x_36; +x_36 = !lean_is_exclusive(x_31); +if (x_36 == 0) { -uint8_t x_8; -x_8 = !lean_is_exclusive(x_7); -if (x_8 == 0) +return x_31; +} +else { -lean_object* x_9; lean_object* x_10; -x_9 = lean_ctor_get(x_7, 0); -x_10 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_10, 0, x_9); -lean_ctor_set(x_7, 0, x_10); -return x_7; +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_31, 0); +x_38 = lean_ctor_get(x_31, 1); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_31); +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_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_11 = lean_ctor_get(x_7, 0); -x_12 = lean_ctor_get(x_7, 1); -lean_inc(x_12); -lean_inc(x_11); -lean_dec(x_7); -x_13 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_13, 0, x_11); -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_40; lean_object* x_41; 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_40 = lean_ctor_get(x_24, 0); +x_41 = lean_ctor_get(x_24, 1); +x_42 = lean_ctor_get(x_24, 2); +x_43 = lean_ctor_get(x_24, 3); +x_44 = lean_ctor_get(x_24, 4); +x_45 = lean_ctor_get(x_24, 5); +x_46 = lean_ctor_get(x_24, 6); +lean_inc(x_46); +lean_inc(x_45); +lean_inc(x_44); +lean_inc(x_43); +lean_inc(x_42); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_24); +x_47 = l_Lean_PersistentArray_append___rarg(x_14, x_43); +x_48 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_48, 0, x_40); +lean_ctor_set(x_48, 1, x_41); +lean_ctor_set(x_48, 2, x_42); +lean_ctor_set(x_48, 3, x_47); +lean_ctor_set(x_48, 4, x_44); +lean_ctor_set(x_48, 5, x_45); +lean_ctor_set(x_48, 6, x_46); +x_49 = lean_st_ref_set(x_11, x_48, x_25); +x_50 = lean_ctor_get(x_49, 1); +lean_inc(x_50); +lean_dec(x_49); +x_51 = l_MonadExcept_ofExcept___at_Lean_Meta_check___spec__3(x_20, x_8, x_9, x_10, x_11, x_50); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_20); +if (lean_obj_tag(x_51) == 0) +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_52 = lean_ctor_get(x_51, 0); +lean_inc(x_52); +x_53 = lean_ctor_get(x_51, 1); +lean_inc(x_53); +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_51); + x_54 = lean_box(0); +} +if (lean_is_scalar(x_54)) { + x_55 = lean_alloc_ctor(0, 2, 0); +} else { + x_55 = x_54; } +lean_ctor_set(x_55, 0, x_52); +lean_ctor_set(x_55, 1, x_53); +return x_55; } else { -uint8_t x_15; -x_15 = !lean_is_exclusive(x_7); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; -x_16 = lean_ctor_get(x_7, 0); -x_17 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_17, 0, x_16); -lean_ctor_set_tag(x_7, 0); -lean_ctor_set(x_7, 0, x_17); -return x_7; +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_56 = lean_ctor_get(x_51, 0); +lean_inc(x_56); +x_57 = lean_ctor_get(x_51, 1); +lean_inc(x_57); +if (lean_is_exclusive(x_51)) { + lean_ctor_release(x_51, 0); + lean_ctor_release(x_51, 1); + x_58 = x_51; +} else { + lean_dec_ref(x_51); + x_58 = lean_box(0); } -else -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_18 = lean_ctor_get(x_7, 0); -x_19 = lean_ctor_get(x_7, 1); -lean_inc(x_19); -lean_inc(x_18); -lean_dec(x_7); -x_20 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_20, 0, 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_19); -return x_21; +if (lean_is_scalar(x_58)) { + x_59 = lean_alloc_ctor(1, 2, 0); +} else { + x_59 = x_58; } +lean_ctor_set(x_59, 0, x_56); +lean_ctor_set(x_59, 1, x_57); +return x_59; } } } -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___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) { -_start: +else { -lean_object* x_13; lean_object* x_14; lean_object* x_15; -lean_inc(x_10); -x_13 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_processPostponed___spec__6(x_1, x_2, x_3, x_6, x_4, x_8, x_9, x_10, x_11, x_12); -x_14 = lean_ctor_get(x_13, 1); -lean_inc(x_14); -lean_dec(x_13); -x_15 = l_MonadExcept_ofExcept___at_Lean_Meta_check___spec__4(x_5, x_8, x_9, x_10, x_11, x_14); -lean_dec(x_10); -return x_15; +lean_object* x_60; double x_61; lean_object* x_62; +x_60 = lean_box(0); +x_61 = lean_unbox_float(x_21); +lean_dec(x_21); +x_62 = l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__3(x_2, x_20, x_14, x_3, x_4, x_5, x_22, x_61, x_60, x_8, x_9, x_10, x_11, x_19); +return x_62; } } -static lean_object* _init_l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("[", 1); -return x_1; +lean_object* x_63; double x_64; lean_object* x_65; +x_63 = lean_box(0); +x_64 = lean_unbox_float(x_21); +lean_dec(x_21); +x_65 = l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__3(x_2, x_20, x_14, x_3, x_4, x_5, x_22, x_64, x_63, x_8, x_9, x_10, x_11, x_19); +return x_65; } } -static lean_object* _init_l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___closed__1; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; } +else +{ +uint8_t x_73; +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +x_73 = !lean_is_exclusive(x_17); +if (x_73 == 0) +{ +return x_17; } -static lean_object* _init_l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___closed__3() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("s] ", 3); -return x_1; +lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_74 = lean_ctor_get(x_17, 0); +x_75 = lean_ctor_get(x_17, 1); +lean_inc(x_75); +lean_inc(x_74); +lean_dec(x_17); +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; } } -static lean_object* _init_l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___closed__3; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; } } LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_check___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) { _start: { -lean_object* x_10; lean_object* x_11; uint8_t x_12; +lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_10 = lean_ctor_get(x_7, 2); +lean_inc(x_10); lean_inc(x_1); -x_10 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_1, 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_unbox(x_11); +x_11 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___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_unbox(x_12); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); lean_dec(x_11); -if (x_12 == 0) +x_15 = l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__4___closed__1; +x_16 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_10, x_15); +if (x_16 == 0) { -lean_object* x_13; lean_object* x_14; +lean_object* x_17; +lean_dec(x_12); +lean_dec(x_10); lean_dec(x_2); lean_dec(x_1); -x_13 = lean_ctor_get(x_10, 1); -lean_inc(x_13); -lean_dec(x_10); -x_14 = lean_apply_5(x_3, x_5, x_6, x_7, x_8, x_13); -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; -x_15 = lean_ctor_get(x_10, 1); -lean_inc(x_15); -lean_dec(x_10); -x_16 = lean_ctor_get(x_7, 5); -lean_inc(x_16); -x_17 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_processPostponed___spec__2___rarg(x_8, 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_alloc_closure((void*)(l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__1), 6, 1); -lean_closure_set(x_20, 0, x_3); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_21 = l_Lean_withOptProfile___at_Lean_Meta_check___spec__2(x_20, x_5, x_6, x_7, x_8, 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; -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_22, 0); -lean_inc(x_24); -x_25 = lean_ctor_get(x_22, 1); -lean_inc(x_25); -lean_dec(x_22); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_24); -x_26 = lean_apply_6(x_2, x_24, x_5, x_6, x_7, x_8, x_23); -if (lean_obj_tag(x_26) == 0) +x_17 = lean_apply_5(x_3, x_5, x_6, x_7, x_8, x_14); +if (lean_obj_tag(x_17) == 0) { -if (lean_obj_tag(x_25) == 0) +uint8_t x_18; +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; -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_box(0); -x_30 = l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__2(x_18, x_1, x_16, x_4, x_24, x_27, x_29, x_5, x_6, x_7, x_8, x_28); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_24); -return x_30; +return x_17; } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; double x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; 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_31 = lean_ctor_get(x_26, 0); -lean_inc(x_31); -x_32 = lean_ctor_get(x_26, 1); -lean_inc(x_32); -lean_dec(x_26); -x_33 = lean_ctor_get(x_25, 0); -lean_inc(x_33); -lean_dec(x_25); -x_34 = lean_unbox_float(x_33); -lean_dec(x_33); -x_35 = lean_float_to_string(x_34); -x_36 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_36, 0, x_35); -x_37 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_37, 0, x_36); -x_38 = l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___closed__2; -x_39 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_39, 1, x_37); -x_40 = l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___closed__4; -x_41 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_41, 0, x_39); -lean_ctor_set(x_41, 1, x_40); -x_42 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_31); -x_43 = l_Lean_Meta_throwLetTypeMismatchMessage___rarg___closed__12; -x_44 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_44, 0, x_42); -lean_ctor_set(x_44, 1, x_43); -x_45 = lean_box(0); -x_46 = l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__2(x_18, x_1, x_16, x_4, x_24, x_44, x_45, x_5, x_6, x_7, x_8, x_32); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_24); -return x_46; +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_17, 0); +x_20 = lean_ctor_get(x_17, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_17); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); +return x_21; } } else { -uint8_t x_47; -lean_dec(x_25); -lean_dec(x_24); -lean_dec(x_18); -lean_dec(x_16); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_1); -x_47 = !lean_is_exclusive(x_26); -if (x_47 == 0) +uint8_t x_22; +x_22 = !lean_is_exclusive(x_17); +if (x_22 == 0) { -return x_26; +return x_17; } else { -lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_48 = lean_ctor_get(x_26, 0); -x_49 = lean_ctor_get(x_26, 1); -lean_inc(x_49); -lean_inc(x_48); -lean_dec(x_26); -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; +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_17, 0); +x_24 = lean_ctor_get(x_17, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_17); +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 { -uint8_t x_51; -lean_dec(x_18); -lean_dec(x_16); -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_51 = !lean_is_exclusive(x_21); -if (x_51 == 0) -{ -return x_21; +lean_object* x_26; uint8_t x_27; lean_object* x_28; +x_26 = lean_box(0); +x_27 = lean_unbox(x_12); +lean_dec(x_12); +x_28 = l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__4(x_3, x_2, x_1, x_4, x_10, x_27, x_26, x_5, x_6, x_7, x_8, x_14); +return x_28; +} } else { -lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_52 = lean_ctor_get(x_21, 0); -x_53 = lean_ctor_get(x_21, 1); -lean_inc(x_53); -lean_inc(x_52); -lean_dec(x_21); -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_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; +x_29 = lean_ctor_get(x_11, 1); +lean_inc(x_29); +lean_dec(x_11); +x_30 = lean_box(0); +x_31 = lean_unbox(x_12); +lean_dec(x_12); +x_32 = l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__4(x_3, x_2, x_1, x_4, x_10, x_31, x_30, x_5, x_6, x_7, x_8, x_29); +return x_32; } } } @@ -8153,11 +8295,11 @@ x_11 = l_Lean_withTraceNode___at_Lean_Meta_check___spec__1(x_8, x_7, x_9, x_10, return x_11; } } -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_check___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_MonadExcept_ofExcept___at_Lean_Meta_check___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_MonadExcept_ofExcept___at_Lean_Meta_check___spec__4(x_1, x_2, x_3, x_4, x_5, x_6); +x_7 = l_MonadExcept_ofExcept___at_Lean_Meta_check___spec__3(x_1, x_2, x_3, x_4, x_5, x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); @@ -8181,6 +8323,32 @@ lean_dec(x_5); return x_14; } } +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___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) { +_start: +{ +uint8_t x_15; uint8_t x_16; double x_17; lean_object* x_18; +x_15 = lean_unbox(x_5); +lean_dec(x_5); +x_16 = lean_unbox(x_7); +lean_dec(x_7); +x_17 = lean_unbox_float(x_8); +lean_dec(x_8); +x_18 = l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__3(x_1, x_2, x_3, x_4, x_15, x_6, x_16, x_17, x_9, x_10, x_11, x_12, x_13, x_14); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___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) { +_start: +{ +uint8_t x_13; uint8_t x_14; lean_object* x_15; +x_13 = lean_unbox(x_4); +lean_dec(x_4); +x_14 = lean_unbox(x_6); +lean_dec(x_6); +x_15 = l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__4(x_1, x_2, x_3, x_13, x_5, x_14, x_7, x_8, x_9, x_10, x_11, x_12); +return x_15; +} +} LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_check___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) { _start: { @@ -8499,19 +8667,19 @@ l___private_Lean_Meta_Check_0__Lean_Meta_checkAux_checkLambdaLet___closed__1 = _ lean_mark_persistent(l___private_Lean_Meta_Check_0__Lean_Meta_checkAux_checkLambdaLet___closed__1); l___private_Lean_Meta_Check_0__Lean_Meta_checkAux_checkForall___closed__1 = _init_l___private_Lean_Meta_Check_0__Lean_Meta_checkAux_checkForall___closed__1(); lean_mark_persistent(l___private_Lean_Meta_Check_0__Lean_Meta_checkAux_checkForall___closed__1); -l_Lean_withSeconds___at_Lean_Meta_check___spec__3___closed__1 = _init_l_Lean_withSeconds___at_Lean_Meta_check___spec__3___closed__1(); -l_Lean_withOptProfile___at_Lean_Meta_check___spec__2___closed__1 = _init_l_Lean_withOptProfile___at_Lean_Meta_check___spec__2___closed__1(); -lean_mark_persistent(l_Lean_withOptProfile___at_Lean_Meta_check___spec__2___closed__1); -l_Lean_withOptProfile___at_Lean_Meta_check___spec__2___closed__2 = _init_l_Lean_withOptProfile___at_Lean_Meta_check___spec__2___closed__2(); -lean_mark_persistent(l_Lean_withOptProfile___at_Lean_Meta_check___spec__2___closed__2); -l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___closed__1 = _init_l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___closed__1(); -lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___closed__1); -l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___closed__2 = _init_l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___closed__2(); -lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___closed__2); -l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___closed__3 = _init_l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___closed__3(); -lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___closed__3); -l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___closed__4 = _init_l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___closed__4(); -lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___closed__4); +l_Lean_withSeconds___at_Lean_Meta_check___spec__2___closed__1 = _init_l_Lean_withSeconds___at_Lean_Meta_check___spec__2___closed__1(); +l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__3___closed__1 = _init_l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__3___closed__1); +l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__3___closed__2 = _init_l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__3___closed__2); +l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__3___closed__3 = _init_l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__3___closed__3(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__3___closed__3); +l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__3___closed__4 = _init_l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__3___closed__4(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__3___closed__4); +l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__3___closed__5 = _init_l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__3___closed__5(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__3___closed__5); +l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__4___closed__1 = _init_l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Meta_check___spec__1___lambda__4___closed__1); l_Lean_Meta_check___lambda__1___closed__1 = _init_l_Lean_Meta_check___lambda__1___closed__1(); lean_mark_persistent(l_Lean_Meta_check___lambda__1___closed__1); l_Lean_Meta_check___lambda__1___closed__2 = _init_l_Lean_Meta_check___lambda__1___closed__2(); diff --git a/stage0/stdlib/Lean/Meta/Closure.c b/stage0/stdlib/Lean/Meta/Closure.c index 29269b899465..ce24a9516d97 100644 --- a/stage0/stdlib/Lean/Meta/Closure.c +++ b/stage0/stdlib/Lean/Meta/Closure.c @@ -3817,7 +3817,7 @@ static lean_object* _init_l_Lean_Meta_Closure_collectExprAux___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_Meta_Closure_collectExprAux___closed__1; x_2 = l_Lean_Meta_Closure_collectExprAux___closed__2; -x_3 = lean_unsigned_to_nat(1503u); +x_3 = lean_unsigned_to_nat(1507u); x_4 = lean_unsigned_to_nat(17u); x_5 = l_Lean_Meta_Closure_collectExprAux___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -3846,7 +3846,7 @@ static lean_object* _init_l_Lean_Meta_Closure_collectExprAux___closed__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; x_1 = l_Lean_Meta_Closure_collectExprAux___closed__1; x_2 = l_Lean_Meta_Closure_collectExprAux___closed__5; -x_3 = lean_unsigned_to_nat(1514u); +x_3 = lean_unsigned_to_nat(1518u); x_4 = lean_unsigned_to_nat(18u); x_5 = l_Lean_Meta_Closure_collectExprAux___closed__6; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -3875,7 +3875,7 @@ static lean_object* _init_l_Lean_Meta_Closure_collectExprAux___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_Lean_Meta_Closure_collectExprAux___closed__1; x_2 = l_Lean_Meta_Closure_collectExprAux___closed__8; -x_3 = lean_unsigned_to_nat(1465u); +x_3 = lean_unsigned_to_nat(1469u); x_4 = lean_unsigned_to_nat(18u); x_5 = l_Lean_Meta_Closure_collectExprAux___closed__9; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -3904,7 +3904,7 @@ static lean_object* _init_l_Lean_Meta_Closure_collectExprAux___closed__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; x_1 = l_Lean_Meta_Closure_collectExprAux___closed__1; x_2 = l_Lean_Meta_Closure_collectExprAux___closed__11; -x_3 = lean_unsigned_to_nat(1560u); +x_3 = lean_unsigned_to_nat(1564u); x_4 = lean_unsigned_to_nat(20u); x_5 = l_Lean_Meta_Closure_collectExprAux___closed__12; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -3933,7 +3933,7 @@ static lean_object* _init_l_Lean_Meta_Closure_collectExprAux___closed__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; x_1 = l_Lean_Meta_Closure_collectExprAux___closed__1; x_2 = l_Lean_Meta_Closure_collectExprAux___closed__14; -x_3 = lean_unsigned_to_nat(1540u); +x_3 = lean_unsigned_to_nat(1544u); x_4 = lean_unsigned_to_nat(24u); x_5 = l_Lean_Meta_Closure_collectExprAux___closed__15; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -3962,7 +3962,7 @@ static lean_object* _init_l_Lean_Meta_Closure_collectExprAux___closed__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; x_1 = l_Lean_Meta_Closure_collectExprAux___closed__1; x_2 = l_Lean_Meta_Closure_collectExprAux___closed__17; -x_3 = lean_unsigned_to_nat(1569u); +x_3 = lean_unsigned_to_nat(1573u); x_4 = lean_unsigned_to_nat(22u); x_5 = l_Lean_Meta_Closure_collectExprAux___closed__18; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); diff --git a/stage0/stdlib/Lean/Meta/Coe.c b/stage0/stdlib/Lean/Meta/Coe.c index 13765139db54..e484d19d428a 100644 --- a/stage0/stdlib/Lean/Meta/Coe.c +++ b/stage0/stdlib/Lean/Meta/Coe.c @@ -14,7 +14,6 @@ extern "C" { #endif lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); -uint8_t l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Meta_expandCoe___lambda__1___closed__1; lean_object* l_Lean_Meta_mkAppOptM(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Coe___hyg_4____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -74,7 +73,6 @@ lean_object* l_Lean_Expr_forallE___override(lean_object*, lean_object*, lean_obj static lean_object* l_Lean_Meta_coerceMonadLift_x3f___closed__3; lean_object* lean_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_coerceToFunction_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Coe___hyg_291_(lean_object*); static lean_object* l_Lean_Meta_coerceMonadLift_x3f___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_coerceToSort_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -82,6 +80,7 @@ lean_object* l_Lean_Expr_constName_x21(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_coerceToFunction_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_coerceMonadLift_x3f___closed__16; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Meta_coerceToFunction_x3f___closed__5; static lean_object* l_Lean_Meta_expandCoe___closed__1; static lean_object* l_Lean_Meta_coerceSimple_x3f___closed__10; @@ -91,7 +90,6 @@ static lean_object* l_Lean_Meta_coerceMonadLift_x3f___closed__14; lean_object* l_Lean_registerTagAttribute(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); static lean_object* l_Lean_Meta_coerceSimple_x3f___closed__8; LEAN_EXPORT lean_object* l_Lean_Meta_coeDeclAttr; -lean_object* l_Lean_Option_register___at_Std_Format_initFn____x40_Lean_Data_Format___hyg_95____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_coerceToSort_x3f___closed__4; static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Coe___hyg_291____closed__6; static lean_object* l_Lean_Meta_coerceToSort_x3f___closed__2; @@ -105,6 +103,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_expandCoe___lambda__1___boxed(lean_object*, static lean_object* l_Lean_Meta_coerceToFunction_x3f___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_isTypeApp_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_decLevel(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_coerceSimple_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_coerceMonadLift_x3f___closed__9; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); @@ -118,6 +117,7 @@ static lean_object* l_Lean_Meta_coerceMonadLift_x3f___closed__5; LEAN_EXPORT lean_object* l_Lean_Meta_autoLift; LEAN_EXPORT lean_object* l_Lean_Meta_expandCoe___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_coerceToFunction_x3f___closed__8; +lean_object* l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Coe___hyg_291____closed__2; lean_object* l_Lean_Expr_getAppFn(lean_object*); lean_object* l_Lean_mkArrow(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -828,7 +828,7 @@ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Coe___hyg_291____closed__2; x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_Coe___hyg_291____closed__5; x_4 = l_Lean_Meta_initFn____x40_Lean_Meta_Coe___hyg_291____closed__6; -x_5 = l_Lean_Option_register___at_Std_Format_initFn____x40_Lean_Data_Format___hyg_95____spec__1(x_2, x_3, x_4, x_1); +x_5 = l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__1(x_2, x_3, x_4, x_1); return x_5; } } @@ -2499,13 +2499,13 @@ lean_inc(x_19); x_20 = lean_ctor_get(x_17, 1); lean_inc(x_20); lean_dec(x_17); -x_21 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_19, x_2, x_3, x_4, x_5, x_18); +x_21 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_19, x_2, x_3, x_4, x_5, x_18); 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_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_20, x_2, x_3, x_4, x_5, x_23); +x_24 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_20, x_2, x_3, x_4, x_5, x_23); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); @@ -2659,13 +2659,13 @@ lean_inc(x_62); x_63 = lean_ctor_get(x_60, 1); lean_inc(x_63); lean_dec(x_60); -x_64 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_62, x_2, x_3, x_4, x_5, x_61); +x_64 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_62, x_2, x_3, x_4, x_5, x_61); 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_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_63, x_2, x_3, x_4, x_5, x_66); +x_67 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_63, x_2, x_3, x_4, x_5, x_66); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); @@ -3061,7 +3061,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_coerceMonadLift_x3f(lean_object* x_1, lean_ _start: { lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_8 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_2, x_3, x_4, x_5, x_6, x_7); +x_8 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_2, x_3, x_4, x_5, x_6, x_7); x_9 = lean_ctor_get(x_8, 0); lean_inc(x_9); x_10 = lean_ctor_get(x_8, 1); @@ -3081,7 +3081,7 @@ lean_inc(x_12); x_13 = lean_ctor_get(x_11, 1); lean_inc(x_13); lean_dec(x_11); -x_14 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_12, x_3, x_4, x_5, x_6, x_13); +x_14 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_12, x_3, x_4, x_5, x_6, x_13); x_15 = lean_ctor_get(x_14, 0); lean_inc(x_15); x_16 = lean_ctor_get(x_14, 1); @@ -3238,7 +3238,7 @@ lean_dec(x_48); x_49 = lean_ctor_get(x_5, 2); lean_inc(x_49); x_50 = l_Lean_Meta_coerceMonadLift_x3f___closed__1; -x_51 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_49, x_50); +x_51 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_49, x_50); lean_dec(x_49); if (x_51 == 0) { @@ -5043,7 +5043,7 @@ lean_dec(x_43); x_353 = lean_ctor_get(x_5, 2); lean_inc(x_353); x_354 = l_Lean_Meta_coerceMonadLift_x3f___closed__1; -x_355 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_353, x_354); +x_355 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_353, x_354); lean_dec(x_353); if (x_355 == 0) { @@ -7136,7 +7136,7 @@ if (lean_is_exclusive(x_686)) { x_691 = lean_ctor_get(x_5, 2); lean_inc(x_691); x_692 = l_Lean_Meta_coerceMonadLift_x3f___closed__1; -x_693 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_691, x_692); +x_693 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_691, x_692); lean_dec(x_691); if (x_693 == 0) { @@ -9190,7 +9190,7 @@ if (lean_is_exclusive(x_1008)) { x_1013 = lean_ctor_get(x_5, 2); lean_inc(x_1013); x_1014 = l_Lean_Meta_coerceMonadLift_x3f___closed__1; -x_1015 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_1013, x_1014); +x_1015 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_1013, x_1014); lean_dec(x_1013); if (x_1015 == 0) { diff --git a/stage0/stdlib/Lean/Meta/CongrTheorems.c b/stage0/stdlib/Lean/Meta/CongrTheorems.c index 7a92b9755b6a..7dbbf63f102f 100644 --- a/stage0/stdlib/Lean/Meta/CongrTheorems.c +++ b/stage0/stdlib/Lean/Meta/CongrTheorems.c @@ -102,7 +102,6 @@ static lean_object* l___private_Lean_Meta_CongrTheorems_0__Lean_Meta_reprCongrAr lean_object* l_Lean_Expr_appArg_x21(lean_object*); static lean_object* l_Lean_Meta_mkHCongrWithArity___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_mkCongrSimpCore_x3f_mk_x3f_go___lambda__1(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*); -lean_object* l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_hasMVar(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mkCongrSimpCore_x3f_mkProof_go___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* l_Lean_LocalContext_getFVar_x21(lean_object*, lean_object*); @@ -145,7 +144,6 @@ static lean_object* l___private_Lean_Meta_CongrTheorems_0__Lean_Meta_reprCongrAr lean_object* l_Lean_Expr_replaceFVar(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkHEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at___private_Lean_Meta_CongrTheorems_0__Lean_Meta_mkCast_go___spec__10___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mkCongrSimpCore_x3f_mk_x3f_go(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___private_Lean_Meta_CongrTheorems_0__Lean_Meta_reprCongrArgKind____x40_Lean_Meta_CongrTheorems___hyg_23____closed__32; @@ -197,6 +195,7 @@ lean_object* l_Lean_LocalDecl_userName(lean_object*); lean_object* lean_name_append_index_after(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_CongrTheorems_0__Lean_Meta_reprCongrArgKind____x40_Lean_Meta_CongrTheorems___hyg_23____closed__24; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at___private_Lean_Meta_CongrTheorems_0__Lean_Meta_mkCast_go___spec__4___boxed(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mkCongrSimpCore_x3f_mkProof_go___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_Meta_mkHCongrWithArity___lambda__2___closed__2; lean_object* l_Lean_Expr_bindingDomain_x21(lean_object*); @@ -216,6 +215,7 @@ LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Meta_getCongrSimpKinds LEAN_EXPORT lean_object* l_Lean_Meta_mkHCongrWithArity___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 lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_CongrTheorems_0__Lean_Meta_mkCast_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_CongrTheorems_0__Lean_Meta_mkCast_go___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_CongrTheorems_0__Lean_Meta_getClassSubobjectMask_x3f___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_mkCongrSimpCore_x3f_mkProof_go___closed__1; @@ -2833,7 +2833,7 @@ x_23 = lean_alloc_closure((void*)(l_Lean_Meta_mkHCongrWithArity_withNewEqs___rar lean_closure_set(x_23, 0, x_1); lean_closure_set(x_23, 1, x_4); lean_closure_set(x_23, 2, x_22); -x_24 = l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__3___rarg(x_18, x_20, x_23, x_6, x_7, x_8, x_9, x_21); +x_24 = l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2___rarg(x_18, x_20, x_23, x_6, x_7, x_8, x_9, x_21); return x_24; } } @@ -9465,7 +9465,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_mkCongrSimp_x3f(lean_object* x_1, uint8_t x _start: { lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_8 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_1, x_3, x_4, x_5, x_6, x_7); +x_8 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_1, x_3, x_4, x_5, x_6, x_7); x_9 = lean_ctor_get(x_8, 0); lean_inc(x_9); x_10 = lean_ctor_get(x_8, 1); diff --git a/stage0/stdlib/Lean/Meta/DiscrTree.c b/stage0/stdlib/Lean/Meta/DiscrTree.c index 5725e8714ee6..f55db6d89902 100644 --- a/stage0/stdlib/Lean/Meta/DiscrTree.c +++ b/stage0/stdlib/Lean/Meta/DiscrTree.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Meta.DiscrTree -// Imports: Init Lean.Meta.WHNF Lean.Meta.DiscrTreeTypes +// Imports: Init Lean.Meta.WHNF Lean.Meta.Transform Lean.Meta.DiscrTreeTypes #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -118,6 +118,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_DiscrTree_Trie_format___rarg(uint8_t, lean_ LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___rarg(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_DiscrTree_insertCore___spec__7___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_DiscrTree_instInhabitedTrie___rarg___boxed(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getKeyArgs___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_PersistentHashMap_findAtAux___at_Lean_Meta_DiscrTree_getMatchWithExtra_mayMatchPrefix___spec__6___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_binSearchAux___at_Lean_Meta_DiscrTree_getUnify_process___spec__3___rarg(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Std_Format_join(lean_object*); @@ -157,6 +158,7 @@ lean_object* l_Lean_Expr_mvar___override(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_DiscrTree_getMatchWithExtra_mayMatchPrefix___spec__8___rarg(uint8_t, lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_DiscrTree_insertCore___spec__8(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_DiscrTree_instToFormatDiscrTree(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_elimLooseBVarsByBeta___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_DiscrTree_format___rarg(uint8_t, lean_object*, lean_object*); static lean_object* l_Lean_Meta_DiscrTree_mkNoindexAnnotation___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_DiscrTree_getUnify___spec__7___rarg___boxed(lean_object*, lean_object*, lean_object*); @@ -201,16 +203,17 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_DiscrTree_getMatc LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_DiscrTree_insertCore___spec__6___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_DiscrTree_format(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_DiscrTree_Key_arity___boxed(lean_object*); +static lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_elimLooseBVarsByBeta___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_DiscrTree_instToFormatTrie___rarg(uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_DiscrTree_getUnify___spec__19(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getKeyArgs___lambda__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_DiscrTree_insertCore___spec__12(lean_object*); LEAN_EXPORT lean_object* l_Array_binSearchAux___at___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getMatchLoop___spec__1___rarg___boxed(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_elimLooseBVarsByBeta(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_DiscrTree_insertCore___spec__21(lean_object*); lean_object* l_Lean_Expr_getRevArg_x21(lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_DiscrTree_Trie_format___spec__1___rarg(uint8_t, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getKeyArgs___closed__2; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_Meta_DiscrTree_getUnify___spec__11___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_DiscrTree_insertCore___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_DiscrTree_getUnify___spec__1___rarg(uint8_t, lean_object*, lean_object*); @@ -264,6 +267,7 @@ lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_MetavarContext_getExpr LEAN_EXPORT uint8_t l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_isNumeral(lean_object*); static lean_object* l_Lean_Meta_DiscrTree_Trie_format___rarg___closed__6; lean_object* lean_array_pop(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getKeyArgs___lambda__2(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_DiscrTree_insertCore___spec__15___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_DiscrTree_insertCore___spec__7(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Meta_DiscrTree_getUnify___spec__10___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -274,6 +278,7 @@ lean_object* lean_array_to_list(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_DiscrTree_getUnify___spec__18___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_binSearchAux___at_Lean_Meta_DiscrTree_getUnify_process___spec__5___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_hasLooseBVars(lean_object*); +uint8_t l_Lean_Expr_isHeadBetaTarget(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Array_binSearchAux___at___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getMatchLoop___spec__4(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getStarResult___spec__3(lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_DiscrTree_Key_lt___rarg(lean_object*, lean_object*); @@ -287,6 +292,7 @@ lean_object* l_Lean_Expr_constName_x21(lean_object*); static lean_object* l_Lean_Meta_DiscrTree_instInhabitedTrie___rarg___closed__1; extern lean_object* l_Lean_instInhabitedExpr; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_DiscrTree_insertCore___spec__4___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getKeyArgs___lambda__2___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_findKey(uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Meta_DiscrTree_format___spec__1(lean_object*, uint8_t); lean_object* l_Lean_mkAnnotation(lean_object*, lean_object*); @@ -294,6 +300,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_DiscrTree_insert___rarg___boxed(lean_object LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_DiscrTree_getUnify___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_DiscrTree_instDecidableLtKeyInstLTKey(uint8_t); static lean_object* l_Lean_Meta_DiscrTree_empty___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_elimLooseBVarsByBeta___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_DiscrTree_reduce(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); @@ -337,6 +344,8 @@ LEAN_EXPORT lean_object* l_Array_binSearchAux___at___private_Lean_Meta_DiscrTree LEAN_EXPORT lean_object* l_Lean_Meta_DiscrTree_instToFormatTrie___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isRec___at___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getKeyArgs___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_binSearchAux___at___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_findKey___spec__1(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_elimLooseBVarsByBeta___closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_pushArgs___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_Meta_DiscrTree_getMatchWithExtra_go___rarg(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_PersistentHashMap_findAux___at_Lean_Meta_DiscrTree_insertCore___spec__13(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_DiscrTree_insertCore___spec__19___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -369,7 +378,9 @@ static lean_object* l_Lean_Meta_DiscrTree_insertCore___rarg___closed__3; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_DiscrTree_getUnify___spec__17___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getKeyArgs___boxed(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_getAppRevArgsAux(lean_object*, lean_object*); +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_EXPORT lean_object* l_Lean_Meta_DiscrTree_getUnify___rarg___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___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_elimLooseBVarsByBeta___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_DiscrTree_Key_ctorIdx(uint8_t); static lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_tmpMVarId___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getUnifyKeyArgs___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -388,6 +399,7 @@ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_DiscrTre lean_object* lean_string_length(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_ignoreArg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_DiscrTree_insertCore___spec__14___rarg(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_pushArgs___lambda__2(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*); uint8_t l_Array_contains___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_DiscrTree_instInhabitedTrie(lean_object*); @@ -416,6 +428,7 @@ LEAN_EXPORT lean_object* l_Array_binSearchAux___at_Lean_Meta_DiscrTree_getUnify_ LEAN_EXPORT lean_object* l_Lean_Meta_DiscrTree_getUnify_process(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_DiscrTree_insertCore___rarg(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_DiscrTree_getMatchWithExtra_go(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_elimLooseBVarsByBeta___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_tmpStar; lean_object* lean_panic_fn(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_DiscrTree_instToFormatKey(uint8_t); @@ -465,6 +478,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_DiscrTree_insertCore___spec__5___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_DiscrTree_Trie_format___rarg___closed__1; size_t lean_usize_sub(size_t, size_t); +LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_elimLooseBVarsByBeta___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Meta_DiscrTree_format___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_DiscrTree_insertCore___spec__6(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_DiscrTree_insertCore___spec__3___rarg(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -505,6 +519,7 @@ static lean_object* l_Lean_Meta_DiscrTree_Trie_format___rarg___closed__4; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_DiscrTree_getUnify___spec__20___rarg(uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_DiscrTree_getUnify_process___spec__4(lean_object*); uint64_t l_Lean_Meta_DiscrTree_Key_hash___rarg(lean_object*); +lean_object* l_Lean_Expr_headBeta(lean_object*); lean_object* lean_array_get_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_DiscrTree_Trie_format___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_DiscrTree_insertCore___spec__13___rarg(uint8_t, lean_object*, size_t, lean_object*); @@ -3494,6 +3509,118 @@ x_3 = lean_box(x_2); return x_3; } } +static lean_object* _init_l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_elimLooseBVarsByBeta___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +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_DiscrTree_0__Lean_Meta_DiscrTree_elimLooseBVarsByBeta___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +uint8_t x_5; +x_5 = l_Lean_Expr_hasLooseBVars(x_1); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; +x_6 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_6, 0, x_1); +x_7 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_4); +return x_7; +} +else +{ +uint8_t x_8; uint8_t x_9; +x_8 = 0; +x_9 = l_Lean_Expr_isHeadBetaTarget(x_1, x_8); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; +lean_dec(x_1); +x_10 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_elimLooseBVarsByBeta___lambda__1___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_4); +return x_11; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = l_Lean_Expr_headBeta(x_1); +x_13 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_13, 0, x_12); +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; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_elimLooseBVarsByBeta___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 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_5, 0, x_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); +return x_6; +} +} +static lean_object* _init_l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_elimLooseBVarsByBeta___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_elimLooseBVarsByBeta___lambda__1___boxed), 4, 0); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_elimLooseBVarsByBeta___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_elimLooseBVarsByBeta___lambda__2___boxed), 4, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_elimLooseBVarsByBeta(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; +x_5 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_elimLooseBVarsByBeta___closed__1; +x_6 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_elimLooseBVarsByBeta___closed__2; +x_7 = l_Lean_Core_transform___at_Lean_Core_betaReduce___spec__1(x_1, x_5, x_6, x_2, x_3, x_4); +return x_7; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_elimLooseBVarsByBeta___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___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_elimLooseBVarsByBeta___lambda__1(x_1, x_2, x_3, x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_elimLooseBVarsByBeta___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___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_elimLooseBVarsByBeta___lambda__2(x_1, x_2, x_3, x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_5; +} +} LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_reduceUntilBadKey_step(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) { _start: { @@ -3949,6 +4076,41 @@ return x_36; } } } +LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_pushArgs___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, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +uint8_t x_10; +x_10 = l_Lean_Expr_hasLooseBVars(x_4); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_11 = lean_box(5); +x_12 = lean_array_push(x_2, x_3); +x_13 = lean_array_push(x_12, x_4); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_11); +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_9); +return x_15; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +lean_dec(x_4); +lean_dec(x_3); +x_16 = lean_box(4); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_2); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_9); +return x_18; +} +} +} LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_pushArgs(uint8_t 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: { @@ -4338,11 +4500,8 @@ return x_91; case 7: { lean_object* x_92; lean_object* x_93; uint8_t x_94; +lean_free_object(x_11); lean_dec(x_13); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); x_92 = lean_ctor_get(x_15, 1); lean_inc(x_92); x_93 = lean_ctor_get(x_15, 2); @@ -4351,507 +4510,519 @@ lean_dec(x_15); x_94 = l_Lean_Expr_hasLooseBVars(x_93); if (x_94 == 0) { -lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; -x_95 = lean_box(5); -x_96 = lean_array_push(x_3, x_92); -x_97 = lean_array_push(x_96, x_93); -x_98 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_98, 0, x_95); -lean_ctor_set(x_98, 1, x_97); -lean_ctor_set(x_11, 0, x_98); -return x_11; +lean_object* x_95; +x_95 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_pushArgs___lambda__2(x_1, x_3, x_92, x_93, x_5, x_6, x_7, x_8, x_14); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +return x_95; } else { -lean_object* x_99; lean_object* x_100; -lean_dec(x_93); +lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_96 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_elimLooseBVarsByBeta___closed__1; +x_97 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_elimLooseBVarsByBeta___closed__2; +lean_inc(x_8); +lean_inc(x_7); +x_98 = l_Lean_Core_transform___at_Lean_Core_betaReduce___spec__1(x_93, x_96, x_97, x_7, x_8, x_14); +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, 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_Meta_DiscrTree_0__Lean_Meta_DiscrTree_pushArgs___lambda__2(x_1, x_3, x_92, x_99, x_5, x_6, x_7, x_8, x_100); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +return x_101; +} +else +{ +uint8_t x_102; lean_dec(x_92); -x_99 = lean_box(4); -x_100 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_100, 0, x_99); -lean_ctor_set(x_100, 1, x_3); -lean_ctor_set(x_11, 0, x_100); -return x_11; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +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_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; +} +} } } case 9: { -lean_object* x_101; lean_object* x_102; lean_object* x_103; +lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_dec(x_13); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_101 = lean_ctor_get(x_15, 0); -lean_inc(x_101); +x_106 = lean_ctor_get(x_15, 0); +lean_inc(x_106); lean_dec(x_15); -x_102 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_102, 0, x_101); -x_103 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_103, 0, x_102); -lean_ctor_set(x_103, 1, x_3); -lean_ctor_set(x_11, 0, x_103); +x_107 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_107, 0, 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_3); +lean_ctor_set(x_11, 0, x_108); return x_11; } case 11: { -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; uint8_t x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; +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; uint8_t x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_free_object(x_11); -x_104 = lean_ctor_get(x_15, 0); -lean_inc(x_104); -x_105 = lean_ctor_get(x_15, 1); -lean_inc(x_105); -x_106 = lean_ctor_get(x_15, 2); -lean_inc(x_106); -x_107 = lean_st_ref_get(x_8, x_14); -x_108 = lean_ctor_get(x_107, 0); -lean_inc(x_108); -x_109 = lean_ctor_get(x_107, 1); +x_109 = lean_ctor_get(x_15, 0); lean_inc(x_109); -lean_dec(x_107); -x_110 = lean_ctor_get(x_108, 0); +x_110 = lean_ctor_get(x_15, 1); lean_inc(x_110); -lean_dec(x_108); -lean_inc(x_104); -x_111 = lean_is_class(x_110, x_104); -x_112 = lean_unsigned_to_nat(0u); -x_113 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_13, x_112); +x_111 = lean_ctor_get(x_15, 2); +lean_inc(x_111); +x_112 = lean_st_ref_get(x_8, x_14); +x_113 = lean_ctor_get(x_112, 0); lean_inc(x_113); -x_114 = lean_alloc_ctor(6, 3, 0); -lean_ctor_set(x_114, 0, x_104); -lean_ctor_set(x_114, 1, x_105); -lean_ctor_set(x_114, 2, x_113); +x_114 = lean_ctor_get(x_112, 1); +lean_inc(x_114); +lean_dec(x_112); +x_115 = lean_ctor_get(x_113, 0); +lean_inc(x_115); +lean_dec(x_113); +lean_inc(x_109); +x_116 = lean_is_class(x_115, x_109); +x_117 = lean_unsigned_to_nat(0u); +x_118 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_13, x_117); +lean_inc(x_118); +x_119 = lean_alloc_ctor(6, 3, 0); +lean_ctor_set(x_119, 0, x_109); +lean_ctor_set(x_119, 1, x_110); +lean_ctor_set(x_119, 2, x_118); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_113); -x_115 = l_Lean_Meta_getFunInfoNArgs(x_15, x_113, x_5, x_6, x_7, x_8, x_109); -if (x_111 == 0) +lean_inc(x_118); +x_120 = l_Lean_Meta_getFunInfoNArgs(x_15, x_118, x_5, x_6, x_7, x_8, x_114); +if (x_116 == 0) { -if (lean_obj_tag(x_115) == 0) +if (lean_obj_tag(x_120) == 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; -x_116 = lean_ctor_get(x_115, 0); -lean_inc(x_116); -x_117 = lean_ctor_get(x_115, 1); -lean_inc(x_117); -lean_dec(x_115); -x_118 = lean_array_push(x_3, x_106); -x_119 = lean_ctor_get(x_116, 0); -lean_inc(x_119); -lean_dec(x_116); -x_120 = lean_unsigned_to_nat(1u); -x_121 = lean_nat_sub(x_113, x_120); -lean_dec(x_113); -x_122 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_pushArgsAux(x_119, x_121, x_13, x_118, x_5, x_6, x_7, x_8, x_117); -lean_dec(x_119); -if (lean_obj_tag(x_122) == 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; +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_array_push(x_3, x_111); +x_124 = lean_ctor_get(x_121, 0); +lean_inc(x_124); +lean_dec(x_121); +x_125 = lean_unsigned_to_nat(1u); +x_126 = lean_nat_sub(x_118, x_125); +lean_dec(x_118); +x_127 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_pushArgsAux(x_124, x_126, x_13, x_123, x_5, x_6, x_7, x_8, x_122); +lean_dec(x_124); +if (lean_obj_tag(x_127) == 0) { -uint8_t x_123; -x_123 = !lean_is_exclusive(x_122); -if (x_123 == 0) +uint8_t x_128; +x_128 = !lean_is_exclusive(x_127); +if (x_128 == 0) { -lean_object* x_124; lean_object* x_125; -x_124 = lean_ctor_get(x_122, 0); -x_125 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_125, 0, x_114); -lean_ctor_set(x_125, 1, x_124); -lean_ctor_set(x_122, 0, x_125); -return x_122; +lean_object* x_129; lean_object* x_130; +x_129 = lean_ctor_get(x_127, 0); +x_130 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_130, 0, x_119); +lean_ctor_set(x_130, 1, x_129); +lean_ctor_set(x_127, 0, x_130); +return x_127; } else { -lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; -x_126 = lean_ctor_get(x_122, 0); -x_127 = lean_ctor_get(x_122, 1); -lean_inc(x_127); -lean_inc(x_126); -lean_dec(x_122); -x_128 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_128, 0, x_114); -lean_ctor_set(x_128, 1, 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_127); -return x_129; +lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; +x_131 = lean_ctor_get(x_127, 0); +x_132 = lean_ctor_get(x_127, 1); +lean_inc(x_132); +lean_inc(x_131); +lean_dec(x_127); +x_133 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_133, 0, x_119); +lean_ctor_set(x_133, 1, 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_132); +return x_134; } } else { -uint8_t x_130; -lean_dec(x_114); -x_130 = !lean_is_exclusive(x_122); -if (x_130 == 0) +uint8_t x_135; +lean_dec(x_119); +x_135 = !lean_is_exclusive(x_127); +if (x_135 == 0) { -return x_122; +return x_127; } else { -lean_object* x_131; lean_object* x_132; lean_object* x_133; -x_131 = lean_ctor_get(x_122, 0); -x_132 = lean_ctor_get(x_122, 1); -lean_inc(x_132); -lean_inc(x_131); -lean_dec(x_122); -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_object* x_136; lean_object* x_137; lean_object* x_138; +x_136 = lean_ctor_get(x_127, 0); +x_137 = lean_ctor_get(x_127, 1); +lean_inc(x_137); +lean_inc(x_136); +lean_dec(x_127); +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_134; -lean_dec(x_114); -lean_dec(x_113); -lean_dec(x_106); +uint8_t x_139; +lean_dec(x_119); +lean_dec(x_118); +lean_dec(x_111); lean_dec(x_13); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_134 = !lean_is_exclusive(x_115); -if (x_134 == 0) +x_139 = !lean_is_exclusive(x_120); +if (x_139 == 0) { -return x_115; +return x_120; } else { -lean_object* x_135; lean_object* x_136; lean_object* x_137; -x_135 = lean_ctor_get(x_115, 0); -x_136 = lean_ctor_get(x_115, 1); -lean_inc(x_136); -lean_inc(x_135); -lean_dec(x_115); -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_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; } } } else { -if (lean_obj_tag(x_115) == 0) +if (lean_obj_tag(x_120) == 0) { -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_138 = lean_ctor_get(x_115, 0); -lean_inc(x_138); -x_139 = lean_ctor_get(x_115, 1); -lean_inc(x_139); -lean_dec(x_115); -x_140 = l_Lean_Meta_DiscrTree_mkNoindexAnnotation(x_106); -x_141 = lean_array_push(x_3, x_140); -x_142 = lean_ctor_get(x_138, 0); -lean_inc(x_142); -lean_dec(x_138); -x_143 = lean_unsigned_to_nat(1u); -x_144 = lean_nat_sub(x_113, x_143); -lean_dec(x_113); -x_145 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_pushArgsAux(x_142, x_144, x_13, x_141, x_5, x_6, x_7, x_8, x_139); -lean_dec(x_142); -if (lean_obj_tag(x_145) == 0) -{ -uint8_t x_146; -x_146 = !lean_is_exclusive(x_145); -if (x_146 == 0) -{ -lean_object* x_147; lean_object* x_148; -x_147 = lean_ctor_get(x_145, 0); -x_148 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_148, 0, x_114); -lean_ctor_set(x_148, 1, x_147); -lean_ctor_set(x_145, 0, x_148); -return x_145; +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_143 = lean_ctor_get(x_120, 0); +lean_inc(x_143); +x_144 = lean_ctor_get(x_120, 1); +lean_inc(x_144); +lean_dec(x_120); +x_145 = l_Lean_Meta_DiscrTree_mkNoindexAnnotation(x_111); +x_146 = lean_array_push(x_3, x_145); +x_147 = lean_ctor_get(x_143, 0); +lean_inc(x_147); +lean_dec(x_143); +x_148 = lean_unsigned_to_nat(1u); +x_149 = lean_nat_sub(x_118, x_148); +lean_dec(x_118); +x_150 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_pushArgsAux(x_147, x_149, x_13, x_146, x_5, x_6, x_7, x_8, x_144); +lean_dec(x_147); +if (lean_obj_tag(x_150) == 0) +{ +uint8_t x_151; +x_151 = !lean_is_exclusive(x_150); +if (x_151 == 0) +{ +lean_object* x_152; lean_object* x_153; +x_152 = lean_ctor_get(x_150, 0); +x_153 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_153, 0, x_119); +lean_ctor_set(x_153, 1, x_152); +lean_ctor_set(x_150, 0, x_153); +return x_150; } else { -lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; -x_149 = lean_ctor_get(x_145, 0); -x_150 = lean_ctor_get(x_145, 1); -lean_inc(x_150); -lean_inc(x_149); -lean_dec(x_145); -x_151 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_151, 0, x_114); -lean_ctor_set(x_151, 1, x_149); -x_152 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_152, 0, x_151); -lean_ctor_set(x_152, 1, x_150); -return x_152; +lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; +x_154 = lean_ctor_get(x_150, 0); +x_155 = lean_ctor_get(x_150, 1); +lean_inc(x_155); +lean_inc(x_154); +lean_dec(x_150); +x_156 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_156, 0, x_119); +lean_ctor_set(x_156, 1, x_154); +x_157 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_157, 0, x_156); +lean_ctor_set(x_157, 1, x_155); +return x_157; } } else { -uint8_t x_153; -lean_dec(x_114); -x_153 = !lean_is_exclusive(x_145); -if (x_153 == 0) +uint8_t x_158; +lean_dec(x_119); +x_158 = !lean_is_exclusive(x_150); +if (x_158 == 0) { -return x_145; +return x_150; } else { -lean_object* x_154; lean_object* x_155; lean_object* x_156; -x_154 = lean_ctor_get(x_145, 0); -x_155 = lean_ctor_get(x_145, 1); -lean_inc(x_155); -lean_inc(x_154); -lean_dec(x_145); -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_object* x_159; lean_object* x_160; lean_object* x_161; +x_159 = lean_ctor_get(x_150, 0); +x_160 = lean_ctor_get(x_150, 1); +lean_inc(x_160); +lean_inc(x_159); +lean_dec(x_150); +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 { -uint8_t x_157; -lean_dec(x_114); -lean_dec(x_113); -lean_dec(x_106); +uint8_t x_162; +lean_dec(x_119); +lean_dec(x_118); +lean_dec(x_111); lean_dec(x_13); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_157 = !lean_is_exclusive(x_115); -if (x_157 == 0) +x_162 = !lean_is_exclusive(x_120); +if (x_162 == 0) { -return x_115; +return x_120; } else { -lean_object* x_158; lean_object* x_159; lean_object* x_160; -x_158 = lean_ctor_get(x_115, 0); -x_159 = lean_ctor_get(x_115, 1); -lean_inc(x_159); -lean_inc(x_158); -lean_dec(x_115); -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_object* x_163; lean_object* x_164; lean_object* x_165; +x_163 = lean_ctor_get(x_120, 0); +x_164 = lean_ctor_get(x_120, 1); +lean_inc(x_164); +lean_inc(x_163); +lean_dec(x_120); +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; } } } } default: { -lean_object* x_161; lean_object* x_162; +lean_object* x_166; lean_object* x_167; lean_dec(x_15); lean_dec(x_13); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_161 = lean_box(4); -x_162 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_162, 0, x_161); -lean_ctor_set(x_162, 1, x_3); -lean_ctor_set(x_11, 0, x_162); +x_166 = lean_box(4); +x_167 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_167, 0, x_166); +lean_ctor_set(x_167, 1, x_3); +lean_ctor_set(x_11, 0, x_167); return x_11; } } } else { -lean_object* x_163; lean_object* x_164; lean_object* x_165; -x_163 = lean_ctor_get(x_11, 0); -x_164 = lean_ctor_get(x_11, 1); -lean_inc(x_164); -lean_inc(x_163); +lean_object* x_168; lean_object* x_169; lean_object* x_170; +x_168 = lean_ctor_get(x_11, 0); +x_169 = lean_ctor_get(x_11, 1); +lean_inc(x_169); +lean_inc(x_168); lean_dec(x_11); -x_165 = l_Lean_Expr_getAppFn(x_163); -switch (lean_obj_tag(x_165)) { +x_170 = l_Lean_Expr_getAppFn(x_168); +switch (lean_obj_tag(x_170)) { case 1: { -lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; -x_166 = lean_ctor_get(x_165, 0); -lean_inc(x_166); -x_167 = lean_unsigned_to_nat(0u); -x_168 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_163, x_167); -lean_inc(x_168); -x_169 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_169, 0, x_166); -lean_ctor_set(x_169, 1, x_168); +lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; +x_171 = lean_ctor_get(x_170, 0); +lean_inc(x_171); +x_172 = lean_unsigned_to_nat(0u); +x_173 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_168, x_172); +lean_inc(x_173); +x_174 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_174, 0, x_171); +lean_ctor_set(x_174, 1, x_173); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_168); -x_170 = l_Lean_Meta_getFunInfoNArgs(x_165, x_168, x_5, x_6, x_7, x_8, x_164); -if (lean_obj_tag(x_170) == 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; -x_171 = lean_ctor_get(x_170, 0); -lean_inc(x_171); -x_172 = lean_ctor_get(x_170, 1); -lean_inc(x_172); -lean_dec(x_170); -x_173 = lean_ctor_get(x_171, 0); lean_inc(x_173); -lean_dec(x_171); -x_174 = lean_unsigned_to_nat(1u); -x_175 = lean_nat_sub(x_168, x_174); -lean_dec(x_168); -x_176 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_pushArgsAux(x_173, x_175, x_163, x_3, x_5, x_6, x_7, x_8, x_172); -lean_dec(x_173); -if (lean_obj_tag(x_176) == 0) +x_175 = l_Lean_Meta_getFunInfoNArgs(x_170, x_173, x_5, x_6, x_7, x_8, x_169); +if (lean_obj_tag(x_175) == 0) { -lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; -x_177 = lean_ctor_get(x_176, 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; +x_176 = lean_ctor_get(x_175, 0); +lean_inc(x_176); +x_177 = lean_ctor_get(x_175, 1); lean_inc(x_177); -x_178 = lean_ctor_get(x_176, 1); +lean_dec(x_175); +x_178 = lean_ctor_get(x_176, 0); lean_inc(x_178); -if (lean_is_exclusive(x_176)) { - lean_ctor_release(x_176, 0); - lean_ctor_release(x_176, 1); - x_179 = x_176; +lean_dec(x_176); +x_179 = lean_unsigned_to_nat(1u); +x_180 = lean_nat_sub(x_173, x_179); +lean_dec(x_173); +x_181 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_pushArgsAux(x_178, x_180, x_168, x_3, x_5, x_6, x_7, x_8, x_177); +lean_dec(x_178); +if (lean_obj_tag(x_181) == 0) +{ +lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; +x_182 = lean_ctor_get(x_181, 0); +lean_inc(x_182); +x_183 = lean_ctor_get(x_181, 1); +lean_inc(x_183); +if (lean_is_exclusive(x_181)) { + lean_ctor_release(x_181, 0); + lean_ctor_release(x_181, 1); + x_184 = x_181; } else { - lean_dec_ref(x_176); - x_179 = lean_box(0); + lean_dec_ref(x_181); + x_184 = lean_box(0); } -x_180 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_180, 0, x_169); -lean_ctor_set(x_180, 1, x_177); -if (lean_is_scalar(x_179)) { - x_181 = lean_alloc_ctor(0, 2, 0); +x_185 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_185, 0, x_174); +lean_ctor_set(x_185, 1, x_182); +if (lean_is_scalar(x_184)) { + x_186 = lean_alloc_ctor(0, 2, 0); } else { - x_181 = x_179; + x_186 = x_184; } -lean_ctor_set(x_181, 0, x_180); -lean_ctor_set(x_181, 1, x_178); -return x_181; +lean_ctor_set(x_186, 0, x_185); +lean_ctor_set(x_186, 1, x_183); +return x_186; } else { -lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; -lean_dec(x_169); -x_182 = lean_ctor_get(x_176, 0); -lean_inc(x_182); -x_183 = lean_ctor_get(x_176, 1); -lean_inc(x_183); -if (lean_is_exclusive(x_176)) { - lean_ctor_release(x_176, 0); - lean_ctor_release(x_176, 1); - x_184 = x_176; +lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; +lean_dec(x_174); +x_187 = lean_ctor_get(x_181, 0); +lean_inc(x_187); +x_188 = lean_ctor_get(x_181, 1); +lean_inc(x_188); +if (lean_is_exclusive(x_181)) { + lean_ctor_release(x_181, 0); + lean_ctor_release(x_181, 1); + x_189 = x_181; } else { - lean_dec_ref(x_176); - x_184 = lean_box(0); + lean_dec_ref(x_181); + x_189 = lean_box(0); } -if (lean_is_scalar(x_184)) { - x_185 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_189)) { + x_190 = lean_alloc_ctor(1, 2, 0); } else { - x_185 = x_184; + x_190 = x_189; } -lean_ctor_set(x_185, 0, x_182); -lean_ctor_set(x_185, 1, x_183); -return x_185; +lean_ctor_set(x_190, 0, x_187); +lean_ctor_set(x_190, 1, x_188); +return x_190; } } else { -lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; -lean_dec(x_169); +lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; +lean_dec(x_174); +lean_dec(x_173); lean_dec(x_168); -lean_dec(x_163); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_186 = lean_ctor_get(x_170, 0); -lean_inc(x_186); -x_187 = lean_ctor_get(x_170, 1); -lean_inc(x_187); -if (lean_is_exclusive(x_170)) { - lean_ctor_release(x_170, 0); - lean_ctor_release(x_170, 1); - x_188 = x_170; +x_191 = lean_ctor_get(x_175, 0); +lean_inc(x_191); +x_192 = lean_ctor_get(x_175, 1); +lean_inc(x_192); +if (lean_is_exclusive(x_175)) { + lean_ctor_release(x_175, 0); + lean_ctor_release(x_175, 1); + x_193 = x_175; } else { - lean_dec_ref(x_170); - x_188 = lean_box(0); + lean_dec_ref(x_175); + x_193 = lean_box(0); } -if (lean_is_scalar(x_188)) { - x_189 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_193)) { + x_194 = lean_alloc_ctor(1, 2, 0); } else { - x_189 = x_188; + x_194 = x_193; } -lean_ctor_set(x_189, 0, x_186); -lean_ctor_set(x_189, 1, x_187); -return x_189; +lean_ctor_set(x_194, 0, x_191); +lean_ctor_set(x_194, 1, x_192); +return x_194; } } case 2: { -lean_object* x_190; lean_object* x_191; uint8_t x_192; -lean_dec(x_163); -x_190 = lean_ctor_get(x_165, 0); -lean_inc(x_190); -lean_dec(x_165); -x_191 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_tmpMVarId; -x_192 = lean_name_eq(x_190, x_191); -if (x_192 == 0) +lean_object* x_195; lean_object* x_196; uint8_t x_197; +lean_dec(x_168); +x_195 = lean_ctor_get(x_170, 0); +lean_inc(x_195); +lean_dec(x_170); +x_196 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_tmpMVarId; +x_197 = lean_name_eq(x_195, x_196); +if (x_197 == 0) { -lean_object* x_193; -x_193 = l_Lean_MVarId_isReadOnlyOrSyntheticOpaque(x_190, x_5, x_6, x_7, x_8, x_164); +lean_object* x_198; +x_198 = l_Lean_MVarId_isReadOnlyOrSyntheticOpaque(x_195, x_5, x_6, x_7, x_8, x_169); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -if (lean_obj_tag(x_193) == 0) +if (lean_obj_tag(x_198) == 0) { -lean_object* x_194; uint8_t x_195; -x_194 = lean_ctor_get(x_193, 0); -lean_inc(x_194); -x_195 = lean_unbox(x_194); -lean_dec(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; -x_196 = lean_ctor_get(x_193, 1); -lean_inc(x_196); -if (lean_is_exclusive(x_193)) { - lean_ctor_release(x_193, 0); - lean_ctor_release(x_193, 1); - x_197 = x_193; -} else { - lean_dec_ref(x_193); - x_197 = lean_box(0); -} -x_198 = lean_box(3); -x_199 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_199, 0, x_198); -lean_ctor_set(x_199, 1, x_3); -if (lean_is_scalar(x_197)) { - x_200 = lean_alloc_ctor(0, 2, 0); -} else { - x_200 = x_197; -} -lean_ctor_set(x_200, 0, x_199); -lean_ctor_set(x_200, 1, x_196); -return x_200; -} -else +lean_object* x_199; uint8_t x_200; +x_199 = lean_ctor_get(x_198, 0); +lean_inc(x_199); +x_200 = lean_unbox(x_199); +lean_dec(x_199); +if (x_200 == 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_193, 1); +x_201 = lean_ctor_get(x_198, 1); lean_inc(x_201); -if (lean_is_exclusive(x_193)) { - lean_ctor_release(x_193, 0); - lean_ctor_release(x_193, 1); - x_202 = x_193; +if (lean_is_exclusive(x_198)) { + lean_ctor_release(x_198, 0); + lean_ctor_release(x_198, 1); + x_202 = x_198; } else { - lean_dec_ref(x_193); + lean_dec_ref(x_198); x_202 = lean_box(0); } -x_203 = lean_box(4); +x_203 = lean_box(3); x_204 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_204, 0, x_203); lean_ctor_set(x_204, 1, x_3); @@ -4864,534 +5035,592 @@ lean_ctor_set(x_205, 0, x_204); lean_ctor_set(x_205, 1, x_201); return x_205; } +else +{ +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_198, 1); +lean_inc(x_206); +if (lean_is_exclusive(x_198)) { + lean_ctor_release(x_198, 0); + lean_ctor_release(x_198, 1); + x_207 = x_198; +} else { + lean_dec_ref(x_198); + x_207 = lean_box(0); +} +x_208 = lean_box(4); +x_209 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_209, 0, x_208); +lean_ctor_set(x_209, 1, x_3); +if (lean_is_scalar(x_207)) { + x_210 = lean_alloc_ctor(0, 2, 0); +} else { + x_210 = x_207; +} +lean_ctor_set(x_210, 0, x_209); +lean_ctor_set(x_210, 1, x_206); +return x_210; +} } else { -lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; +lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_dec(x_3); -x_206 = lean_ctor_get(x_193, 0); -lean_inc(x_206); -x_207 = lean_ctor_get(x_193, 1); -lean_inc(x_207); -if (lean_is_exclusive(x_193)) { - lean_ctor_release(x_193, 0); - lean_ctor_release(x_193, 1); - x_208 = x_193; +x_211 = lean_ctor_get(x_198, 0); +lean_inc(x_211); +x_212 = lean_ctor_get(x_198, 1); +lean_inc(x_212); +if (lean_is_exclusive(x_198)) { + lean_ctor_release(x_198, 0); + lean_ctor_release(x_198, 1); + x_213 = x_198; } else { - lean_dec_ref(x_193); - x_208 = lean_box(0); + lean_dec_ref(x_198); + x_213 = lean_box(0); } -if (lean_is_scalar(x_208)) { - x_209 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_213)) { + x_214 = lean_alloc_ctor(1, 2, 0); } else { - x_209 = x_208; + x_214 = x_213; } -lean_ctor_set(x_209, 0, x_206); -lean_ctor_set(x_209, 1, x_207); -return x_209; +lean_ctor_set(x_214, 0, x_211); +lean_ctor_set(x_214, 1, x_212); +return x_214; } } else { -lean_object* x_210; lean_object* x_211; lean_object* x_212; -lean_dec(x_190); +lean_object* x_215; lean_object* x_216; lean_object* x_217; +lean_dec(x_195); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_210 = lean_box(3); -x_211 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_211, 0, x_210); -lean_ctor_set(x_211, 1, x_3); -x_212 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_212, 0, x_211); -lean_ctor_set(x_212, 1, x_164); -return x_212; +x_215 = lean_box(3); +x_216 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_216, 0, x_215); +lean_ctor_set(x_216, 1, x_3); +x_217 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_217, 0, x_216); +lean_ctor_set(x_217, 1, x_169); +return x_217; } } case 4: { if (x_2 == 0) { -lean_object* x_213; lean_object* x_214; -x_213 = lean_ctor_get(x_165, 0); -lean_inc(x_213); +lean_object* x_218; lean_object* x_219; +x_218 = lean_ctor_get(x_170, 0); +lean_inc(x_218); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_163); -x_214 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_shouldAddAsStar(x_213, x_163, x_5, x_6, x_7, x_8, x_164); -if (lean_obj_tag(x_214) == 0) +lean_inc(x_168); +x_219 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_shouldAddAsStar(x_218, x_168, x_5, x_6, x_7, x_8, x_169); +if (lean_obj_tag(x_219) == 0) { -lean_object* x_215; uint8_t x_216; -x_215 = lean_ctor_get(x_214, 0); -lean_inc(x_215); -x_216 = lean_unbox(x_215); -lean_dec(x_215); -if (x_216 == 0) +lean_object* x_220; uint8_t x_221; +x_220 = lean_ctor_get(x_219, 0); +lean_inc(x_220); +x_221 = lean_unbox(x_220); +lean_dec(x_220); +if (x_221 == 0) { -lean_object* x_217; lean_object* x_218; lean_object* x_219; -x_217 = lean_ctor_get(x_214, 1); -lean_inc(x_217); -lean_dec(x_214); -x_218 = lean_box(0); -x_219 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_pushArgs___lambda__1(x_163, x_1, x_213, x_165, x_3, x_218, x_5, x_6, x_7, x_8, x_217); -return x_219; +lean_object* x_222; lean_object* x_223; lean_object* x_224; +x_222 = lean_ctor_get(x_219, 1); +lean_inc(x_222); +lean_dec(x_219); +x_223 = lean_box(0); +x_224 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_pushArgs___lambda__1(x_168, x_1, x_218, x_170, x_3, x_223, x_5, x_6, x_7, x_8, x_222); +return x_224; } else { -lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; -lean_dec(x_213); -lean_dec(x_165); -lean_dec(x_163); +lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; +lean_dec(x_218); +lean_dec(x_170); +lean_dec(x_168); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_220 = lean_ctor_get(x_214, 1); -lean_inc(x_220); -if (lean_is_exclusive(x_214)) { - lean_ctor_release(x_214, 0); - lean_ctor_release(x_214, 1); - x_221 = x_214; +x_225 = lean_ctor_get(x_219, 1); +lean_inc(x_225); +if (lean_is_exclusive(x_219)) { + lean_ctor_release(x_219, 0); + lean_ctor_release(x_219, 1); + x_226 = x_219; } else { - lean_dec_ref(x_214); - x_221 = lean_box(0); -} -x_222 = lean_box(3); -x_223 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_223, 0, x_222); -lean_ctor_set(x_223, 1, x_3); -if (lean_is_scalar(x_221)) { - x_224 = lean_alloc_ctor(0, 2, 0); + lean_dec_ref(x_219); + x_226 = lean_box(0); +} +x_227 = lean_box(3); +x_228 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_228, 0, x_227); +lean_ctor_set(x_228, 1, x_3); +if (lean_is_scalar(x_226)) { + x_229 = lean_alloc_ctor(0, 2, 0); } else { - x_224 = x_221; + x_229 = x_226; } -lean_ctor_set(x_224, 0, x_223); -lean_ctor_set(x_224, 1, x_220); -return x_224; +lean_ctor_set(x_229, 0, x_228); +lean_ctor_set(x_229, 1, x_225); +return x_229; } } else { -lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; -lean_dec(x_213); -lean_dec(x_165); -lean_dec(x_163); +lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; +lean_dec(x_218); +lean_dec(x_170); +lean_dec(x_168); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_225 = lean_ctor_get(x_214, 0); -lean_inc(x_225); -x_226 = lean_ctor_get(x_214, 1); -lean_inc(x_226); -if (lean_is_exclusive(x_214)) { - lean_ctor_release(x_214, 0); - lean_ctor_release(x_214, 1); - x_227 = x_214; +x_230 = lean_ctor_get(x_219, 0); +lean_inc(x_230); +x_231 = lean_ctor_get(x_219, 1); +lean_inc(x_231); +if (lean_is_exclusive(x_219)) { + lean_ctor_release(x_219, 0); + lean_ctor_release(x_219, 1); + x_232 = x_219; } else { - lean_dec_ref(x_214); - x_227 = lean_box(0); + lean_dec_ref(x_219); + x_232 = lean_box(0); } -if (lean_is_scalar(x_227)) { - x_228 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_232)) { + x_233 = lean_alloc_ctor(1, 2, 0); } else { - x_228 = x_227; + x_233 = x_232; } -lean_ctor_set(x_228, 0, x_225); -lean_ctor_set(x_228, 1, x_226); -return x_228; +lean_ctor_set(x_233, 0, x_230); +lean_ctor_set(x_233, 1, x_231); +return x_233; +} +} +else +{ +lean_object* x_234; lean_object* x_235; lean_object* x_236; +x_234 = lean_ctor_get(x_170, 0); +lean_inc(x_234); +x_235 = lean_box(0); +x_236 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_pushArgs___lambda__1(x_168, x_1, x_234, x_170, x_3, x_235, x_5, x_6, x_7, x_8, x_169); +return x_236; } } +case 7: +{ +lean_object* x_237; lean_object* x_238; uint8_t x_239; +lean_dec(x_168); +x_237 = lean_ctor_get(x_170, 1); +lean_inc(x_237); +x_238 = lean_ctor_get(x_170, 2); +lean_inc(x_238); +lean_dec(x_170); +x_239 = l_Lean_Expr_hasLooseBVars(x_238); +if (x_239 == 0) +{ +lean_object* x_240; +x_240 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_pushArgs___lambda__2(x_1, x_3, x_237, x_238, x_5, x_6, x_7, x_8, x_169); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +return x_240; +} else { -lean_object* x_229; lean_object* x_230; lean_object* x_231; -x_229 = lean_ctor_get(x_165, 0); -lean_inc(x_229); -x_230 = lean_box(0); -x_231 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_pushArgs___lambda__1(x_163, x_1, x_229, x_165, x_3, x_230, x_5, x_6, x_7, x_8, x_164); -return x_231; -} -} -case 7: +lean_object* x_241; lean_object* x_242; lean_object* x_243; +x_241 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_elimLooseBVarsByBeta___closed__1; +x_242 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_elimLooseBVarsByBeta___closed__2; +lean_inc(x_8); +lean_inc(x_7); +x_243 = l_Lean_Core_transform___at_Lean_Core_betaReduce___spec__1(x_238, x_241, x_242, x_7, x_8, x_169); +if (lean_obj_tag(x_243) == 0) { -lean_object* x_232; lean_object* x_233; uint8_t x_234; -lean_dec(x_163); +lean_object* x_244; lean_object* x_245; lean_object* x_246; +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_246 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_pushArgs___lambda__2(x_1, x_3, x_237, x_244, x_5, x_6, x_7, x_8, x_245); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_232 = lean_ctor_get(x_165, 1); -lean_inc(x_232); -x_233 = lean_ctor_get(x_165, 2); -lean_inc(x_233); -lean_dec(x_165); -x_234 = l_Lean_Expr_hasLooseBVars(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; -x_235 = lean_box(5); -x_236 = lean_array_push(x_3, x_232); -x_237 = lean_array_push(x_236, x_233); -x_238 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_238, 0, x_235); -lean_ctor_set(x_238, 1, x_237); -x_239 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_239, 0, x_238); -lean_ctor_set(x_239, 1, x_164); -return x_239; +return x_246; } else { -lean_object* x_240; lean_object* x_241; lean_object* x_242; -lean_dec(x_233); -lean_dec(x_232); -x_240 = lean_box(4); -x_241 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_241, 0, x_240); -lean_ctor_set(x_241, 1, x_3); -x_242 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_242, 0, x_241); -lean_ctor_set(x_242, 1, x_164); -return x_242; +lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; +lean_dec(x_237); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +x_247 = lean_ctor_get(x_243, 0); +lean_inc(x_247); +x_248 = lean_ctor_get(x_243, 1); +lean_inc(x_248); +if (lean_is_exclusive(x_243)) { + lean_ctor_release(x_243, 0); + lean_ctor_release(x_243, 1); + x_249 = x_243; +} else { + lean_dec_ref(x_243); + x_249 = lean_box(0); +} +if (lean_is_scalar(x_249)) { + x_250 = lean_alloc_ctor(1, 2, 0); +} else { + x_250 = x_249; +} +lean_ctor_set(x_250, 0, x_247); +lean_ctor_set(x_250, 1, x_248); +return x_250; +} } } case 9: { -lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; -lean_dec(x_163); +lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; +lean_dec(x_168); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_243 = lean_ctor_get(x_165, 0); -lean_inc(x_243); -lean_dec(x_165); -x_244 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_244, 0, 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_3); -x_246 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_246, 0, x_245); -lean_ctor_set(x_246, 1, x_164); -return x_246; +x_251 = lean_ctor_get(x_170, 0); +lean_inc(x_251); +lean_dec(x_170); +x_252 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_252, 0, x_251); +x_253 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_253, 0, x_252); +lean_ctor_set(x_253, 1, x_3); +x_254 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_254, 0, x_253); +lean_ctor_set(x_254, 1, x_169); +return x_254; } case 11: { -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; uint8_t x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; -x_247 = lean_ctor_get(x_165, 0); -lean_inc(x_247); -x_248 = lean_ctor_get(x_165, 1); -lean_inc(x_248); -x_249 = lean_ctor_get(x_165, 2); -lean_inc(x_249); -x_250 = lean_st_ref_get(x_8, x_164); -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_251, 0); -lean_inc(x_253); -lean_dec(x_251); -lean_inc(x_247); -x_254 = lean_is_class(x_253, x_247); -x_255 = lean_unsigned_to_nat(0u); -x_256 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_163, x_255); -lean_inc(x_256); -x_257 = lean_alloc_ctor(6, 3, 0); -lean_ctor_set(x_257, 0, x_247); -lean_ctor_set(x_257, 1, x_248); -lean_ctor_set(x_257, 2, x_256); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); +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; uint8_t x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; +x_255 = lean_ctor_get(x_170, 0); +lean_inc(x_255); +x_256 = lean_ctor_get(x_170, 1); lean_inc(x_256); -x_258 = l_Lean_Meta_getFunInfoNArgs(x_165, x_256, x_5, x_6, x_7, x_8, x_252); -if (x_254 == 0) -{ -if (lean_obj_tag(x_258) == 0) -{ -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; +x_257 = lean_ctor_get(x_170, 2); +lean_inc(x_257); +x_258 = lean_st_ref_get(x_8, x_169); 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_array_push(x_3, x_249); -x_262 = lean_ctor_get(x_259, 0); -lean_inc(x_262); +x_261 = lean_ctor_get(x_259, 0); +lean_inc(x_261); lean_dec(x_259); -x_263 = lean_unsigned_to_nat(1u); -x_264 = lean_nat_sub(x_256, x_263); -lean_dec(x_256); -x_265 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_pushArgsAux(x_262, x_264, x_163, x_261, x_5, x_6, x_7, x_8, x_260); -lean_dec(x_262); -if (lean_obj_tag(x_265) == 0) +lean_inc(x_255); +x_262 = lean_is_class(x_261, x_255); +x_263 = lean_unsigned_to_nat(0u); +x_264 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_168, x_263); +lean_inc(x_264); +x_265 = lean_alloc_ctor(6, 3, 0); +lean_ctor_set(x_265, 0, x_255); +lean_ctor_set(x_265, 1, x_256); +lean_ctor_set(x_265, 2, x_264); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_264); +x_266 = l_Lean_Meta_getFunInfoNArgs(x_170, x_264, x_5, x_6, x_7, x_8, x_260); +if (x_262 == 0) { -lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; -x_266 = lean_ctor_get(x_265, 0); -lean_inc(x_266); -x_267 = lean_ctor_get(x_265, 1); +if (lean_obj_tag(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; +x_267 = lean_ctor_get(x_266, 0); lean_inc(x_267); -if (lean_is_exclusive(x_265)) { - lean_ctor_release(x_265, 0); - lean_ctor_release(x_265, 1); - x_268 = x_265; +x_268 = lean_ctor_get(x_266, 1); +lean_inc(x_268); +lean_dec(x_266); +x_269 = lean_array_push(x_3, x_257); +x_270 = lean_ctor_get(x_267, 0); +lean_inc(x_270); +lean_dec(x_267); +x_271 = lean_unsigned_to_nat(1u); +x_272 = lean_nat_sub(x_264, x_271); +lean_dec(x_264); +x_273 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_pushArgsAux(x_270, x_272, x_168, x_269, x_5, x_6, x_7, x_8, x_268); +lean_dec(x_270); +if (lean_obj_tag(x_273) == 0) +{ +lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; +x_274 = lean_ctor_get(x_273, 0); +lean_inc(x_274); +x_275 = lean_ctor_get(x_273, 1); +lean_inc(x_275); +if (lean_is_exclusive(x_273)) { + lean_ctor_release(x_273, 0); + lean_ctor_release(x_273, 1); + x_276 = x_273; } else { - lean_dec_ref(x_265); - x_268 = lean_box(0); -} -x_269 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_269, 0, x_257); -lean_ctor_set(x_269, 1, x_266); -if (lean_is_scalar(x_268)) { - x_270 = lean_alloc_ctor(0, 2, 0); + lean_dec_ref(x_273); + x_276 = lean_box(0); +} +x_277 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_277, 0, x_265); +lean_ctor_set(x_277, 1, x_274); +if (lean_is_scalar(x_276)) { + x_278 = lean_alloc_ctor(0, 2, 0); } else { - x_270 = x_268; + x_278 = x_276; } -lean_ctor_set(x_270, 0, x_269); -lean_ctor_set(x_270, 1, x_267); -return x_270; +lean_ctor_set(x_278, 0, x_277); +lean_ctor_set(x_278, 1, x_275); +return x_278; } else { -lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; -lean_dec(x_257); -x_271 = lean_ctor_get(x_265, 0); -lean_inc(x_271); -x_272 = lean_ctor_get(x_265, 1); -lean_inc(x_272); -if (lean_is_exclusive(x_265)) { - lean_ctor_release(x_265, 0); - lean_ctor_release(x_265, 1); - x_273 = x_265; +lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; +lean_dec(x_265); +x_279 = lean_ctor_get(x_273, 0); +lean_inc(x_279); +x_280 = lean_ctor_get(x_273, 1); +lean_inc(x_280); +if (lean_is_exclusive(x_273)) { + lean_ctor_release(x_273, 0); + lean_ctor_release(x_273, 1); + x_281 = x_273; } else { - lean_dec_ref(x_265); - x_273 = lean_box(0); + lean_dec_ref(x_273); + x_281 = lean_box(0); } -if (lean_is_scalar(x_273)) { - x_274 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_281)) { + x_282 = lean_alloc_ctor(1, 2, 0); } else { - x_274 = x_273; + x_282 = x_281; } -lean_ctor_set(x_274, 0, x_271); -lean_ctor_set(x_274, 1, x_272); -return x_274; +lean_ctor_set(x_282, 0, x_279); +lean_ctor_set(x_282, 1, x_280); +return x_282; } } else { -lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; +lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; +lean_dec(x_265); +lean_dec(x_264); lean_dec(x_257); -lean_dec(x_256); -lean_dec(x_249); -lean_dec(x_163); +lean_dec(x_168); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_275 = lean_ctor_get(x_258, 0); -lean_inc(x_275); -x_276 = lean_ctor_get(x_258, 1); -lean_inc(x_276); -if (lean_is_exclusive(x_258)) { - lean_ctor_release(x_258, 0); - lean_ctor_release(x_258, 1); - x_277 = x_258; +x_283 = lean_ctor_get(x_266, 0); +lean_inc(x_283); +x_284 = lean_ctor_get(x_266, 1); +lean_inc(x_284); +if (lean_is_exclusive(x_266)) { + lean_ctor_release(x_266, 0); + lean_ctor_release(x_266, 1); + x_285 = x_266; } else { - lean_dec_ref(x_258); - x_277 = lean_box(0); + lean_dec_ref(x_266); + x_285 = lean_box(0); } -if (lean_is_scalar(x_277)) { - x_278 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_285)) { + x_286 = lean_alloc_ctor(1, 2, 0); } else { - x_278 = x_277; + x_286 = x_285; } -lean_ctor_set(x_278, 0, x_275); -lean_ctor_set(x_278, 1, x_276); -return x_278; +lean_ctor_set(x_286, 0, x_283); +lean_ctor_set(x_286, 1, x_284); +return x_286; } } else { -if (lean_obj_tag(x_258) == 0) -{ -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_279 = lean_ctor_get(x_258, 0); -lean_inc(x_279); -x_280 = lean_ctor_get(x_258, 1); -lean_inc(x_280); -lean_dec(x_258); -x_281 = l_Lean_Meta_DiscrTree_mkNoindexAnnotation(x_249); -x_282 = lean_array_push(x_3, x_281); -x_283 = lean_ctor_get(x_279, 0); -lean_inc(x_283); -lean_dec(x_279); -x_284 = lean_unsigned_to_nat(1u); -x_285 = lean_nat_sub(x_256, x_284); -lean_dec(x_256); -x_286 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_pushArgsAux(x_283, x_285, x_163, x_282, x_5, x_6, x_7, x_8, x_280); -lean_dec(x_283); -if (lean_obj_tag(x_286) == 0) +if (lean_obj_tag(x_266) == 0) { -lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; -x_287 = lean_ctor_get(x_286, 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; +x_287 = lean_ctor_get(x_266, 0); lean_inc(x_287); -x_288 = lean_ctor_get(x_286, 1); +x_288 = lean_ctor_get(x_266, 1); lean_inc(x_288); -if (lean_is_exclusive(x_286)) { - lean_ctor_release(x_286, 0); - lean_ctor_release(x_286, 1); - x_289 = x_286; +lean_dec(x_266); +x_289 = l_Lean_Meta_DiscrTree_mkNoindexAnnotation(x_257); +x_290 = lean_array_push(x_3, x_289); +x_291 = lean_ctor_get(x_287, 0); +lean_inc(x_291); +lean_dec(x_287); +x_292 = lean_unsigned_to_nat(1u); +x_293 = lean_nat_sub(x_264, x_292); +lean_dec(x_264); +x_294 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_pushArgsAux(x_291, x_293, x_168, x_290, x_5, x_6, x_7, x_8, x_288); +lean_dec(x_291); +if (lean_obj_tag(x_294) == 0) +{ +lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; +x_295 = lean_ctor_get(x_294, 0); +lean_inc(x_295); +x_296 = lean_ctor_get(x_294, 1); +lean_inc(x_296); +if (lean_is_exclusive(x_294)) { + lean_ctor_release(x_294, 0); + lean_ctor_release(x_294, 1); + x_297 = x_294; } else { - lean_dec_ref(x_286); - x_289 = lean_box(0); -} -x_290 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_290, 0, x_257); -lean_ctor_set(x_290, 1, x_287); -if (lean_is_scalar(x_289)) { - x_291 = lean_alloc_ctor(0, 2, 0); + lean_dec_ref(x_294); + x_297 = lean_box(0); +} +x_298 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_298, 0, x_265); +lean_ctor_set(x_298, 1, x_295); +if (lean_is_scalar(x_297)) { + x_299 = lean_alloc_ctor(0, 2, 0); } else { - x_291 = x_289; + x_299 = x_297; } -lean_ctor_set(x_291, 0, x_290); -lean_ctor_set(x_291, 1, x_288); -return x_291; +lean_ctor_set(x_299, 0, x_298); +lean_ctor_set(x_299, 1, x_296); +return x_299; } else { -lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; -lean_dec(x_257); -x_292 = lean_ctor_get(x_286, 0); -lean_inc(x_292); -x_293 = lean_ctor_get(x_286, 1); -lean_inc(x_293); -if (lean_is_exclusive(x_286)) { - lean_ctor_release(x_286, 0); - lean_ctor_release(x_286, 1); - x_294 = x_286; +lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; +lean_dec(x_265); +x_300 = lean_ctor_get(x_294, 0); +lean_inc(x_300); +x_301 = lean_ctor_get(x_294, 1); +lean_inc(x_301); +if (lean_is_exclusive(x_294)) { + lean_ctor_release(x_294, 0); + lean_ctor_release(x_294, 1); + x_302 = x_294; } else { - lean_dec_ref(x_286); - x_294 = lean_box(0); + lean_dec_ref(x_294); + x_302 = lean_box(0); } -if (lean_is_scalar(x_294)) { - x_295 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_302)) { + x_303 = lean_alloc_ctor(1, 2, 0); } else { - x_295 = x_294; + x_303 = x_302; } -lean_ctor_set(x_295, 0, x_292); -lean_ctor_set(x_295, 1, x_293); -return x_295; +lean_ctor_set(x_303, 0, x_300); +lean_ctor_set(x_303, 1, x_301); +return x_303; } } else { -lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; +lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; +lean_dec(x_265); +lean_dec(x_264); lean_dec(x_257); -lean_dec(x_256); -lean_dec(x_249); -lean_dec(x_163); +lean_dec(x_168); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_296 = lean_ctor_get(x_258, 0); -lean_inc(x_296); -x_297 = lean_ctor_get(x_258, 1); -lean_inc(x_297); -if (lean_is_exclusive(x_258)) { - lean_ctor_release(x_258, 0); - lean_ctor_release(x_258, 1); - x_298 = x_258; +x_304 = lean_ctor_get(x_266, 0); +lean_inc(x_304); +x_305 = lean_ctor_get(x_266, 1); +lean_inc(x_305); +if (lean_is_exclusive(x_266)) { + lean_ctor_release(x_266, 0); + lean_ctor_release(x_266, 1); + x_306 = x_266; } else { - lean_dec_ref(x_258); - x_298 = lean_box(0); + lean_dec_ref(x_266); + x_306 = lean_box(0); } -if (lean_is_scalar(x_298)) { - x_299 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_306)) { + x_307 = lean_alloc_ctor(1, 2, 0); } else { - x_299 = x_298; + x_307 = x_306; } -lean_ctor_set(x_299, 0, x_296); -lean_ctor_set(x_299, 1, x_297); -return x_299; +lean_ctor_set(x_307, 0, x_304); +lean_ctor_set(x_307, 1, x_305); +return x_307; } } } default: { -lean_object* x_300; lean_object* x_301; lean_object* x_302; -lean_dec(x_165); -lean_dec(x_163); +lean_object* x_308; lean_object* x_309; lean_object* x_310; +lean_dec(x_170); +lean_dec(x_168); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_300 = lean_box(4); -x_301 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_301, 0, x_300); -lean_ctor_set(x_301, 1, x_3); -x_302 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_302, 0, x_301); -lean_ctor_set(x_302, 1, x_164); -return x_302; +x_308 = lean_box(4); +x_309 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_309, 0, x_308); +lean_ctor_set(x_309, 1, x_3); +x_310 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_310, 0, x_309); +lean_ctor_set(x_310, 1, x_169); +return x_310; } } } } else { -uint8_t x_303; +uint8_t x_311; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_303 = !lean_is_exclusive(x_11); -if (x_303 == 0) +x_311 = !lean_is_exclusive(x_11); +if (x_311 == 0) { return x_11; } else { -lean_object* x_304; lean_object* x_305; lean_object* x_306; -x_304 = lean_ctor_get(x_11, 0); -x_305 = lean_ctor_get(x_11, 1); -lean_inc(x_305); -lean_inc(x_304); +lean_object* x_312; lean_object* x_313; lean_object* x_314; +x_312 = lean_ctor_get(x_11, 0); +x_313 = lean_ctor_get(x_11, 1); +lean_inc(x_313); +lean_inc(x_312); lean_dec(x_11); -x_306 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_306, 0, x_304); -lean_ctor_set(x_306, 1, x_305); -return x_306; +x_314 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_314, 0, x_312); +lean_ctor_set(x_314, 1, x_313); +return x_314; } } } else { -lean_object* x_307; lean_object* x_308; lean_object* x_309; +lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_307 = lean_box(3); -x_308 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_308, 0, x_307); -lean_ctor_set(x_308, 1, x_3); -x_309 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_309, 0, x_308); -lean_ctor_set(x_309, 1, x_9); -return x_309; +x_315 = lean_box(3); +x_316 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_316, 0, x_315); +lean_ctor_set(x_316, 1, x_3); +x_317 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_317, 0, x_316); +lean_ctor_set(x_317, 1, x_9); +return x_317; } } } @@ -5406,6 +5635,20 @@ lean_dec(x_6); return x_13; } } +LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_pushArgs___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) { +_start: +{ +uint8_t x_10; lean_object* x_11; +x_10 = lean_unbox(x_1); +lean_dec(x_1); +x_11 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_pushArgs___lambda__2(x_10, 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); +return x_11; +} +} LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_pushArgs___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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: { @@ -9016,7 +9259,7 @@ static lean_object* _init_l_Lean_Meta_DiscrTree_insertCore___rarg___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_Meta_DiscrTree_insertCore___rarg___closed__1; x_2 = l_Lean_Meta_DiscrTree_insertCore___rarg___closed__2; -x_3 = lean_unsigned_to_nat(366u); +x_3 = lean_unsigned_to_nat(396u); x_4 = lean_unsigned_to_nat(23u); x_5 = l_Lean_Meta_DiscrTree_insertCore___rarg___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -9600,21 +9843,58 @@ lean_ctor_set(x_16, 1, x_9); return x_16; } } -static lean_object* _init_l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getKeyArgs___closed__1() { +static lean_object* _init_l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getKeyArgs___lambda__2___closed__1() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_levelZero; -x_2 = l_Lean_Expr_sort___override(x_1); +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___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getKeyArgs___closed__2() { +LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getKeyArgs___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, lean_object* x_7, lean_object* x_8) { +_start: +{ +uint8_t x_9; +x_9 = l_Lean_Expr_hasLooseBVars(x_3); +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; +x_10 = lean_box(5); +x_11 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getKeyArgs___lambda__2___closed__1; +x_12 = lean_array_push(x_11, x_2); +x_13 = lean_array_push(x_12, x_3); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_10); +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_8); +return x_15; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +lean_dec(x_3); +lean_dec(x_2); +x_16 = lean_box(4); +x_17 = l_Lean_Meta_DiscrTree_instInhabitedTrie___rarg___closed__1; +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_17); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_8); +return x_19; +} +} +} +static lean_object* _init_l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getKeyArgs___closed__1() { _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); +x_1 = l_Lean_levelZero; +x_2 = l_Lean_Expr_sort___override(x_1); return x_2; } } @@ -10058,11 +10338,8 @@ return x_119; case 7: { lean_object* x_120; lean_object* x_121; uint8_t x_122; +lean_free_object(x_10); lean_dec(x_12); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); x_120 = lean_ctor_get(x_14, 1); lean_inc(x_120); x_121 = lean_ctor_get(x_14, 2); @@ -10071,676 +10348,740 @@ lean_dec(x_14); x_122 = l_Lean_Expr_hasLooseBVars(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; -x_123 = lean_box(5); -x_124 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getKeyArgs___closed__2; -x_125 = lean_array_push(x_124, x_120); -x_126 = lean_array_push(x_125, x_121); -x_127 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_127, 0, x_123); -lean_ctor_set(x_127, 1, x_126); -lean_ctor_set(x_10, 0, x_127); -return x_10; +lean_object* x_123; +x_123 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getKeyArgs___lambda__2(x_1, x_120, x_121, x_5, x_6, x_7, x_8, x_13); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +return x_123; } else { -lean_object* x_128; lean_object* x_129; lean_object* x_130; -lean_dec(x_121); +lean_object* x_124; lean_object* x_125; lean_object* x_126; +x_124 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_elimLooseBVarsByBeta___closed__1; +x_125 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_elimLooseBVarsByBeta___closed__2; +lean_inc(x_8); +lean_inc(x_7); +x_126 = l_Lean_Core_transform___at_Lean_Core_betaReduce___spec__1(x_121, x_124, x_125, x_7, x_8, x_13); +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, 0); +lean_inc(x_127); +x_128 = lean_ctor_get(x_126, 1); +lean_inc(x_128); +lean_dec(x_126); +x_129 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getKeyArgs___lambda__2(x_1, x_120, x_127, x_5, x_6, x_7, x_8, x_128); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +return x_129; +} +else +{ +uint8_t x_130; lean_dec(x_120); -x_128 = lean_box(4); -x_129 = l_Lean_Meta_DiscrTree_instInhabitedTrie___rarg___closed__1; -x_130 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_130, 0, x_128); -lean_ctor_set(x_130, 1, x_129); -lean_ctor_set(x_10, 0, x_130); -return x_10; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +x_130 = !lean_is_exclusive(x_126); +if (x_130 == 0) +{ +return x_126; +} +else +{ +lean_object* x_131; lean_object* x_132; lean_object* x_133; +x_131 = lean_ctor_get(x_126, 0); +x_132 = lean_ctor_get(x_126, 1); +lean_inc(x_132); +lean_inc(x_131); +lean_dec(x_126); +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; +} +} } } case 9: { -lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; +lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_dec(x_12); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_131 = lean_ctor_get(x_14, 0); -lean_inc(x_131); +x_134 = lean_ctor_get(x_14, 0); +lean_inc(x_134); lean_dec(x_14); -x_132 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_132, 0, x_131); -x_133 = l_Lean_Meta_DiscrTree_instInhabitedTrie___rarg___closed__1; -x_134 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_134, 0, x_132); -lean_ctor_set(x_134, 1, x_133); -lean_ctor_set(x_10, 0, x_134); +x_135 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_135, 0, x_134); +x_136 = l_Lean_Meta_DiscrTree_instInhabitedTrie___rarg___closed__1; +x_137 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_137, 0, x_135); +lean_ctor_set(x_137, 1, x_136); +lean_ctor_set(x_10, 0, x_137); return x_10; } case 11: { -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_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_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_135 = lean_ctor_get(x_14, 0); -lean_inc(x_135); -x_136 = lean_ctor_get(x_14, 1); -lean_inc(x_136); -x_137 = lean_ctor_get(x_14, 2); -lean_inc(x_137); -lean_dec(x_14); -x_138 = lean_unsigned_to_nat(0u); -x_139 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_12, x_138); +x_138 = lean_ctor_get(x_14, 0); +lean_inc(x_138); +x_139 = lean_ctor_get(x_14, 1); lean_inc(x_139); -x_140 = lean_alloc_ctor(6, 3, 0); -lean_ctor_set(x_140, 0, x_135); -lean_ctor_set(x_140, 1, x_136); -lean_ctor_set(x_140, 2, x_139); -x_141 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_createNodes___closed__1; -x_142 = lean_array_push(x_141, x_137); -x_143 = lean_mk_empty_array_with_capacity(x_139); -lean_dec(x_139); -x_144 = l___private_Lean_Expr_0__Lean_Expr_getAppRevArgsAux(x_12, x_143); -x_145 = l_Array_append___rarg(x_142, x_144); -x_146 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_146, 0, x_140); -lean_ctor_set(x_146, 1, x_145); -lean_ctor_set(x_10, 0, x_146); +x_140 = lean_ctor_get(x_14, 2); +lean_inc(x_140); +lean_dec(x_14); +x_141 = lean_unsigned_to_nat(0u); +x_142 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_12, x_141); +lean_inc(x_142); +x_143 = lean_alloc_ctor(6, 3, 0); +lean_ctor_set(x_143, 0, x_138); +lean_ctor_set(x_143, 1, x_139); +lean_ctor_set(x_143, 2, x_142); +x_144 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_createNodes___closed__1; +x_145 = lean_array_push(x_144, x_140); +x_146 = lean_mk_empty_array_with_capacity(x_142); +lean_dec(x_142); +x_147 = l___private_Lean_Expr_0__Lean_Expr_getAppRevArgsAux(x_12, x_146); +x_148 = l_Array_append___rarg(x_145, x_147); +x_149 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_149, 0, x_143); +lean_ctor_set(x_149, 1, x_148); +lean_ctor_set(x_10, 0, x_149); return x_10; } default: { -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_14); lean_dec(x_12); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_147 = lean_box(4); -x_148 = l_Lean_Meta_DiscrTree_instInhabitedTrie___rarg___closed__1; -x_149 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_149, 0, x_147); -lean_ctor_set(x_149, 1, x_148); -lean_ctor_set(x_10, 0, x_149); +x_150 = lean_box(4); +x_151 = l_Lean_Meta_DiscrTree_instInhabitedTrie___rarg___closed__1; +x_152 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_152, 0, x_150); +lean_ctor_set(x_152, 1, x_151); +lean_ctor_set(x_10, 0, x_152); return x_10; } } } else { -lean_object* x_150; lean_object* x_151; lean_object* x_152; -x_150 = lean_ctor_get(x_10, 0); -x_151 = lean_ctor_get(x_10, 1); -lean_inc(x_151); -lean_inc(x_150); +lean_object* x_153; lean_object* x_154; lean_object* x_155; +x_153 = lean_ctor_get(x_10, 0); +x_154 = lean_ctor_get(x_10, 1); +lean_inc(x_154); +lean_inc(x_153); lean_dec(x_10); -x_152 = l_Lean_Expr_getAppFn(x_150); -switch (lean_obj_tag(x_152)) { +x_155 = l_Lean_Expr_getAppFn(x_153); +switch (lean_obj_tag(x_155)) { case 1: { -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_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_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_153 = lean_ctor_get(x_152, 0); -lean_inc(x_153); -lean_dec(x_152); -x_154 = lean_unsigned_to_nat(0u); -x_155 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_150, x_154); -lean_inc(x_155); -x_156 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_156, 0, x_153); -lean_ctor_set(x_156, 1, x_155); -x_157 = lean_mk_empty_array_with_capacity(x_155); +x_156 = lean_ctor_get(x_155, 0); +lean_inc(x_156); lean_dec(x_155); -x_158 = l___private_Lean_Expr_0__Lean_Expr_getAppRevArgsAux(x_150, x_157); -x_159 = lean_alloc_ctor(0, 2, 0); +x_157 = lean_unsigned_to_nat(0u); +x_158 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_153, x_157); +lean_inc(x_158); +x_159 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_159, 0, x_156); lean_ctor_set(x_159, 1, 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_151); -return x_160; +x_160 = lean_mk_empty_array_with_capacity(x_158); +lean_dec(x_158); +x_161 = l___private_Lean_Expr_0__Lean_Expr_getAppRevArgsAux(x_153, x_160); +x_162 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_162, 0, x_159); +lean_ctor_set(x_162, 1, x_161); +x_163 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_163, 0, x_162); +lean_ctor_set(x_163, 1, x_154); +return x_163; } case 2: { -lean_dec(x_150); +lean_dec(x_153); if (x_3 == 0) { -lean_object* x_161; uint8_t x_162; -x_161 = lean_ctor_get(x_5, 0); -lean_inc(x_161); -x_162 = lean_ctor_get_uint8(x_161, 4); -lean_dec(x_161); -if (x_162 == 0) +lean_object* x_164; uint8_t x_165; +x_164 = lean_ctor_get(x_5, 0); +lean_inc(x_164); +x_165 = lean_ctor_get_uint8(x_164, 4); +lean_dec(x_164); +if (x_165 == 0) { -lean_object* x_163; lean_object* x_164; -x_163 = lean_ctor_get(x_152, 0); -lean_inc(x_163); -lean_dec(x_152); -x_164 = l_Lean_MVarId_isReadOnlyOrSyntheticOpaque(x_163, x_5, x_6, x_7, x_8, x_151); +lean_object* x_166; lean_object* x_167; +x_166 = lean_ctor_get(x_155, 0); +lean_inc(x_166); +lean_dec(x_155); +x_167 = l_Lean_MVarId_isReadOnlyOrSyntheticOpaque(x_166, x_5, x_6, x_7, x_8, x_154); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -if (lean_obj_tag(x_164) == 0) +if (lean_obj_tag(x_167) == 0) { -lean_object* x_165; uint8_t x_166; -x_165 = lean_ctor_get(x_164, 0); -lean_inc(x_165); -x_166 = lean_unbox(x_165); -lean_dec(x_165); -if (x_166 == 0) +lean_object* x_168; uint8_t x_169; +x_168 = lean_ctor_get(x_167, 0); +lean_inc(x_168); +x_169 = lean_unbox(x_168); +lean_dec(x_168); +if (x_169 == 0) { -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_164, 1); -lean_inc(x_167); -if (lean_is_exclusive(x_164)) { - lean_ctor_release(x_164, 0); - lean_ctor_release(x_164, 1); - x_168 = x_164; +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_170 = lean_ctor_get(x_167, 1); +lean_inc(x_170); +if (lean_is_exclusive(x_167)) { + lean_ctor_release(x_167, 0); + lean_ctor_release(x_167, 1); + x_171 = x_167; } else { - lean_dec_ref(x_164); - x_168 = lean_box(0); -} -x_169 = lean_box(3); -x_170 = l_Lean_Meta_DiscrTree_instInhabitedTrie___rarg___closed__1; -x_171 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_171, 0, x_169); -lean_ctor_set(x_171, 1, x_170); -if (lean_is_scalar(x_168)) { - x_172 = lean_alloc_ctor(0, 2, 0); + lean_dec_ref(x_167); + x_171 = lean_box(0); +} +x_172 = lean_box(3); +x_173 = l_Lean_Meta_DiscrTree_instInhabitedTrie___rarg___closed__1; +x_174 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_174, 0, x_172); +lean_ctor_set(x_174, 1, x_173); +if (lean_is_scalar(x_171)) { + x_175 = lean_alloc_ctor(0, 2, 0); } else { - x_172 = x_168; + x_175 = x_171; } -lean_ctor_set(x_172, 0, x_171); -lean_ctor_set(x_172, 1, x_167); -return x_172; +lean_ctor_set(x_175, 0, x_174); +lean_ctor_set(x_175, 1, x_170); +return x_175; } 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; -x_173 = lean_ctor_get(x_164, 1); -lean_inc(x_173); -if (lean_is_exclusive(x_164)) { - lean_ctor_release(x_164, 0); - lean_ctor_release(x_164, 1); - x_174 = x_164; +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_176 = lean_ctor_get(x_167, 1); +lean_inc(x_176); +if (lean_is_exclusive(x_167)) { + lean_ctor_release(x_167, 0); + lean_ctor_release(x_167, 1); + x_177 = x_167; } else { - lean_dec_ref(x_164); - x_174 = lean_box(0); + lean_dec_ref(x_167); + x_177 = lean_box(0); } -x_175 = lean_box(4); -x_176 = l_Lean_Meta_DiscrTree_instInhabitedTrie___rarg___closed__1; -x_177 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_177, 0, x_175); -lean_ctor_set(x_177, 1, x_176); -if (lean_is_scalar(x_174)) { - x_178 = lean_alloc_ctor(0, 2, 0); +x_178 = lean_box(4); +x_179 = l_Lean_Meta_DiscrTree_instInhabitedTrie___rarg___closed__1; +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_177)) { + x_181 = lean_alloc_ctor(0, 2, 0); } else { - x_178 = x_174; + x_181 = x_177; } -lean_ctor_set(x_178, 0, x_177); -lean_ctor_set(x_178, 1, x_173); -return x_178; +lean_ctor_set(x_181, 0, x_180); +lean_ctor_set(x_181, 1, x_176); +return x_181; } } else { -lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; -x_179 = lean_ctor_get(x_164, 0); -lean_inc(x_179); -x_180 = lean_ctor_get(x_164, 1); -lean_inc(x_180); -if (lean_is_exclusive(x_164)) { - lean_ctor_release(x_164, 0); - lean_ctor_release(x_164, 1); - x_181 = x_164; +lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; +x_182 = lean_ctor_get(x_167, 0); +lean_inc(x_182); +x_183 = lean_ctor_get(x_167, 1); +lean_inc(x_183); +if (lean_is_exclusive(x_167)) { + lean_ctor_release(x_167, 0); + lean_ctor_release(x_167, 1); + x_184 = x_167; } else { - lean_dec_ref(x_164); - x_181 = lean_box(0); + lean_dec_ref(x_167); + x_184 = lean_box(0); } -if (lean_is_scalar(x_181)) { - x_182 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_184)) { + x_185 = lean_alloc_ctor(1, 2, 0); } else { - x_182 = x_181; + x_185 = x_184; } -lean_ctor_set(x_182, 0, x_179); -lean_ctor_set(x_182, 1, x_180); -return x_182; +lean_ctor_set(x_185, 0, x_182); +lean_ctor_set(x_185, 1, x_183); +return x_185; } } else { -lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; -lean_dec(x_152); +lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; +lean_dec(x_155); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_183 = lean_box(3); -x_184 = l_Lean_Meta_DiscrTree_instInhabitedTrie___rarg___closed__1; -x_185 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_185, 0, x_183); -lean_ctor_set(x_185, 1, x_184); -x_186 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_186, 0, x_185); -lean_ctor_set(x_186, 1, x_151); -return x_186; +x_186 = lean_box(3); +x_187 = l_Lean_Meta_DiscrTree_instInhabitedTrie___rarg___closed__1; +x_188 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_188, 0, x_186); +lean_ctor_set(x_188, 1, x_187); +x_189 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_189, 0, x_188); +lean_ctor_set(x_189, 1, x_154); +return x_189; } } else { -lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; -lean_dec(x_152); +lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; +lean_dec(x_155); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_187 = lean_box(4); -x_188 = l_Lean_Meta_DiscrTree_instInhabitedTrie___rarg___closed__1; -x_189 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_189, 0, x_187); -lean_ctor_set(x_189, 1, x_188); -x_190 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_190, 0, x_189); -lean_ctor_set(x_190, 1, x_151); -return x_190; +x_190 = lean_box(4); +x_191 = l_Lean_Meta_DiscrTree_instInhabitedTrie___rarg___closed__1; +x_192 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_192, 0, x_190); +lean_ctor_set(x_192, 1, 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_154); +return x_193; } } case 4: { -lean_object* x_191; lean_object* x_192; lean_object* x_193; uint8_t x_194; -x_191 = lean_ctor_get(x_152, 0); -lean_inc(x_191); -lean_dec(x_152); -x_192 = l_Lean_Meta_getConfig(x_5, x_6, x_7, x_8, x_151); -x_193 = lean_ctor_get(x_192, 0); -lean_inc(x_193); -x_194 = lean_ctor_get_uint8(x_193, 4); -lean_dec(x_193); -if (x_194 == 0) -{ -lean_object* x_195; lean_object* x_196; lean_object* x_197; -x_195 = lean_ctor_get(x_192, 1); -lean_inc(x_195); -lean_dec(x_192); -x_196 = lean_box(0); -x_197 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getKeyArgs___lambda__1(x_150, x_1, x_191, x_196, x_5, x_6, x_7, x_8, x_195); +lean_object* x_194; lean_object* x_195; lean_object* x_196; uint8_t x_197; +x_194 = lean_ctor_get(x_155, 0); +lean_inc(x_194); +lean_dec(x_155); +x_195 = l_Lean_Meta_getConfig(x_5, x_6, x_7, x_8, x_154); +x_196 = lean_ctor_get(x_195, 0); +lean_inc(x_196); +x_197 = lean_ctor_get_uint8(x_196, 4); +lean_dec(x_196); +if (x_197 == 0) +{ +lean_object* x_198; lean_object* x_199; lean_object* x_200; +x_198 = lean_ctor_get(x_195, 1); +lean_inc(x_198); +lean_dec(x_195); +x_199 = lean_box(0); +x_200 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getKeyArgs___lambda__1(x_153, x_1, x_194, x_199, x_5, x_6, x_7, x_8, x_198); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -return x_197; +return x_200; } else { -lean_object* x_198; uint8_t x_199; -x_198 = lean_ctor_get(x_192, 1); -lean_inc(x_198); -lean_dec(x_192); -x_199 = l_Lean_Expr_hasExprMVar(x_150); -if (x_199 == 0) +lean_object* x_201; uint8_t x_202; +x_201 = lean_ctor_get(x_195, 1); +lean_inc(x_201); +lean_dec(x_195); +x_202 = l_Lean_Expr_hasExprMVar(x_153); +if (x_202 == 0) { -lean_object* x_200; lean_object* x_201; -x_200 = lean_box(0); -x_201 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getKeyArgs___lambda__1(x_150, x_1, x_191, x_200, x_5, x_6, x_7, x_8, x_198); +lean_object* x_203; lean_object* x_204; +x_203 = lean_box(0); +x_204 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getKeyArgs___lambda__1(x_153, x_1, x_194, x_203, x_5, x_6, x_7, x_8, x_201); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -return x_201; +return x_204; } else { -lean_object* x_202; lean_object* x_203; uint8_t x_204; -lean_inc(x_191); -x_202 = l_Lean_isReducible___at___private_Lean_Meta_Basic_0__Lean_Meta_getDefInfoTemp___spec__1(x_191, x_5, x_6, x_7, x_8, x_198); -x_203 = lean_ctor_get(x_202, 0); -lean_inc(x_203); -x_204 = lean_unbox(x_203); -lean_dec(x_203); -if (x_204 == 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; -x_205 = lean_ctor_get(x_202, 1); -lean_inc(x_205); -lean_dec(x_202); -x_206 = lean_st_ref_get(x_8, x_205); -x_207 = lean_ctor_get(x_206, 0); -lean_inc(x_207); -x_208 = lean_ctor_get(x_206, 1); -lean_inc(x_208); +lean_object* x_205; lean_object* x_206; uint8_t x_207; +lean_inc(x_194); +x_205 = l_Lean_isReducible___at___private_Lean_Meta_Basic_0__Lean_Meta_getDefInfoTemp___spec__1(x_194, x_5, x_6, x_7, x_8, x_201); +x_206 = lean_ctor_get(x_205, 0); +lean_inc(x_206); +x_207 = lean_unbox(x_206); lean_dec(x_206); -x_209 = lean_ctor_get(x_207, 0); -lean_inc(x_209); -lean_dec(x_207); -x_210 = l_Lean_Meta_isMatcherAppCore_x3f(x_209, x_150); -if (lean_obj_tag(x_210) == 0) +if (x_207 == 0) { -lean_object* x_211; lean_object* x_212; uint8_t x_213; -lean_inc(x_191); -x_211 = l_Lean_isRec___at___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getKeyArgs___spec__1(x_191, x_5, x_6, x_7, x_8, x_208); -x_212 = lean_ctor_get(x_211, 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; +x_208 = lean_ctor_get(x_205, 1); +lean_inc(x_208); +lean_dec(x_205); +x_209 = lean_st_ref_get(x_8, x_208); +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_210, 0); lean_inc(x_212); -x_213 = lean_unbox(x_212); -lean_dec(x_212); -if (x_213 == 0) +lean_dec(x_210); +x_213 = l_Lean_Meta_isMatcherAppCore_x3f(x_212, x_153); +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_211, 1); -lean_inc(x_214); -lean_dec(x_211); -x_215 = lean_box(0); -x_216 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getKeyArgs___lambda__1(x_150, x_1, x_191, x_215, x_5, x_6, x_7, x_8, x_214); +lean_object* x_214; lean_object* x_215; uint8_t x_216; +lean_inc(x_194); +x_214 = l_Lean_isRec___at___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getKeyArgs___spec__1(x_194, x_5, x_6, x_7, x_8, x_211); +x_215 = lean_ctor_get(x_214, 0); +lean_inc(x_215); +x_216 = lean_unbox(x_215); +lean_dec(x_215); +if (x_216 == 0) +{ +lean_object* x_217; lean_object* x_218; lean_object* x_219; +x_217 = lean_ctor_get(x_214, 1); +lean_inc(x_217); +lean_dec(x_214); +x_218 = lean_box(0); +x_219 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getKeyArgs___lambda__1(x_153, x_1, x_194, x_218, x_5, x_6, x_7, x_8, x_217); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -return x_216; +return x_219; } else { -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_dec(x_191); -lean_dec(x_150); +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_dec(x_194); +lean_dec(x_153); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_217 = lean_ctor_get(x_211, 1); -lean_inc(x_217); -lean_dec(x_211); -x_218 = l_Lean_Meta_throwIsDefEqStuck___rarg(x_217); -x_219 = lean_ctor_get(x_218, 0); -lean_inc(x_219); -x_220 = lean_ctor_get(x_218, 1); +x_220 = lean_ctor_get(x_214, 1); lean_inc(x_220); -if (lean_is_exclusive(x_218)) { - lean_ctor_release(x_218, 0); - lean_ctor_release(x_218, 1); - x_221 = x_218; +lean_dec(x_214); +x_221 = l_Lean_Meta_throwIsDefEqStuck___rarg(x_220); +x_222 = lean_ctor_get(x_221, 0); +lean_inc(x_222); +x_223 = lean_ctor_get(x_221, 1); +lean_inc(x_223); +if (lean_is_exclusive(x_221)) { + lean_ctor_release(x_221, 0); + lean_ctor_release(x_221, 1); + x_224 = x_221; } else { - lean_dec_ref(x_218); - x_221 = lean_box(0); + lean_dec_ref(x_221); + x_224 = lean_box(0); } -if (lean_is_scalar(x_221)) { - x_222 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_224)) { + x_225 = lean_alloc_ctor(1, 2, 0); } else { - x_222 = x_221; + x_225 = x_224; } -lean_ctor_set(x_222, 0, x_219); -lean_ctor_set(x_222, 1, x_220); -return x_222; +lean_ctor_set(x_225, 0, x_222); +lean_ctor_set(x_225, 1, x_223); +return x_225; } } else { -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; size_t x_238; lean_object* x_239; lean_object* x_240; -x_223 = lean_ctor_get(x_210, 0); -lean_inc(x_223); -lean_dec(x_210); -x_224 = lean_unsigned_to_nat(0u); -x_225 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_150, x_224); -x_226 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getKeyArgs___closed__1; -lean_inc(x_225); -x_227 = lean_mk_array(x_225, x_226); -x_228 = lean_unsigned_to_nat(1u); -x_229 = lean_nat_sub(x_225, x_228); -lean_dec(x_225); -lean_inc(x_150); -x_230 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_150, x_227, x_229); -x_231 = l_Lean_Meta_Match_MatcherInfo_getFirstDiscrPos(x_223); -x_232 = lean_ctor_get(x_223, 1); -lean_inc(x_232); -lean_dec(x_223); -x_233 = lean_nat_add(x_231, x_232); -lean_dec(x_232); -x_234 = l_Array_toSubarray___rarg(x_230, x_231, x_233); -x_235 = lean_ctor_get(x_234, 2); +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; size_t x_239; lean_object* x_240; size_t x_241; lean_object* x_242; lean_object* x_243; +x_226 = lean_ctor_get(x_213, 0); +lean_inc(x_226); +lean_dec(x_213); +x_227 = lean_unsigned_to_nat(0u); +x_228 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_153, x_227); +x_229 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getKeyArgs___closed__1; +lean_inc(x_228); +x_230 = lean_mk_array(x_228, x_229); +x_231 = lean_unsigned_to_nat(1u); +x_232 = lean_nat_sub(x_228, x_231); +lean_dec(x_228); +lean_inc(x_153); +x_233 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_153, x_230, x_232); +x_234 = l_Lean_Meta_Match_MatcherInfo_getFirstDiscrPos(x_226); +x_235 = lean_ctor_get(x_226, 1); lean_inc(x_235); -x_236 = lean_usize_of_nat(x_235); +lean_dec(x_226); +x_236 = lean_nat_add(x_234, x_235); lean_dec(x_235); -x_237 = lean_ctor_get(x_234, 1); -lean_inc(x_237); -x_238 = lean_usize_of_nat(x_237); +x_237 = l_Array_toSubarray___rarg(x_233, x_234, x_236); +x_238 = lean_ctor_get(x_237, 2); +lean_inc(x_238); +x_239 = lean_usize_of_nat(x_238); +lean_dec(x_238); +x_240 = lean_ctor_get(x_237, 1); +lean_inc(x_240); +x_241 = lean_usize_of_nat(x_240); +lean_dec(x_240); +x_242 = lean_box(0); +x_243 = l_Subarray_forInUnsafe_loop___at___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getKeyArgs___spec__2(x_237, x_239, x_241, x_242, x_5, x_6, x_7, x_8, x_211); lean_dec(x_237); -x_239 = lean_box(0); -x_240 = l_Subarray_forInUnsafe_loop___at___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getKeyArgs___spec__2(x_234, x_236, x_238, x_239, x_5, x_6, x_7, x_8, x_208); -lean_dec(x_234); -if (lean_obj_tag(x_240) == 0) +if (lean_obj_tag(x_243) == 0) { -lean_object* x_241; lean_object* x_242; -x_241 = lean_ctor_get(x_240, 1); -lean_inc(x_241); -lean_dec(x_240); -x_242 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getKeyArgs___lambda__1(x_150, x_1, x_191, x_239, x_5, x_6, x_7, x_8, x_241); +lean_object* x_244; lean_object* x_245; +x_244 = lean_ctor_get(x_243, 1); +lean_inc(x_244); +lean_dec(x_243); +x_245 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getKeyArgs___lambda__1(x_153, x_1, x_194, x_242, x_5, x_6, x_7, x_8, x_244); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -return x_242; +return x_245; } else { -lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; -lean_dec(x_191); -lean_dec(x_150); +lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; +lean_dec(x_194); +lean_dec(x_153); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_243 = lean_ctor_get(x_240, 0); -lean_inc(x_243); -x_244 = lean_ctor_get(x_240, 1); -lean_inc(x_244); -if (lean_is_exclusive(x_240)) { - lean_ctor_release(x_240, 0); - lean_ctor_release(x_240, 1); - x_245 = x_240; +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; } else { - lean_dec_ref(x_240); - x_245 = lean_box(0); + lean_dec_ref(x_243); + x_248 = lean_box(0); } -if (lean_is_scalar(x_245)) { - x_246 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_248)) { + x_249 = lean_alloc_ctor(1, 2, 0); } else { - x_246 = x_245; + x_249 = x_248; } -lean_ctor_set(x_246, 0, x_243); -lean_ctor_set(x_246, 1, x_244); -return x_246; +lean_ctor_set(x_249, 0, x_246); +lean_ctor_set(x_249, 1, x_247); +return x_249; } } } 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_dec(x_191); -lean_dec(x_150); +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_194); +lean_dec(x_153); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_247 = lean_ctor_get(x_202, 1); -lean_inc(x_247); -lean_dec(x_202); -x_248 = l_Lean_Meta_throwIsDefEqStuck___rarg(x_247); -x_249 = lean_ctor_get(x_248, 0); -lean_inc(x_249); -x_250 = lean_ctor_get(x_248, 1); +x_250 = lean_ctor_get(x_205, 1); lean_inc(x_250); -if (lean_is_exclusive(x_248)) { - lean_ctor_release(x_248, 0); - lean_ctor_release(x_248, 1); - x_251 = x_248; +lean_dec(x_205); +x_251 = l_Lean_Meta_throwIsDefEqStuck___rarg(x_250); +x_252 = lean_ctor_get(x_251, 0); +lean_inc(x_252); +x_253 = lean_ctor_get(x_251, 1); +lean_inc(x_253); +if (lean_is_exclusive(x_251)) { + lean_ctor_release(x_251, 0); + lean_ctor_release(x_251, 1); + x_254 = x_251; } else { - lean_dec_ref(x_248); - x_251 = lean_box(0); + lean_dec_ref(x_251); + x_254 = lean_box(0); } -if (lean_is_scalar(x_251)) { - x_252 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_254)) { + x_255 = lean_alloc_ctor(1, 2, 0); } else { - x_252 = x_251; + x_255 = x_254; } -lean_ctor_set(x_252, 0, x_249); -lean_ctor_set(x_252, 1, x_250); -return x_252; +lean_ctor_set(x_255, 0, x_252); +lean_ctor_set(x_255, 1, x_253); +return x_255; } } } } case 7: { -lean_object* x_253; lean_object* x_254; uint8_t x_255; -lean_dec(x_150); +lean_object* x_256; lean_object* x_257; uint8_t x_258; +lean_dec(x_153); +x_256 = lean_ctor_get(x_155, 1); +lean_inc(x_256); +x_257 = lean_ctor_get(x_155, 2); +lean_inc(x_257); +lean_dec(x_155); +x_258 = l_Lean_Expr_hasLooseBVars(x_257); +if (x_258 == 0) +{ +lean_object* x_259; +x_259 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getKeyArgs___lambda__2(x_1, x_256, x_257, x_5, x_6, x_7, x_8, x_154); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_253 = lean_ctor_get(x_152, 1); -lean_inc(x_253); -x_254 = lean_ctor_get(x_152, 2); -lean_inc(x_254); -lean_dec(x_152); -x_255 = l_Lean_Expr_hasLooseBVars(x_254); -if (x_255 == 0) -{ -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_256 = lean_box(5); -x_257 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getKeyArgs___closed__2; -x_258 = lean_array_push(x_257, x_253); -x_259 = lean_array_push(x_258, x_254); -x_260 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_260, 0, x_256); -lean_ctor_set(x_260, 1, 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_151); -return x_261; -} -else -{ -lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; -lean_dec(x_254); -lean_dec(x_253); -x_262 = lean_box(4); -x_263 = l_Lean_Meta_DiscrTree_instInhabitedTrie___rarg___closed__1; -x_264 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_264, 0, x_262); -lean_ctor_set(x_264, 1, x_263); -x_265 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_265, 0, x_264); -lean_ctor_set(x_265, 1, x_151); -return x_265; +return x_259; } +else +{ +lean_object* x_260; lean_object* x_261; lean_object* x_262; +x_260 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_elimLooseBVarsByBeta___closed__1; +x_261 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_elimLooseBVarsByBeta___closed__2; +lean_inc(x_8); +lean_inc(x_7); +x_262 = l_Lean_Core_transform___at_Lean_Core_betaReduce___spec__1(x_257, x_260, x_261, x_7, x_8, x_154); +if (lean_obj_tag(x_262) == 0) +{ +lean_object* x_263; lean_object* x_264; lean_object* x_265; +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); +x_265 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getKeyArgs___lambda__2(x_1, x_256, x_263, x_5, x_6, x_7, x_8, x_264); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +return x_265; } -case 9: +else { -lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; -lean_dec(x_150); +lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; +lean_dec(x_256); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_266 = lean_ctor_get(x_152, 0); +x_266 = lean_ctor_get(x_262, 0); lean_inc(x_266); -lean_dec(x_152); -x_267 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_267, 0, x_266); -x_268 = l_Lean_Meta_DiscrTree_instInhabitedTrie___rarg___closed__1; -x_269 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_269, 0, x_267); -lean_ctor_set(x_269, 1, x_268); -x_270 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_270, 0, x_269); -lean_ctor_set(x_270, 1, x_151); -return x_270; +x_267 = lean_ctor_get(x_262, 1); +lean_inc(x_267); +if (lean_is_exclusive(x_262)) { + lean_ctor_release(x_262, 0); + lean_ctor_release(x_262, 1); + x_268 = x_262; +} else { + lean_dec_ref(x_262); + x_268 = lean_box(0); } -case 11: +if (lean_is_scalar(x_268)) { + x_269 = lean_alloc_ctor(1, 2, 0); +} else { + x_269 = x_268; +} +lean_ctor_set(x_269, 0, x_266); +lean_ctor_set(x_269, 1, x_267); +return x_269; +} +} +} +case 9: { -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_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; +lean_dec(x_153); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_271 = lean_ctor_get(x_152, 0); -lean_inc(x_271); -x_272 = lean_ctor_get(x_152, 1); -lean_inc(x_272); -x_273 = lean_ctor_get(x_152, 2); -lean_inc(x_273); -lean_dec(x_152); -x_274 = lean_unsigned_to_nat(0u); -x_275 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_150, x_274); -lean_inc(x_275); -x_276 = lean_alloc_ctor(6, 3, 0); -lean_ctor_set(x_276, 0, x_271); -lean_ctor_set(x_276, 1, x_272); -lean_ctor_set(x_276, 2, x_275); -x_277 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_createNodes___closed__1; -x_278 = lean_array_push(x_277, x_273); -x_279 = lean_mk_empty_array_with_capacity(x_275); -lean_dec(x_275); -x_280 = l___private_Lean_Expr_0__Lean_Expr_getAppRevArgsAux(x_150, x_279); -x_281 = l_Array_append___rarg(x_278, x_280); -x_282 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_282, 0, x_276); -lean_ctor_set(x_282, 1, x_281); -x_283 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_283, 0, x_282); -lean_ctor_set(x_283, 1, x_151); -return x_283; +x_270 = lean_ctor_get(x_155, 0); +lean_inc(x_270); +lean_dec(x_155); +x_271 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_271, 0, x_270); +x_272 = l_Lean_Meta_DiscrTree_instInhabitedTrie___rarg___closed__1; +x_273 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_273, 0, x_271); +lean_ctor_set(x_273, 1, 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; } -default: +case 11: { -lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; -lean_dec(x_152); -lean_dec(x_150); +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_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_284 = lean_box(4); -x_285 = l_Lean_Meta_DiscrTree_instInhabitedTrie___rarg___closed__1; +x_275 = lean_ctor_get(x_155, 0); +lean_inc(x_275); +x_276 = lean_ctor_get(x_155, 1); +lean_inc(x_276); +x_277 = lean_ctor_get(x_155, 2); +lean_inc(x_277); +lean_dec(x_155); +x_278 = lean_unsigned_to_nat(0u); +x_279 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_153, x_278); +lean_inc(x_279); +x_280 = lean_alloc_ctor(6, 3, 0); +lean_ctor_set(x_280, 0, x_275); +lean_ctor_set(x_280, 1, x_276); +lean_ctor_set(x_280, 2, x_279); +x_281 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_createNodes___closed__1; +x_282 = lean_array_push(x_281, x_277); +x_283 = lean_mk_empty_array_with_capacity(x_279); +lean_dec(x_279); +x_284 = l___private_Lean_Expr_0__Lean_Expr_getAppRevArgsAux(x_153, x_283); +x_285 = l_Array_append___rarg(x_282, x_284); x_286 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_286, 0, x_284); +lean_ctor_set(x_286, 0, x_280); lean_ctor_set(x_286, 1, x_285); x_287 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_287, 0, x_286); -lean_ctor_set(x_287, 1, x_151); +lean_ctor_set(x_287, 1, x_154); return x_287; } +default: +{ +lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; +lean_dec(x_155); +lean_dec(x_153); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +x_288 = lean_box(4); +x_289 = l_Lean_Meta_DiscrTree_instInhabitedTrie___rarg___closed__1; +x_290 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_290, 0, x_288); +lean_ctor_set(x_290, 1, x_289); +x_291 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_291, 0, x_290); +lean_ctor_set(x_291, 1, x_154); +return x_291; +} } } } else { -uint8_t x_288; +uint8_t x_292; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_288 = !lean_is_exclusive(x_10); -if (x_288 == 0) +x_292 = !lean_is_exclusive(x_10); +if (x_292 == 0) { return x_10; } else { -lean_object* x_289; lean_object* x_290; lean_object* x_291; -x_289 = lean_ctor_get(x_10, 0); -x_290 = lean_ctor_get(x_10, 1); -lean_inc(x_290); -lean_inc(x_289); +lean_object* x_293; lean_object* x_294; lean_object* x_295; +x_293 = lean_ctor_get(x_10, 0); +x_294 = lean_ctor_get(x_10, 1); +lean_inc(x_294); +lean_inc(x_293); lean_dec(x_10); -x_291 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_291, 0, x_289); -lean_ctor_set(x_291, 1, x_290); -return x_291; +x_295 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_295, 0, x_293); +lean_ctor_set(x_295, 1, x_294); +return x_295; } } } @@ -10789,6 +11130,20 @@ lean_dec(x_4); return x_11; } } +LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getKeyArgs___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) { +_start: +{ +uint8_t x_9; lean_object* x_10; +x_9 = lean_unbox(x_1); +lean_dec(x_1); +x_10 = l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getKeyArgs___lambda__2(x_9, 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); +return x_10; +} +} LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getKeyArgs___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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: { @@ -22746,6 +23101,7 @@ return x_10; } lean_object* initialize_Init(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_WHNF(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Transform(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_DiscrTreeTypes(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Meta_DiscrTree(uint8_t builtin, lean_object* w) { @@ -22758,6 +23114,9 @@ lean_dec_ref(res); res = initialize_Lean_Meta_WHNF(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Lean_Meta_Transform(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); res = initialize_Lean_Meta_DiscrTreeTypes(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); @@ -22869,6 +23228,12 @@ l_Lean_Meta_DiscrTree_mkNoindexAnnotation___closed__1 = _init_l_Lean_Meta_DiscrT lean_mark_persistent(l_Lean_Meta_DiscrTree_mkNoindexAnnotation___closed__1); l_Lean_Meta_DiscrTree_mkNoindexAnnotation___closed__2 = _init_l_Lean_Meta_DiscrTree_mkNoindexAnnotation___closed__2(); lean_mark_persistent(l_Lean_Meta_DiscrTree_mkNoindexAnnotation___closed__2); +l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_elimLooseBVarsByBeta___lambda__1___closed__1 = _init_l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_elimLooseBVarsByBeta___lambda__1___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_elimLooseBVarsByBeta___lambda__1___closed__1); +l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_elimLooseBVarsByBeta___closed__1 = _init_l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_elimLooseBVarsByBeta___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_elimLooseBVarsByBeta___closed__1); +l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_elimLooseBVarsByBeta___closed__2 = _init_l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_elimLooseBVarsByBeta___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_elimLooseBVarsByBeta___closed__2); l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_initCapacity = _init_l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_initCapacity(); lean_mark_persistent(l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_initCapacity); l_Lean_Meta_DiscrTree_mkPath___closed__1 = _init_l_Lean_Meta_DiscrTree_mkPath___closed__1(); @@ -22887,10 +23252,10 @@ l_Lean_Meta_DiscrTree_insertCore___rarg___closed__3 = _init_l_Lean_Meta_DiscrTre lean_mark_persistent(l_Lean_Meta_DiscrTree_insertCore___rarg___closed__3); l_Lean_Meta_DiscrTree_insertCore___rarg___closed__4 = _init_l_Lean_Meta_DiscrTree_insertCore___rarg___closed__4(); lean_mark_persistent(l_Lean_Meta_DiscrTree_insertCore___rarg___closed__4); +l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getKeyArgs___lambda__2___closed__1 = _init_l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getKeyArgs___lambda__2___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getKeyArgs___lambda__2___closed__1); l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getKeyArgs___closed__1 = _init_l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getKeyArgs___closed__1(); lean_mark_persistent(l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getKeyArgs___closed__1); -l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getKeyArgs___closed__2 = _init_l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getKeyArgs___closed__2(); -lean_mark_persistent(l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_getKeyArgs___closed__2); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lean/Meta/Eqns.c b/stage0/stdlib/Lean/Meta/Eqns.c index 52fb374a9bb9..cf0970cde296 100644 --- a/stage0/stdlib/Lean/Meta/Eqns.c +++ b/stage0/stdlib/Lean/Meta/Eqns.c @@ -52,7 +52,6 @@ extern lean_object* l_Lean_Expr_instBEqExpr; lean_object* l_instBEqProd___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Eqns___hyg_16_(lean_object*); lean_object* l_List_mapTR_loop___at_Lean_mkConstWithLevelParams___spec__1(lean_object*, lean_object*); -lean_object* l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_initializing(lean_object*); static lean_object* l_List_forIn_loop___at_Lean_Meta_getEqnsFor_x3f___spec__8___closed__10; lean_object* l_instHashableArray___rarg___boxed(lean_object*, lean_object*); @@ -102,6 +101,7 @@ static lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_getEqnsFor LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___at___private_Lean_Meta_Eqns_0__Lean_Meta_mkSimpleEqThm___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_registerGetEqnsFn___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Eqns___hyg_1353_(lean_object*); +lean_object* l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_registerGetEqnsFn___lambda__1___closed__1; static lean_object* l_List_forIn_loop___at_Lean_Meta_getEqnsFor_x3f___spec__8___closed__16; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_getEqnsFor_x3f___spec__5(lean_object*, size_t, size_t, lean_object*, lean_object*); @@ -3808,7 +3808,7 @@ lean_closure_set(x_10, 1, x_9); lean_closure_set(x_10, 2, x_8); x_11 = l_Lean_Meta_getEqnsFor_x3f___closed__5; x_12 = l_Lean_Meta_getEqnsFor_x3f___closed__6; -x_13 = l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__3___rarg(x_11, x_12, x_10, x_3, x_4, x_5, x_6, x_7); +x_13 = l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2___rarg(x_11, x_12, x_10, x_3, x_4, x_5, x_6, x_7); return x_13; } } @@ -4554,7 +4554,7 @@ lean_closure_set(x_9, 0, x_1); lean_closure_set(x_9, 1, x_8); x_10 = l_Lean_Meta_getEqnsFor_x3f___closed__5; x_11 = l_Lean_Meta_getEqnsFor_x3f___closed__6; -x_12 = l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__3___rarg(x_10, x_11, x_9, x_3, x_4, x_5, x_6, x_7); +x_12 = l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2___rarg(x_10, x_11, x_9, x_3, x_4, x_5, x_6, x_7); return x_12; } } diff --git a/stage0/stdlib/Lean/Meta/Eval.c b/stage0/stdlib/Lean/Meta/Eval.c index a7f511e58ff0..13cfccb98685 100644 --- a/stage0/stdlib/Lean/Meta/Eval.c +++ b/stage0/stdlib/Lean/Meta/Eval.c @@ -35,7 +35,6 @@ static lean_object* l_Lean_Meta_evalExpr___rarg___lambda__1___closed__2; lean_object* lean_st_ref_get(lean_object*, lean_object*); static lean_object* l_Lean_Meta_evalExpr_x27___rarg___lambda__1___closed__1; lean_object* l_Lean_setEnv___at_Lean_Meta_setInlineAttribute___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_evalExpr_x27___rarg___boxed(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*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); @@ -45,6 +44,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_evalExprCore(lean_object*); static lean_object* l_Lean_Meta_evalExprCore___rarg___closed__1; static lean_object* l_Lean_Meta_evalExprCore___rarg___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_evalExpr_x27___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_indentExpr(lean_object*); LEAN_EXPORT lean_object* l_Lean_evalConst___at_Lean_Meta_evalExprCore___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_evalExprCore___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -403,7 +403,7 @@ lean_inc(x_15); x_16 = lean_ctor_get(x_14, 1); lean_inc(x_16); lean_dec(x_14); -x_17 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_1, x_4, x_5, x_6, x_7, x_16); +x_17 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_1, x_4, x_5, x_6, x_7, x_16); x_18 = lean_ctor_get(x_17, 0); lean_inc(x_18); x_19 = lean_ctor_get(x_17, 1); diff --git a/stage0/stdlib/Lean/Meta/ExprDefEq.c b/stage0/stdlib/Lean/Meta/ExprDefEq.c index 2200a3959bf2..65050867ae32 100644 --- a/stage0/stdlib/Lean/Meta/ExprDefEq.c +++ b/stage0/stdlib/Lean/Meta/ExprDefEq.c @@ -19,7 +19,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkCacheKey extern lean_object* l_Lean_Core_instMonadTraceCoreM; lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_unstuckMVar___closed__8; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__14; lean_object* l_Lean_addTrace___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_collectLetDeps___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_is_matcher(lean_object*, lean_object*); @@ -29,12 +28,12 @@ lean_object* l___private_Init_Util_0__outOfBounds___rarg(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_Meta_CheckAssignment_checkMVar___spec__4___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_getCachedResult___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___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_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__11; LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_collectLetDeclsFrom_visit___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_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__2___closed__3; static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqQuickOther___lambda__2___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqQuickOther___lambda__2(uint8_t, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_ReaderT_instMonadFunctorReaderT(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__10; LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Meta_CheckAssignment_checkFVar___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___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_Meta_CheckAssignment_checkMVar___spec__13___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_cacheResult___spec__1___boxed(lean_object*, lean_object*, lean_object*); @@ -49,31 +48,31 @@ LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDef static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processConstApprox_go___closed__1; static lean_object* l_Lean_Meta_isExprDefEqAuxImpl___lambda__3___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_profiler; 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_isExprDefEqAuxImpl___lambda__3(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_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_whenUndefDo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099_(lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); lean_object* l_Lean_mkAppN(lean_object*, lean_object*); lean_object* lean_is_level_def_eq(lean_object*, lean_object*, 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_Meta_isAbstractedUnassignedMVar___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_cacheResult___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___lambda__2___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqBindingAux(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_Lean_PersistentArray_anyM___at_Lean_Meta_CheckAssignment_checkMVar___spec__14(lean_object*, lean_object*); static lean_object* l_Lean_Meta_CheckAssignment_checkMVar___closed__2; LEAN_EXPORT lean_object* l_Lean_PersistentArray_foldlM___at_Lean_Meta_CheckAssignment_checkMVar___spec__52___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_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___lambda__1___closed__6; LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqProj_isDefEqSingleton___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_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__11; lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashSetImp_contains___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_collectLetDepsAux___spec__1___boxed(lean_object*, lean_object*); lean_object* l_Lean_getProjectionFnInfo_x3f___at_Lean_Meta_getStuckMVar_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_Meta_CheckAssignment_checkMVar___spec__44(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__9; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_Meta_CheckAssignment_checkMVar___spec__27___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqMVarSelf(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_hasLetDeclsInBetween___lambda__1___boxed(lean_object*); -static lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___closed__1; lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_shouldVisit(lean_object*, lean_object*); static lean_object* l_Lean_Meta_CheckAssignment_checkFVar___closed__1; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_Meta_CheckAssignment_checkMVar___spec__11___boxed(lean_object*, lean_object*, lean_object*); @@ -84,6 +83,7 @@ lean_object* l_Lean_Meta_mkFreshExprMVarAt(lean_object*, lean_object*, lean_obje LEAN_EXPORT lean_object* l_Lean_Meta_CheckAssignmentQuick_check_visit___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_instMonadTrace___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Meta_CheckAssignment_checkMVar___spec__50___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_occursCheck_visitMVar___at_Lean_Meta_checkAssignment___spec__3(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*); static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqLeft___closed__2; @@ -125,6 +125,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processAss LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_CheckAssignment_checkMVar___spec__10(lean_object*, lean_object*, size_t, size_t); uint8_t l_Lean_Expr_isApp(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processConstApprox_defaultCase___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +double l_Lean_trace_profiler_threshold_getSecs(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqQuickOther___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_CheckAssignment_checkMVar___spec__2(lean_object*, lean_object*, size_t, size_t); @@ -136,6 +137,7 @@ lean_object* l_Lean_MessageData_ofList(lean_object*); lean_object* l_Lean_PersistentArray_push___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at_Lean_Meta_CheckAssignment_checkMVar___spec__65(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_findDecl_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6556_(lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); lean_object* l_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_hasLetDeclsInBetween___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -171,7 +173,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_isDefEqStringLit(lean_object*, lean_object* LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Meta_checkAssignment___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_withProofIrrelTransparency(lean_object*); lean_object* l_Lean_exceptBoolEmoji___rarg(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6071_(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processConstApprox_go(lean_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_isDefEqNat(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_replaceRef(lean_object*, lean_object*); @@ -183,14 +184,12 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_expandDela LEAN_EXPORT lean_object* l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqApp___spec__2(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___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isAssignable(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqQuickOther___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_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6051_(lean_object*); lean_object* l_Lean_Expr_proj___override(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_bvar___override(lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgs___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_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at_Lean_Meta_CheckAssignment_checkMVar___spec__31(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_addLetDeps___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*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__10; lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_isExprDefEqAuxImpl___lambda__3___closed__15; LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isExprDefEqExpensive(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -206,6 +205,7 @@ LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at_Lean_Meta_CheckAssignmen static lean_object* l_Lean_Meta_isExprDefEqAuxImpl___lambda__3___closed__17; LEAN_EXPORT lean_object* l_Lean_Meta_CheckAssignment_checkFVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_environment_find(lean_object*, lean_object*); +uint8_t lean_float_decLt(double, double); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Meta_CheckAssignment_checkMVar___spec__18___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processConstApprox_defaultCase(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_CheckAssignmentQuick_check_visit(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -215,7 +215,6 @@ static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_unstuckMVar___c lean_object* l_Lean_Meta_throwIsDefEqStuck___rarg(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_cacheResult___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_ConstantInfo_hasValue(lean_object*); -static lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqDelta(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isAppOf(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEta(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -227,6 +226,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_Ex LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_tryHeuristic___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_Meta_ExprDefEq_0__Lean_Meta_isDefEqQuickOther___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_forallBoundedTelescope___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__2___rarg(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_ExprDefEq___hyg_20099____closed__6; LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_assignConst___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___private_Lean_Meta_ExprDefEq_0__Lean_Meta_unstuckMVar___closed__14; LEAN_EXPORT lean_object* l_Lean_LocalContext_foldlM___at_Lean_Meta_CheckAssignment_checkMVar___spec__51(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -236,19 +236,22 @@ lean_object* l_Lean_stringToMessageData(lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___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* l_Lean_Expr_mvar___override(lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Meta_CheckAssignment_checkMVar___spec__33___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__9; LEAN_EXPORT lean_object* l_Lean_HashSetImp_moveEntries___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_collectLetDeclsFrom_visit___spec__12(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l___private_Lean_Util_Trace_0__Lean_checkTraceOption(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_instInhabitedPersistentArrayNode(lean_object*); +static lean_object* l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6576____closed__1; static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_CheckAssignment_addAssignmentInfo___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_assignConst___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processAssignmentFOApproxAux(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__7; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_checkAssignment___spec__6(lean_object*, 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_isExprDefEqAuxImpl___lambda__3___closed__10; extern lean_object* l_Lean_Expr_instBEqExpr; static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_hasLetDeclsInBetween___closed__1; +static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__5; lean_object* l_instBEqProd___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_unfoldNonProjFnDefEq_packedInstanceOf_x3f___lambda__1(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_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_hasLetDeclsInBetween___spec__2(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_Meta_ExprDefEq_0__Lean_Meta_processAssignment_process___spec__1(lean_object*, lean_object*, size_t, size_t); @@ -268,6 +271,7 @@ static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_assignConst___l lean_object* l_Lean_Meta_getNumPostponed___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_unstuckMVar___closed__7; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_CheckAssignment_checkMVar___spec__64(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_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__9; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_CheckAssignment_checkMVar___spec__56___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_addTrace___at_Lean_Meta_CheckAssignment_checkFVar___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_Array_foldlMUnsafe_fold___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_collectLetDeps___spec__2(lean_object*, size_t, size_t, lean_object*); @@ -282,13 +286,13 @@ LEAN_EXPORT uint8_t l_Lean_Meta_CheckAssignmentQuick_check(uint8_t, lean_object* uint8_t l_Lean_Expr_hasMVar(lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_CheckAssignment_checkMVar___spec__49(lean_object*, lean_object*, size_t, size_t); static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqRight___closed__2; +LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___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___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processAssignment(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_simpAssignmentArgAux___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMapImp_moveEntries___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_collectLetDeclsFrom_visit___spec__4(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isLambda(lean_object*); static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_unstuckMVar___closed__3; lean_object* l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__4; LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_collectLetDeclsFrom_visit___spec__5(lean_object*, lean_object*); static lean_object* l_Lean_Meta_isExprDefEqAuxImpl___lambda__3___closed__9; LEAN_EXPORT lean_object* l_ReaderT_pure___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processAssignment___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -319,6 +323,7 @@ lean_object* l___private_Lean_Expr_0__Lean_mkAppRangeAux(lean_object*, lean_obje LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_hasLetDeclsInBetween___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_CheckAssignmentQuick_check_visit___spec__1(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isExprDefEqExpensive___lambda__2(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_ExprDefEq___hyg_20099____closed__4; static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__2___closed__7; static size_t l_Lean_PersistentHashMap_findAux___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_getCachedResult___spec__2___closed__2; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgs___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*); @@ -334,10 +339,10 @@ static lean_object* l_Lean_Meta_CheckAssignment_throwOutOfScopeFVar___rarg___clo LEAN_EXPORT lean_object* l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqQuickMVarMVar___spec__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_local_ctx_find(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_unfoldComparingHeadsDefEq___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_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6556____closed__2; lean_object* lean_st_ref_take(lean_object*, lean_object*); lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_processPostponed___spec__2___rarg(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___lambda__2___closed__2; -static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDeltaCandidate_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___closed__2; static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_unstuckMVar___closed__4; @@ -353,6 +358,7 @@ LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_v LEAN_EXPORT lean_object* l_Lean_HashMap_insert___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_collectLetDeclsFrom_visit___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_assignConst(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_expr_eqv(lean_object*, lean_object*); +uint8_t l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_12510_(uint8_t, uint8_t); static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processConstApprox_go___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_CheckAssignment_check(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_collectLetDeclsFrom_visit___spec__8___boxed(lean_object*, lean_object*); @@ -363,30 +369,28 @@ lean_object* l_Lean_instantiateMVars___at___private_Lean_MetavarContext_0__Lean_ LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at_Lean_Meta_CheckAssignment_checkMVar___spec__6___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_contains___at_Lean_Meta_CheckAssignment_checkFVar___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at_Lean_Meta_CheckAssignment_checkMVar___spec__14___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6071____closed__1; LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyM___at_Lean_Meta_CheckAssignment_checkMVar___spec__22(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_collectLetDepsAux(lean_object*, 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_mkHashMap___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_collectLetDeclsFrom___spec__2(lean_object*); lean_object* l_List_mapTR_loop___at_Lean_MessageData_instCoeListExprMessageData___spec__1(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6051____closed__2; static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_unstuckMVar___closed__11; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___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_anyMUnsafe_any___at_Lean_Meta_CheckAssignment_checkFVar___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_HashMapImp_find_x3f___at_Lean_instantiateExprMVars___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_expandDelayedAssigned_x3f___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___private_Lean_Meta_ExprDefEq_0__Lean_Meta_unstuckMVar___closed__18; -static lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___closed__5; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processAssignment_process___spec__2(lean_object*, lean_object*, size_t, size_t); static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___closed__1; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__13; +LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__4(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* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean_object*); lean_object* l_Lean_registerInternalExceptionId(lean_object*, lean_object*); lean_object* l_Lean_Meta_SavedState_restore(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_CheckAssignment_throwOutOfScopeFVar___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at_Lean_Meta_CheckAssignment_checkMVar___spec__7(lean_object*, lean_object*); -static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___closed__4; +static lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Lean_HashSetImp_expand___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_collectLetDeclsFrom_visit___spec__11(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_CheckAssignment_throwCheckAssignmentFailure___rarg(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processConstApprox___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_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -416,15 +420,16 @@ static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqQuickOth LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_FVarId_getDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_forallE___override(lean_object*, lean_object*, lean_object*, uint8_t); +static lean_object* l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6556____closed__1; lean_object* l_ReaderT_instMonadLiftReaderT(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Meta_CheckAssignment_checkMVar___spec__40___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_expandDelayedAssigned_x3f___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_anyMUnsafe_any___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processAssignment_process___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__10; lean_object* lean_st_ref_get(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___lambda__1___closed__4; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_CheckAssignment_checkMVar___spec__58(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_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__2___closed__2; -static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__5; LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_expandDelayedAssigned_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_array_pop(lean_object*); static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___closed__3; @@ -441,7 +446,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_unfoldNonP lean_object* lean_array_to_list(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_unstuckMVar___closed__2; static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___closed__2; -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_LocalContext_containsFVar(lean_object*, lean_object*); uint8_t l_Lean_Expr_hasLooseBVars(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_getCachedResult___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -458,6 +462,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_Ex LEAN_EXPORT lean_object* l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_tryHeuristic___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Expr_instHashableExpr; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_CheckAssignmentQuick_check_visit___spec__2(lean_object*, lean_object*, size_t, size_t); +static lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__5; extern lean_object* l_Lean_levelZero; LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_cacheResult(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_isDelayedAssigned___at_Lean_Meta_CheckAssignment_checkApp___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -487,26 +492,25 @@ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at___private_Lean_ LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgs___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* l_Lean_Expr_letE___override(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); lean_object* l_instHashableProd___rarg___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__14; LEAN_EXPORT lean_object* l_Lean_Meta_isDefEqBindingDomain(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_CheckAssignmentQuick_check___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_unstuckMVar___closed__13; static lean_object* l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__2___closed__1; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__8; LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at_Lean_Meta_CheckAssignment_checkMVar___spec__39(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___closed__1; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__6; LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqOnFailure(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_collectLetDeclsFrom(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_occursCheck_visit___at_Lean_Meta_checkAssignment___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_Meta_CheckAssignment_State_cache___default; +uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_Meta_CheckAssignment_checkMVar___spec__35(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_CheckAssignment_throwCheckAssignmentFailure___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_checkAssignment___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processConstApprox_go_cont___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_unfold___rarg(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_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__8; lean_object* l_Lean_MetavarContext_getExprAssignmentCore_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_CheckAssignment_checkApp___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_foldlMUnsafe_fold___at_Lean_Meta_CheckAssignment_checkMVar___spec__60___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*); @@ -516,6 +520,7 @@ LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_Meta_CheckAssi static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqQuickOther___closed__4; uint8_t l_Lean_LocalDecl_isLet(lean_object*); uint8_t l_Lean_Meta_ParamInfo_isInstImplicit(lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__2; lean_object* l_Lean_Meta_saveState___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at_Lean_Meta_CheckAssignment_checkMVar___spec__22___boxed(lean_object*, lean_object*); @@ -525,6 +530,7 @@ lean_object* l_Lean_LocalDecl_fvarId(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_CheckAssignment_checkMVar___spec__64___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_ExprDefEq_0__Lean_Meta_isDefEqProofIrrel(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_ReaderT_pure___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processAssignment___spec__1(lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__12; LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___lambda__3(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_CheckAssignment_checkMVar___spec__48(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -537,6 +543,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_unstuckMVa LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_collectLetDeclsFrom_visit___spec__8(lean_object*, lean_object*); extern lean_object* l_Lean_Meta_instMonadMetaM; LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_unstuckMVar___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_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__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_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqUnitLike(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_tryHeuristic___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -546,7 +553,6 @@ LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_Exp static lean_object* l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__2___closed__3; static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processAssignmentFOApprox_loop___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_unstuckMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___closed__6; static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqQuickOther___lambda__2___closed__2; static lean_object* l_Lean_Meta_isExprDefEqAuxImpl___lambda__3___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_unfoldComparingHeadsDefEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -554,13 +560,14 @@ LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDef LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processAssignment_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_Meta_isExprDefEqAuxImpl___lambda__3___closed__7; lean_object* lean_usize_to_nat(size_t); -static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__3; size_t lean_data_hashmap_mk_idx(lean_object*, uint64_t); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_hasLetDeclsInBetween(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_collectLetDeclsFrom___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_isExprDefEqAuxImpl___lambda__3___closed__5; lean_object* l_Lean_LocalContext_mkLocalDecl(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t); -static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__7; +static lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__2; +lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_processPostponed___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___lambda__2(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_inheritedTraceOptions; LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqProjInst(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofExpr(lean_object*); @@ -569,10 +576,10 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_tryHeurist static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_hasLetDeclsInBetween___spec__1___closed__1; lean_object* l_Lean_Meta_getConfig(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processAssignmentFOApprox_loop___spec__1(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__2; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_Meta_CheckAssignment_checkMVar___spec__20___boxed(lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Core_instMonadQuotationCoreM; lean_object* l_Lean_Expr_toCtorIfLit(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6576_(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_Meta_CheckAssignment_checkMVar___spec__19___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_CheckAssignment_cache(lean_object*, 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_go___spec__3(uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -585,6 +592,7 @@ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at___private_Lean_Me static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___closed__2; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Meta_CheckAssignment_checkMVar___spec__49___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqRight___closed__1; +static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__7; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgs___spec__2___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___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__1___closed__3; lean_object* l_Lean_Meta_mkProjFn(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -594,13 +602,11 @@ static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqQuickOth static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___closed__2; LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyM___at_Lean_Meta_CheckAssignment_checkMVar___spec__30(lean_object*, lean_object*); static lean_object* l_Lean_Meta_CheckAssignment_State_cache___default___closed__1; -lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Meta_processPostponed___spec__7(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_Lean_Meta_isDefEqBindingDomain_loop(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at_Lean_Meta_CheckAssignment_checkMVar___spec__38___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlMAux___at_Lean_Meta_CheckAssignment_checkMVar___spec__61___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_Meta_CheckAssignment_assignToConstFun___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_processPostponed___spec__8(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_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_Meta_CheckAssignment_checkMVar___spec__45(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_LocalContext_isSubPrefixOf(lean_object*, lean_object*, lean_object*); @@ -611,21 +617,26 @@ LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_CheckAssignment_checkM LEAN_EXPORT uint8_t l_List_elem___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_collectLetDeclsFrom_visit___spec__10(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_Meta_CheckAssignment_checkMVar___spec__11(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass(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_ExprDefEq___hyg_20099____closed__7; lean_object* l_Lean_Expr_constLevels_x21(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_isExprDefEqAuxImpl___lambda__2(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqLeft___closed__3; LEAN_EXPORT lean_object* l_Lean_HashSetImp_insert___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_collectLetDeclsFrom_visit___spec__9(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_addLetDeps___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_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__3; +static lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__6; LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqBinding___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_simpAssignmentArg___closed__1; +lean_object* l_Lean_withSeconds___at_Lean_Meta_processPostponed___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlMAux___at_Lean_Meta_CheckAssignment_checkMVar___spec__54(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_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_collectLetDepsAux___boxed(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_Meta_CheckAssignment_checkApp___spec__3___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_ReaderT_pure___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processAssignment___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__9; lean_object* l_Lean_Expr_app___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_getCachedResult___spec__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_isExprDefEqAuxImpl___lambda__3___closed__13; -static lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___closed__4; +static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__2; static lean_object* l_Lean_Meta_CheckAssignment_throwCheckAssignmentFailure___rarg___closed__2; LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_CheckAssignment_checkApp___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_EXPORT lean_object* l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__2___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -646,21 +657,25 @@ lean_object* l_Lean_MVarId_isAssigned___at___private_Lean_Meta_SynthInstance_0__ LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Meta_CheckAssignment_checkMVar___spec__48___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_float_to_string(double); static lean_object* l_Lean_occursCheck_visitMVar___at_Lean_Meta_checkAssignment___spec__3___closed__1; +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_tryUnificationHints(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__5; LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_unstuckMVar___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqOnFailure___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* l_Lean_PersistentHashMap_insert___at_Lean_MVarId_assign___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_unstuckMVar___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqOnFailure___spec__2___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_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__13; lean_object* l___private_Lean_Expr_0__Lean_Expr_etaExpandedAux(lean_object*, lean_object*); +static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_CheckAssignment_throwOutOfScopeFVar___rarg(lean_object*); static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_hasLetDeclsInBetween___closed__2; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_CheckAssignment_checkFVar___spec__2(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Lean_mkHashSet___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_collectLetDeps___spec__1___boxed(lean_object*); +static lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__4___closed__1; uint8_t l___private_Lean_Expr_0__Lean_beqBinderInfo____x40_Lean_Expr___hyg_473_(uint8_t, uint8_t); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_simpAssignmentArgAux(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at_Lean_Meta_CheckAssignment_checkMVar___spec__53___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_anyMUnsafe_any___at_Lean_Meta_CheckAssignmentQuick_check_visit___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__2; lean_object* l_Lean_indentExpr(lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_etaEq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_unfoldNonProjFnDefEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -668,11 +683,11 @@ static lean_object* l_Lean_Meta_isExprDefEqAuxImpl___lambda__3___closed__12; lean_object* l_Lean_PersistentHashMap_mkEmptyEntries(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_CheckAssignment_checkFVar___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Core_checkMaxHeartbeats(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___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___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqBindingAux___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_foldlMUnsafe_fold___at_Lean_Meta_CheckAssignment_checkMVar___spec__62___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_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_Meta_CheckAssignment_checkMVar___spec__35___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqLeft___lambda__1___boxed(lean_object*, lean_object*, 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_12336_(uint8_t, uint8_t); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_Meta_CheckAssignment_checkMVar___spec__28(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_unfoldDefinition_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqProj_isDefEqSingleton___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*); @@ -690,7 +705,7 @@ static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processAssignme lean_object* lean_array_set(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_isTracingEnabledFor___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyM___at_Lean_Meta_CheckAssignment_checkMVar___spec__46(lean_object*, lean_object*); -static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__10; +static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__8; lean_object* l_Lean_Meta_isProofQuick(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkFreshFVarId___at___private_Lean_Meta_Basic_0__Lean_Meta_forallTelescopeReducingAuxAux_process___spec__1___rarg(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*); @@ -700,7 +715,6 @@ lean_object* l_Lean_LocalDecl_type(lean_object*); static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___closed__3; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_CheckAssignment_checkMVar___spec__33(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Lean_getExprMVarAssignment_x3f___at_Lean_Meta_checkAssignment___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___closed__3; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_Meta_CheckAssignment_checkMVar___spec__4(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_foldlM___at_Lean_Meta_CheckAssignment_checkMVar___spec__52(lean_object*, 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_instMonadLiftT__1(lean_object*, lean_object*); @@ -712,7 +726,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_consumeLet LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqProj_isDefEqSingleton___lambda__4(lean_object*, lean_object*, 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_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__2___closed__6; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__12; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_Meta_CheckAssignment_checkMVar___spec__36___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_getAppFn(lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_CheckAssignment_checkMVar___spec__16(lean_object*, lean_object*, size_t, size_t); @@ -737,6 +750,7 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_CheckAssignmen LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_unfoldNonProjFnDefEq_packedInstanceOf_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_isExprDefEqAuxImpl___lambda__3___closed__16; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Meta_CheckAssignment_checkMVar___spec__17___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___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_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_replace___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_collectLetDeclsFrom_visit___spec__14___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); @@ -745,19 +759,20 @@ LEAN_EXPORT lean_object* l_Lean_Meta_CheckAssignment_checkAssignmentExceptionId; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Meta_CheckAssignment_checkMVar___spec__24___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_Meta_CheckAssignment_checkMVar___spec__36(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processAssignment___closed__1; -LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595_(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_CheckAssignment_checkMVar___spec__57(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* l_Array_back___rarg(lean_object*, lean_object*); +static lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__4; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___boxed(lean_object**); lean_object* lean_expr_instantiate_rev(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__3; lean_object* l_Lean_Meta_getFunInfoNArgs(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqApp___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___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processAssignment_process___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_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_unfoldBothDefEq(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_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__1; static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___lambda__1___closed__2; lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_getDelayedMVarAssignment_x3f___spec__1(lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_hasLetDeclsInBetween___spec__1___closed__2; +static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__6; size_t lean_usize_sub(size_t, size_t); static lean_object* l_Lean_Meta_isExprDefEqAuxImpl___lambda__3___closed__14; lean_object* l_Lean_Meta_isClass_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -786,20 +801,23 @@ static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAs LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processConstApprox___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_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_collectLetDeps(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_isExprDefEqAuxImpl___lambda__3___closed__8; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__1; +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__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_Lean_occursCheck___at_Lean_Meta_checkAssignment___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_isDefEqStringLit___closed__2; static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__2___closed__2; lean_object* lean_array_uget(lean_object*, size_t); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__7; LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isListLevelDefEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_fvar___override(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_isAbstractedUnassignedMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_unstuckMVar___closed__5; static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqOnFailure___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_unfoldDefEq___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___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_addLetDeps___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_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__8; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_CheckAssignment_checkMVar___spec__40(lean_object*, lean_object*, size_t, size_t); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__11; LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isAssigned___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_trace_profiler; static lean_object* l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__2___closed__5; LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqRight(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_ReducibilityHints_lt(lean_object*, lean_object*); @@ -807,7 +825,6 @@ LEAN_EXPORT lean_object* l_Lean_HashMapImp_find_x3f___at___private_Lean_Meta_Exp static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___lambda__1___closed__5; lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_CheckAssignment_checkMVar___spec__25(lean_object*, lean_object*, size_t, size_t); -lean_object* l_Lean_withOptProfile___at_Lean_Meta_processPostponed___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_isDefEqNative(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_CheckAssignment_checkMVar___closed__6; @@ -815,22 +832,23 @@ uint64_t l___private_Lean_Expr_0__Lean_hashFVarId____x40_Lean_Expr___hyg_1933_(l size_t lean_usize_shift_left(size_t, size_t); lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___closed__1; +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___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_Meta_checkAssignment(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_CheckAssignment_checkMVar(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_DependsOn_dep_visit___at_Lean_Meta_CheckAssignment_checkMVar___spec__19(lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(lean_object*, lean_object*, lean_object*); -static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__6; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_CheckAssignment_checkMVar___spec__42(lean_object*, lean_object*, size_t, size_t); lean_object* l_Lean_Meta_useEtaStruct(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__11; LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_Meta_CheckAssignment_checkMVar___spec__39___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Meta_CheckAssignment_instMonadCacheExprCheckAssignmentM___closed__3; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgs___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_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_CheckAssignment_checkMVar___spec__41(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_Meta_CheckAssignment_checkMVar___spec__5(lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Meta_processPostponed___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_unstuckMVar___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqQuickOther(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__5; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Meta_CheckAssignment_checkMVar___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_isExprDefEqAuxImpl___lambda__3___closed__20; lean_object* l_Lean_Expr_headBeta(lean_object*); @@ -840,9 +858,9 @@ lean_object* l_Lean_LocalDecl_toExpr(lean_object*); static lean_object* l_Lean_Meta_CheckAssignment_checkMVar___closed__3; static lean_object* l_Lean_Meta_CheckAssignment_checkMVar___closed__7; static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___closed__5; -static lean_object* l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6071____closed__2; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Meta_CheckAssignment_checkMVar___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgs___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_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6576____closed__2; lean_object* l_Lean_Meta_throwUnknownMVar___rarg(lean_object*, lean_object*, lean_object*, 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___private_Lean_Meta_ExprDefEq_0__Lean_Meta_trySynthPending(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -865,7 +883,6 @@ lean_object* l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isExprDefEqExpensive___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at_Lean_Meta_CheckAssignment_checkMVar___spec__30___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_isEtaUnassignedMVar___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6051____closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_cacheResult___spec__2(lean_object*, size_t, size_t, lean_object*, uint8_t); uint8_t l_Lean_Expr_isConst(lean_object*); lean_object* l_Lean_Expr_lam___override(lean_object*, lean_object*, lean_object*, uint8_t); @@ -906,7 +923,9 @@ uint8_t l_Lean_isStructureLike(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processAssignmentFOApprox_loop___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_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___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* l_Nat_repr(lean_object*); +lean_object* l_Lean_Meta_isProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_FVarId_getValue_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_CheckAssignment_throwOutOfScopeFVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_expandDelayedAssigned_x3f___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(lean_object*, lean_object*); @@ -932,6 +951,191 @@ LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_Meta_CheckAssi uint8_t l_Array_isEmpty___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Core_instMonadOptionsCoreM___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_isAbstractedUnassignedMVar(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: +{ +switch (lean_obj_tag(x_1)) { +case 2: +{ +lean_object* x_7; lean_object* x_8; +x_7 = lean_ctor_get(x_1, 0); +lean_inc(x_7); +lean_dec(x_1); +lean_inc(x_7); +x_8 = l_Lean_MVarId_isReadOnlyOrSyntheticOpaque(x_7, x_2, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_unbox(x_9); +lean_dec(x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_11 = lean_ctor_get(x_8, 1); +lean_inc(x_11); +lean_dec(x_8); +x_12 = l_Lean_MVarId_isAssigned___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__1(x_7, x_2, x_3, x_4, x_5, x_11); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_unbox(x_13); +lean_dec(x_13); +if (x_14 == 0) +{ +uint8_t x_15; +x_15 = !lean_is_exclusive(x_12); +if (x_15 == 0) +{ +lean_object* x_16; uint8_t x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_12, 0); +lean_dec(x_16); +x_17 = 1; +x_18 = lean_box(x_17); +lean_ctor_set(x_12, 0, x_18); +return x_12; +} +else +{ +lean_object* x_19; uint8_t x_20; lean_object* x_21; lean_object* x_22; +x_19 = lean_ctor_get(x_12, 1); +lean_inc(x_19); +lean_dec(x_12); +x_20 = 1; +x_21 = lean_box(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_19); +return x_22; +} +} +else +{ +uint8_t x_23; +x_23 = !lean_is_exclusive(x_12); +if (x_23 == 0) +{ +lean_object* x_24; uint8_t x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_12, 0); +lean_dec(x_24); +x_25 = 0; +x_26 = lean_box(x_25); +lean_ctor_set(x_12, 0, x_26); +return x_12; +} +else +{ +lean_object* x_27; uint8_t x_28; lean_object* x_29; lean_object* x_30; +x_27 = lean_ctor_get(x_12, 1); +lean_inc(x_27); +lean_dec(x_12); +x_28 = 0; +x_29 = lean_box(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_27); +return x_30; +} +} +} +else +{ +uint8_t x_31; +lean_dec(x_7); +x_31 = !lean_is_exclusive(x_8); +if (x_31 == 0) +{ +lean_object* x_32; uint8_t x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_8, 0); +lean_dec(x_32); +x_33 = 0; +x_34 = lean_box(x_33); +lean_ctor_set(x_8, 0, x_34); +return x_8; +} +else +{ +lean_object* x_35; uint8_t x_36; lean_object* x_37; lean_object* x_38; +x_35 = lean_ctor_get(x_8, 1); +lean_inc(x_35); +lean_dec(x_8); +x_36 = 0; +x_37 = lean_box(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_35); +return x_38; +} +} +} +else +{ +uint8_t x_39; +lean_dec(x_7); +x_39 = !lean_is_exclusive(x_8); +if (x_39 == 0) +{ +return x_8; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_8, 0); +x_41 = lean_ctor_get(x_8, 1); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_8); +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; +} +} +} +case 5: +{ +lean_object* x_43; +x_43 = lean_ctor_get(x_1, 0); +lean_inc(x_43); +lean_dec(x_1); +x_1 = x_43; +goto _start; +} +case 6: +{ +lean_object* x_45; +x_45 = lean_ctor_get(x_1, 2); +lean_inc(x_45); +lean_dec(x_1); +x_1 = x_45; +goto _start; +} +default: +{ +uint8_t x_47; lean_object* x_48; lean_object* x_49; +lean_dec(x_1); +x_47 = 0; +x_48 = lean_box(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_6); +return x_49; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_isAbstractedUnassignedMVar___boxed(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_Meta_isAbstractedUnassignedMVar(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); +return x_7; +} +} static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__1___closed__1() { _start: { @@ -1230,7 +1434,7 @@ return x_53; } } } -static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__1() { +static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__1() { _start: { lean_object* x_1; @@ -1238,7 +1442,7 @@ x_1 = lean_mk_string_from_bytes("Meta", 4); return x_1; } } -static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__2() { +static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__2() { _start: { lean_object* x_1; @@ -1246,7 +1450,7 @@ x_1 = lean_mk_string_from_bytes("isDefEq", 7); return x_1; } } -static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__3() { +static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__3() { _start: { lean_object* x_1; @@ -1254,7 +1458,7 @@ x_1 = lean_mk_string_from_bytes("eta", 3); return x_1; } } -static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__4() { +static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__4() { _start: { lean_object* x_1; @@ -1262,19 +1466,19 @@ x_1 = lean_mk_string_from_bytes("struct", 6); return x_1; } } -static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__5() { +static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__1; -x_2 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__2; -x_3 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__3; -x_4 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__4; +x_1 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__1; +x_2 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__2; +x_3 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__3; +x_4 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__4; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__6() { +static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__6() { _start: { lean_object* x_1; @@ -1282,16 +1486,16 @@ x_1 = lean_mk_string_from_bytes(" =?= ", 5); return x_1; } } -static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__7() { +static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__7() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__6; +x_1 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__6; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__8() { +static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__8() { _start: { lean_object* x_1; @@ -1299,16 +1503,16 @@ x_1 = lean_mk_string_from_bytes(" @ [", 4); return x_1; } } -static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__9() { +static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__9() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__8; +x_1 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__8; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__10() { +static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__10() { _start: { lean_object* x_1; @@ -1316,15 +1520,135 @@ x_1 = lean_mk_string_from_bytes("], ", 3); return x_1; } } -static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__11() { +static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__11() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__10; +x_1 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__10; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___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) { +_start: +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_dec(x_9); +x_15 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__5; +x_16 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_15, 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_unbox(x_17); +lean_dec(x_17); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +x_19 = lean_ctor_get(x_16, 1); +lean_inc(x_19); +lean_dec(x_16); +x_20 = lean_box(0); +x_21 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__2(x_1, x_15, x_2, x_3, x_4, x_5, x_20, x_10, x_11, x_12, x_13, x_19); +lean_dec(x_5); +lean_dec(x_4); +return x_21; +} +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; uint8_t x_41; +x_22 = lean_ctor_get(x_16, 1); +lean_inc(x_22); +lean_dec(x_16); +x_23 = l_Lean_MessageData_ofExpr(x_6); +x_24 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__2___closed__9; +x_25 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_23); +x_26 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__7; +x_27 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = l_Lean_MessageData_ofExpr(x_7); +x_29 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +x_30 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__9; +x_31 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_30); +x_32 = l_Nat_repr(x_8); +x_33 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_33, 0, x_32); +x_34 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_34, 0, x_33); +x_35 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_35, 0, x_31); +lean_ctor_set(x_35, 1, x_34); +x_36 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__11; +x_37 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_37, 0, x_35); +lean_ctor_set(x_37, 1, x_36); +lean_inc(x_1); +x_38 = l_Lean_MessageData_ofExpr(x_1); +x_39 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +x_40 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_26); +x_41 = lean_nat_dec_lt(x_2, x_4); +if (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; lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_42 = l_Lean_instInhabitedExpr; +x_43 = l___private_Init_Util_0__outOfBounds___rarg(x_42); +x_44 = l_Lean_MessageData_ofExpr(x_43); +x_45 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_45, 0, x_40); +lean_ctor_set(x_45, 1, x_44); +x_46 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_24); +x_47 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_15, x_46, x_10, x_11, x_12, x_13, x_22); +x_48 = lean_ctor_get(x_47, 0); +lean_inc(x_48); +x_49 = lean_ctor_get(x_47, 1); +lean_inc(x_49); +lean_dec(x_47); +x_50 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__2(x_1, x_15, x_2, x_3, x_4, x_5, x_48, x_10, x_11, x_12, x_13, x_49); +lean_dec(x_48); +lean_dec(x_5); +lean_dec(x_4); +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_array_fget(x_5, x_2); +x_52 = l_Lean_MessageData_ofExpr(x_51); +x_53 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_53, 0, x_40); +lean_ctor_set(x_53, 1, x_52); +x_54 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_54, 0, x_53); +lean_ctor_set(x_54, 1, x_24); +x_55 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_15, x_54, x_10, x_11, x_12, x_13, x_22); +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 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__2(x_1, x_15, x_2, x_3, x_4, x_5, x_56, x_10, x_11, x_12, x_13, x_57); +lean_dec(x_56); +lean_dec(x_5); +lean_dec(x_4); +return x_58; +} +} +} +} LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___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_object* x_18, lean_object* x_19) { _start: { @@ -1337,7 +1661,7 @@ x_21 = lean_unsigned_to_nat(0u); x_22 = lean_nat_dec_eq(x_10, 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_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_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_dec(x_14); x_23 = lean_unsigned_to_nat(1u); x_24 = lean_nat_sub(x_10, x_23); @@ -1354,42 +1678,52 @@ lean_inc(x_35); x_36 = lean_ctor_get(x_34, 1); lean_inc(x_36); lean_dec(x_34); -x_37 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__5; -x_38 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_37, x_15, x_16, x_17, x_18, x_36); -x_39 = lean_ctor_get(x_38, 0); -lean_inc(x_39); -x_40 = lean_unbox(x_39); -lean_dec(x_39); -if (x_40 == 0) +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_35); +x_37 = l_Lean_Meta_isProof(x_35, x_15, x_16, x_17, x_18, x_36); +if (lean_obj_tag(x_37) == 0) { -lean_object* x_41; lean_object* x_42; lean_object* x_43; -lean_dec(x_33); -x_41 = lean_ctor_get(x_38, 1); -lean_inc(x_41); +lean_object* x_38; uint8_t x_39; +x_38 = lean_ctor_get(x_37, 0); +lean_inc(x_38); +x_39 = lean_unbox(x_38); lean_dec(x_38); -x_42 = lean_box(0); +if (x_39 == 0) +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_37, 1); +lean_inc(x_40); +lean_dec(x_37); +x_41 = lean_box(0); lean_inc(x_18); lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); +lean_inc(x_2); +lean_inc(x_1); +lean_inc(x_6); +lean_inc(x_8); lean_inc(x_9); lean_inc(x_11); -x_43 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__2(x_35, x_37, x_11, x_9, x_8, x_6, x_42, x_15, x_16, x_17, x_18, x_41); -if (lean_obj_tag(x_43) == 0) +x_42 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3(x_35, x_11, x_9, x_8, x_6, x_1, x_2, x_33, x_41, x_15, x_16, x_17, x_18, x_40); +if (lean_obj_tag(x_42) == 0) { -lean_object* x_44; lean_object* x_45; -x_44 = lean_ctor_get(x_43, 0); +lean_object* x_43; lean_object* x_44; +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_25 = x_44; -x_26 = x_45; +lean_dec(x_42); +x_25 = x_43; +x_26 = x_44; goto block_32; } else { -uint8_t x_46; +uint8_t x_45; lean_dec(x_24); lean_dec(x_18); lean_dec(x_17); @@ -1397,118 +1731,101 @@ lean_dec(x_16); lean_dec(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_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_46 = !lean_is_exclusive(x_43); -if (x_46 == 0) +x_45 = !lean_is_exclusive(x_42); +if (x_45 == 0) { -return x_43; +return x_42; } else { -lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_47 = lean_ctor_get(x_43, 0); -x_48 = lean_ctor_get(x_43, 1); -lean_inc(x_48); +lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_46 = lean_ctor_get(x_42, 0); +x_47 = lean_ctor_get(x_42, 1); lean_inc(x_47); -lean_dec(x_43); -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_42); +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; } } } 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; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; uint8_t x_69; -x_50 = lean_ctor_get(x_38, 1); -lean_inc(x_50); -lean_dec(x_38); -lean_inc(x_1); -x_51 = l_Lean_MessageData_ofExpr(x_1); -x_52 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__2___closed__9; -x_53 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_53, 0, x_52); -lean_ctor_set(x_53, 1, x_51); -x_54 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__7; -x_55 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_55, 0, x_53); -lean_ctor_set(x_55, 1, x_54); -lean_inc(x_2); -x_56 = l_Lean_MessageData_ofExpr(x_2); -x_57 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_57, 0, x_55); -lean_ctor_set(x_57, 1, x_56); -x_58 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__9; -x_59 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_59, 0, x_57); -lean_ctor_set(x_59, 1, x_58); -x_60 = l_Nat_repr(x_33); -x_61 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_61, 0, x_60); -x_62 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_62, 0, x_61); -x_63 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_63, 0, x_59); -lean_ctor_set(x_63, 1, x_62); -x_64 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__11; -x_65 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_65, 0, x_63); -lean_ctor_set(x_65, 1, x_64); -lean_inc(x_35); -x_66 = l_Lean_MessageData_ofExpr(x_35); -x_67 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_67, 0, x_65); -lean_ctor_set(x_67, 1, x_66); -x_68 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_68, 0, x_67); -lean_ctor_set(x_68, 1, x_54); -x_69 = lean_nat_dec_lt(x_11, x_8); -if (x_69 == 0) +lean_object* x_49; uint8_t x_50; +x_49 = lean_ctor_get(x_37, 1); +lean_inc(x_49); +lean_dec(x_37); +x_50 = lean_nat_dec_lt(x_11, x_8); +if (x_50 == 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; -x_70 = l_Lean_instInhabitedExpr; -x_71 = l___private_Init_Util_0__outOfBounds___rarg(x_70); -x_72 = l_Lean_MessageData_ofExpr(x_71); -x_73 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_73, 0, x_68); -lean_ctor_set(x_73, 1, x_72); -x_74 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_74, 0, x_73); -lean_ctor_set(x_74, 1, x_52); -x_75 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_37, x_74, x_15, x_16, x_17, x_18, x_50); -x_76 = lean_ctor_get(x_75, 0); -lean_inc(x_76); -x_77 = lean_ctor_get(x_75, 1); -lean_inc(x_77); -lean_dec(x_75); +lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_51 = l_Lean_instInhabitedExpr; +x_52 = l___private_Init_Util_0__outOfBounds___rarg(x_51); +x_53 = l_Lean_Meta_isAbstractedUnassignedMVar(x_52, x_15, x_16, x_17, x_18, x_49); +if (lean_obj_tag(x_53) == 0) +{ +lean_object* x_54; uint8_t x_55; +x_54 = lean_ctor_get(x_53, 0); +lean_inc(x_54); +x_55 = lean_unbox(x_54); +lean_dec(x_54); +if (x_55 == 0) +{ +lean_object* x_56; lean_object* x_57; +lean_dec(x_35); +lean_dec(x_33); +x_56 = lean_ctor_get(x_53, 1); +lean_inc(x_56); +lean_dec(x_53); +lean_inc(x_9); +x_57 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_57, 0, x_9); +x_25 = x_57; +x_26 = x_56; +goto block_32; +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_53, 1); +lean_inc(x_58); +lean_dec(x_53); +x_59 = lean_box(0); lean_inc(x_18); lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); +lean_inc(x_2); +lean_inc(x_1); +lean_inc(x_6); +lean_inc(x_8); lean_inc(x_9); lean_inc(x_11); -x_78 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__2(x_35, x_37, x_11, x_9, x_8, x_6, x_76, x_15, x_16, x_17, x_18, x_77); -lean_dec(x_76); -if (lean_obj_tag(x_78) == 0) +x_60 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3(x_35, x_11, x_9, x_8, x_6, x_1, x_2, x_33, x_59, x_15, x_16, x_17, x_18, x_58); +if (lean_obj_tag(x_60) == 0) { -lean_object* x_79; lean_object* x_80; -x_79 = lean_ctor_get(x_78, 0); -lean_inc(x_79); -x_80 = lean_ctor_get(x_78, 1); -lean_inc(x_80); -lean_dec(x_78); -x_25 = x_79; -x_26 = x_80; +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_25 = x_61; +x_26 = x_62; goto block_32; } else { -uint8_t x_81; +uint8_t x_63; lean_dec(x_24); lean_dec(x_18); lean_dec(x_17); @@ -1516,71 +1833,133 @@ lean_dec(x_16); lean_dec(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_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_81 = !lean_is_exclusive(x_78); -if (x_81 == 0) +x_63 = !lean_is_exclusive(x_60); +if (x_63 == 0) { -return x_78; +return x_60; } else { -lean_object* x_82; lean_object* x_83; lean_object* x_84; -x_82 = lean_ctor_get(x_78, 0); -x_83 = lean_ctor_get(x_78, 1); -lean_inc(x_83); -lean_inc(x_82); -lean_dec(x_78); -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_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; +} } } } 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_array_fget(x_6, x_11); -x_86 = l_Lean_MessageData_ofExpr(x_85); -x_87 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_87, 0, x_68); -lean_ctor_set(x_87, 1, x_86); -x_88 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_88, 0, x_87); -lean_ctor_set(x_88, 1, x_52); -x_89 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_37, x_88, x_15, x_16, x_17, x_18, x_50); -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_89); +uint8_t x_67; +lean_dec(x_35); +lean_dec(x_33); +lean_dec(x_24); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(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_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_67 = !lean_is_exclusive(x_53); +if (x_67 == 0) +{ +return x_53; +} +else +{ +lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_68 = lean_ctor_get(x_53, 0); +x_69 = lean_ctor_get(x_53, 1); +lean_inc(x_69); +lean_inc(x_68); +lean_dec(x_53); +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; +} +} +} +else +{ +lean_object* x_71; lean_object* x_72; +x_71 = lean_array_fget(x_6, x_11); +x_72 = l_Lean_Meta_isAbstractedUnassignedMVar(x_71, x_15, x_16, x_17, x_18, x_49); +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); +if (x_74 == 0) +{ +lean_object* x_75; lean_object* x_76; +lean_dec(x_35); +lean_dec(x_33); +x_75 = lean_ctor_get(x_72, 1); +lean_inc(x_75); +lean_dec(x_72); +lean_inc(x_9); +x_76 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_76, 0, x_9); +x_25 = x_76; +x_26 = x_75; +goto block_32; +} +else +{ +lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_77 = lean_ctor_get(x_72, 1); +lean_inc(x_77); +lean_dec(x_72); +x_78 = lean_box(0); lean_inc(x_18); lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); +lean_inc(x_2); +lean_inc(x_1); +lean_inc(x_6); +lean_inc(x_8); lean_inc(x_9); lean_inc(x_11); -x_92 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__2(x_35, x_37, x_11, x_9, x_8, x_6, x_90, x_15, x_16, x_17, x_18, x_91); -lean_dec(x_90); -if (lean_obj_tag(x_92) == 0) +x_79 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3(x_35, x_11, x_9, x_8, x_6, x_1, x_2, x_33, x_78, x_15, x_16, x_17, x_18, x_77); +if (lean_obj_tag(x_79) == 0) { -lean_object* x_93; lean_object* x_94; -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_25 = x_93; -x_26 = x_94; +lean_object* x_80; lean_object* x_81; +x_80 = lean_ctor_get(x_79, 0); +lean_inc(x_80); +x_81 = lean_ctor_get(x_79, 1); +lean_inc(x_81); +lean_dec(x_79); +x_25 = x_80; +x_26 = x_81; goto block_32; } else { -uint8_t x_95; +uint8_t x_82; lean_dec(x_24); lean_dec(x_18); lean_dec(x_17); @@ -1588,30 +1967,111 @@ lean_dec(x_16); lean_dec(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_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_95 = !lean_is_exclusive(x_92); -if (x_95 == 0) +x_82 = !lean_is_exclusive(x_79); +if (x_82 == 0) { -return x_92; +return x_79; } else { -lean_object* x_96; lean_object* x_97; lean_object* x_98; -x_96 = lean_ctor_get(x_92, 0); -x_97 = lean_ctor_get(x_92, 1); -lean_inc(x_97); -lean_inc(x_96); -lean_dec(x_92); -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; +lean_object* x_83; lean_object* x_84; lean_object* x_85; +x_83 = lean_ctor_get(x_79, 0); +x_84 = lean_ctor_get(x_79, 1); +lean_inc(x_84); +lean_inc(x_83); +lean_dec(x_79); +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; +} +} +} +} +else +{ +uint8_t x_86; +lean_dec(x_35); +lean_dec(x_33); +lean_dec(x_24); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(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_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_86 = !lean_is_exclusive(x_72); +if (x_86 == 0) +{ +return x_72; +} +else +{ +lean_object* x_87; lean_object* x_88; lean_object* x_89; +x_87 = lean_ctor_get(x_72, 0); +x_88 = lean_ctor_get(x_72, 1); +lean_inc(x_88); +lean_inc(x_87); +lean_dec(x_72); +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_90; +lean_dec(x_35); +lean_dec(x_33); +lean_dec(x_24); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(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_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_90 = !lean_is_exclusive(x_37); +if (x_90 == 0) +{ +return x_37; } +else +{ +lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_91 = lean_ctor_get(x_37, 0); +x_92 = lean_ctor_get(x_37, 1); +lean_inc(x_92); +lean_inc(x_91); +lean_dec(x_37); +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; } } block_32: @@ -1626,7 +2086,9 @@ lean_dec(x_16); lean_dec(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_4); lean_dec(x_3); lean_dec(x_2); @@ -1657,7 +2119,7 @@ goto _start; } else { -lean_object* x_99; +lean_object* x_94; lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); @@ -1665,20 +2127,22 @@ 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_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_99 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_99, 0, x_14); -lean_ctor_set(x_99, 1, x_19); -return x_99; +x_94 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_94, 0, x_14); +lean_ctor_set(x_94, 1, x_19); +return x_94; } } else { -lean_object* x_100; +lean_object* x_95; lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); @@ -1686,15 +2150,17 @@ 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_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_100 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_100, 0, x_14); -lean_ctor_set(x_100, 1, x_19); -return x_100; +x_95 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_95, 0, x_14); +lean_ctor_set(x_95, 1, x_19); +return x_95; } } } @@ -1818,7 +2284,7 @@ lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); lean_inc(x_9); -lean_inc(x_8); +lean_inc_n(x_8, 2); lean_inc(x_5); x_82 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_8, x_5, x_8, x_81, x_9, x_11, x_12, x_13, x_14, x_32); lean_dec(x_8); @@ -2121,7 +2587,7 @@ lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); lean_inc(x_9); -lean_inc(x_8); +lean_inc_n(x_8, 2); lean_inc(x_5); x_151 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_8, x_5, x_8, x_150, x_9, x_11, x_12, x_13, x_14, x_110); lean_dec(x_8); @@ -2438,7 +2904,7 @@ lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); lean_inc(x_9); -lean_inc(x_8); +lean_inc_n(x_8, 2); lean_inc(x_5); x_225 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_8, x_5, x_8, x_224, x_9, x_11, x_12, x_13, x_14, x_184); lean_dec(x_8); @@ -2788,7 +3254,7 @@ lean_dec(x_10); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_16 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__5; +x_16 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__5; x_17 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_16, x_5, x_6, x_7, x_8, x_9); x_18 = lean_ctor_get(x_17, 0); lean_inc(x_18); @@ -2851,7 +3317,7 @@ lean_dec(x_10); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_39 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__5; +x_39 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__5; x_40 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_39, x_5, x_6, x_7, x_8, x_35); x_41 = lean_ctor_get(x_40, 0); lean_inc(x_41); @@ -2994,7 +3460,6 @@ x_81 = lean_array_get_size(x_78); x_82 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___closed__7; x_83 = 1; x_84 = l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__2(x_1, x_2, x_3, x_4, x_10, x_78, x_80, x_81, x_82, x_83, x_5, x_6, x_7, x_8, x_73); -lean_dec(x_78); return x_84; } } @@ -3151,8 +3616,6 @@ lean_object* x_20; x_20 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___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, x_18, x_19); lean_dec(x_13); lean_dec(x_12); -lean_dec(x_8); -lean_dec(x_6); lean_dec(x_5); return x_20; } @@ -3177,7 +3640,6 @@ uint8_t x_16; lean_object* x_17; x_16 = lean_unbox(x_10); lean_dec(x_10); x_17 = l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_16, x_11, x_12, x_13, x_14, x_15); -lean_dec(x_6); return x_17; } } @@ -5979,8 +6441,8 @@ LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDef lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; x_11 = lean_array_push(x_4, x_1); x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_11); -lean_ctor_set(x_12, 1, x_2); +lean_ctor_set(x_12, 0, x_2); +lean_ctor_set(x_12, 1, x_11); x_13 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_13, 0, x_3); lean_ctor_set(x_13, 1, x_12); @@ -5992,7 +6454,7 @@ lean_ctor_set(x_15, 1, x_10); return x_15; } } -static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___closed__1() { +static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___lambda__2___closed__1() { _start: { lean_object* x_1; lean_object* x_2; @@ -6002,17 +6464,159 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___closed__2() { +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___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_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 = lean_is_expr_def_eq(x_1, x_2, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_12) == 0) +{ +lean_object* x_13; uint8_t x_14; +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_unbox(x_13); +lean_dec(x_13); +if (x_14 == 0) +{ +uint8_t x_15; +lean_dec(x_5); +x_15 = !lean_is_exclusive(x_12); +if (x_15 == 0) +{ +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_12, 0); +lean_dec(x_16); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_3); +lean_ctor_set(x_17, 1, x_4); +x_18 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___lambda__2___closed__1; +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_17); +x_20 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_12, 0, x_20); +return x_12; +} +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; +x_21 = lean_ctor_get(x_12, 1); +lean_inc(x_21); +lean_dec(x_12); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_3); +lean_ctor_set(x_22, 1, x_4); +x_23 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___lambda__2___closed__1; +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_22); +x_25 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_25, 0, x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_21); +return x_26; +} +} +else +{ +uint8_t x_27; +x_27 = !lean_is_exclusive(x_12); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_28 = lean_ctor_get(x_12, 0); +lean_dec(x_28); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_3); +lean_ctor_set(x_29, 1, x_4); +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 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_12, 0, x_31); +return x_12; +} +else +{ +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_12, 1); +lean_inc(x_32); +lean_dec(x_12); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_3); +lean_ctor_set(x_33, 1, x_4); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_5); +lean_ctor_set(x_34, 1, x_33); +x_35 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_35, 0, 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_32); +return x_36; +} +} +} +else +{ +uint8_t x_37; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_37 = !lean_is_exclusive(x_12); +if (x_37 == 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_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; +} +} +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___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) { +_start: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_11 = lean_array_push(x_4, x_1); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_2); +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 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_14, 0, 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_10); +return x_15; +} +} +static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__1; -x_2 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__2; +x_1 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__1; +x_2 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__2; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___closed__3() { +static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___closed__2() { _start: { lean_object* x_1; @@ -6020,11 +6624,11 @@ x_1 = lean_mk_string_from_bytes("found messy ", 12); return x_1; } } -static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___closed__4() { +static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___closed__3() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___closed__3; +x_1 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___closed__2; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } @@ -6041,52 +6645,57 @@ x_17 = lean_unsigned_to_nat(0u); x_18 = lean_nat_dec_eq(x_6, 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_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; lean_object* x_34; uint8_t x_35; lean_object* x_36; uint8_t x_37; lean_object* x_38; +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; lean_object* x_43; uint8_t x_44; lean_object* x_45; uint8_t x_46; lean_object* x_47; x_19 = lean_unsigned_to_nat(1u); x_20 = lean_nat_sub(x_6, x_19); lean_dec(x_6); -x_29 = lean_ctor_get(x_10, 1); -lean_inc(x_29); +x_38 = lean_ctor_get(x_10, 1); +lean_inc(x_38); lean_dec(x_10); -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; +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_29); - x_32 = lean_box(0); + lean_dec_ref(x_38); + x_41 = lean_box(0); } -x_33 = lean_nat_dec_lt(x_7, x_4); -x_34 = lean_array_get_size(x_2); -x_35 = lean_nat_dec_lt(x_7, x_34); -lean_dec(x_34); -x_36 = lean_array_get_size(x_3); -x_37 = lean_nat_dec_lt(x_7, x_36); -lean_dec(x_36); -if (x_33 == 0) +x_42 = lean_nat_dec_lt(x_7, x_4); +x_43 = lean_array_get_size(x_2); +x_44 = lean_nat_dec_lt(x_7, x_43); +lean_dec(x_43); +x_45 = lean_array_get_size(x_3); +x_46 = lean_nat_dec_lt(x_7, x_45); +lean_dec(x_45); +if (x_42 == 0) { -lean_object* x_166; lean_object* x_167; -x_166 = l_Lean_Meta_instInhabitedParamInfo; -x_167 = l___private_Init_Util_0__outOfBounds___rarg(x_166); -x_38 = x_167; -goto block_165; +lean_object* x_232; lean_object* x_233; +x_232 = l_Lean_Meta_instInhabitedParamInfo; +x_233 = l___private_Init_Util_0__outOfBounds___rarg(x_232); +x_47 = x_233; +goto block_231; } else { -lean_object* x_168; -x_168 = lean_array_fget(x_1, x_7); -x_38 = x_168; -goto block_165; +lean_object* x_234; +x_234 = lean_array_fget(x_1, x_7); +x_47 = x_234; +goto block_231; } -block_28: +block_37: { if (lean_obj_tag(x_21) == 0) { -lean_object* x_23; lean_object* x_24; +lean_object* x_22; +x_22 = lean_ctor_get(x_21, 0); +lean_inc(x_22); +if (lean_obj_tag(x_22) == 0) +{ +uint8_t x_23; lean_dec(x_20); lean_dec(x_14); lean_dec(x_13); @@ -6094,566 +6703,909 @@ lean_dec(x_12); lean_dec(x_11); lean_dec(x_7); lean_dec(x_5); -x_23 = lean_ctor_get(x_21, 0); -lean_inc(x_23); +x_23 = !lean_is_exclusive(x_21); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; +x_24 = lean_ctor_get(x_21, 0); +lean_dec(x_24); +x_25 = lean_ctor_get(x_22, 0); +lean_inc(x_25); +lean_dec(x_22); +lean_ctor_set(x_21, 0, x_25); +return x_21; +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_21, 1); +lean_inc(x_26); 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; +x_27 = lean_ctor_get(x_22, 0); +lean_inc(x_27); +lean_dec(x_22); +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; +} } else { -lean_object* x_25; lean_object* x_26; -x_25 = lean_ctor_get(x_21, 0); -lean_inc(x_25); +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_21, 1); +lean_inc(x_29); lean_dec(x_21); -x_26 = lean_nat_add(x_7, x_9); +x_30 = lean_ctor_get(x_22, 0); +lean_inc(x_30); +lean_dec(x_22); +x_31 = lean_nat_add(x_7, x_9); lean_dec(x_7); x_6 = x_20; -x_7 = x_26; -x_10 = x_25; -x_15 = x_22; +x_7 = x_31; +x_10 = x_30; +x_15 = x_29; goto _start; } } -block_165: +else { -uint8_t x_39; lean_object* x_40; -x_39 = lean_ctor_get_uint8(x_38, sizeof(void*)*1 + 5); -if (x_35 == 0) +uint8_t x_33; +lean_dec(x_20); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_7); +lean_dec(x_5); +x_33 = !lean_is_exclusive(x_21); +if (x_33 == 0) { -lean_object* x_162; lean_object* x_163; -x_162 = l_Lean_instInhabitedExpr; -x_163 = l___private_Init_Util_0__outOfBounds___rarg(x_162); -x_40 = x_163; -goto block_161; +return x_21; } else { -lean_object* x_164; -x_164 = lean_array_fget(x_2, x_7); -x_40 = x_164; -goto block_161; +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_21, 0); +x_35 = lean_ctor_get(x_21, 1); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_21); +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; +} } -block_161: +} +block_231: { -lean_object* x_41; -if (x_37 == 0) +uint8_t x_48; lean_object* x_49; +x_48 = lean_ctor_get_uint8(x_47, sizeof(void*)*1 + 5); +if (x_44 == 0) { -lean_object* x_158; lean_object* x_159; -x_158 = l_Lean_instInhabitedExpr; -x_159 = l___private_Init_Util_0__outOfBounds___rarg(x_158); -x_41 = x_159; -goto block_157; +lean_object* x_228; lean_object* x_229; +x_228 = l_Lean_instInhabitedExpr; +x_229 = l___private_Init_Util_0__outOfBounds___rarg(x_228); +x_49 = x_229; +goto block_227; } else { -lean_object* x_160; -x_160 = lean_array_fget(x_3, x_7); -x_41 = x_160; -goto block_157; +lean_object* x_230; +x_230 = lean_array_fget(x_2, x_7); +x_49 = x_230; +goto block_227; } -block_157: +block_227: { -uint8_t x_42; lean_object* x_43; -if (x_39 == 0) +lean_object* x_50; +if (x_46 == 0) { -uint8_t x_65; -x_65 = lean_ctor_get_uint8(x_38, sizeof(void*)*1 + 4); -if (x_65 == 0) +lean_object* x_224; lean_object* x_225; +x_224 = l_Lean_instInhabitedExpr; +x_225 = l___private_Init_Util_0__outOfBounds___rarg(x_224); +x_50 = x_225; +goto block_223; +} +else { -uint8_t x_66; -x_66 = l_Lean_Meta_ParamInfo_isExplicit(x_38); -lean_dec(x_38); -if (x_66 == 0) +lean_object* x_226; +x_226 = lean_array_fget(x_3, x_7); +x_50 = x_226; +goto block_223; +} +block_223: { -lean_object* x_67; -lean_inc(x_40); -x_67 = l_Lean_Meta_isEtaUnassignedMVar(x_40, x_11, x_12, x_13, x_14, x_15); -if (lean_obj_tag(x_67) == 0) +uint8_t x_51; lean_object* x_52; +if (x_48 == 0) { -lean_object* x_68; uint8_t x_69; -x_68 = lean_ctor_get(x_67, 0); -lean_inc(x_68); -x_69 = lean_unbox(x_68); -if (x_69 == 0) +uint8_t x_117; +x_117 = lean_ctor_get_uint8(x_47, sizeof(void*)*1 + 4); +if (x_117 == 0) { -lean_object* x_70; lean_object* x_71; -lean_dec(x_68); -x_70 = lean_ctor_get(x_67, 1); -lean_inc(x_70); -lean_dec(x_67); -lean_inc(x_41); -x_71 = l_Lean_Meta_isEtaUnassignedMVar(x_41, x_11, x_12, x_13, x_14, x_70); -if (lean_obj_tag(x_71) == 0) +uint8_t x_118; +x_118 = l_Lean_Meta_ParamInfo_isExplicit(x_47); +if (x_118 == 0) { -lean_object* x_72; lean_object* x_73; uint8_t x_74; -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_unbox(x_72); -lean_dec(x_72); -x_42 = x_74; -x_43 = x_73; -goto block_64; +lean_object* x_119; +lean_inc(x_49); +x_119 = l_Lean_Meta_isEtaUnassignedMVar(x_49, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_119) == 0) +{ +lean_object* x_120; uint8_t x_121; +x_120 = lean_ctor_get(x_119, 0); +lean_inc(x_120); +x_121 = lean_unbox(x_120); +if (x_121 == 0) +{ +lean_object* x_122; lean_object* x_123; +lean_dec(x_120); +x_122 = lean_ctor_get(x_119, 1); +lean_inc(x_122); +lean_dec(x_119); +lean_inc(x_50); +x_123 = l_Lean_Meta_isEtaUnassignedMVar(x_50, x_11, x_12, x_13, x_14, x_122); +if (lean_obj_tag(x_123) == 0) +{ +lean_object* x_124; lean_object* x_125; uint8_t x_126; +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_unbox(x_124); +lean_dec(x_124); +x_51 = x_126; +x_52 = x_125; +goto block_116; } else { -uint8_t x_75; +uint8_t x_127; +lean_dec(x_50); +lean_dec(x_49); +lean_dec(x_47); lean_dec(x_41); lean_dec(x_40); -lean_dec(x_32); -lean_dec(x_31); -lean_dec(x_30); -lean_dec(x_20); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_7); -lean_dec(x_5); -x_75 = !lean_is_exclusive(x_71); -if (x_75 == 0) +lean_dec(x_39); +x_127 = !lean_is_exclusive(x_123); +if (x_127 == 0) { -return x_71; +x_21 = x_123; +goto block_37; } else { -lean_object* x_76; lean_object* x_77; lean_object* x_78; -x_76 = lean_ctor_get(x_71, 0); -x_77 = lean_ctor_get(x_71, 1); -lean_inc(x_77); -lean_inc(x_76); -lean_dec(x_71); -x_78 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_78, 0, x_76); -lean_ctor_set(x_78, 1, x_77); -return x_78; +lean_object* x_128; lean_object* x_129; lean_object* x_130; +x_128 = lean_ctor_get(x_123, 0); +x_129 = lean_ctor_get(x_123, 1); +lean_inc(x_129); +lean_inc(x_128); +lean_dec(x_123); +x_130 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_130, 0, x_128); +lean_ctor_set(x_130, 1, x_129); +x_21 = x_130; +goto block_37; } } } else { -lean_object* x_79; uint8_t x_80; -x_79 = lean_ctor_get(x_67, 1); -lean_inc(x_79); -lean_dec(x_67); -x_80 = lean_unbox(x_68); -lean_dec(x_68); -x_42 = x_80; -x_43 = x_79; -goto block_64; +lean_object* x_131; uint8_t x_132; +x_131 = lean_ctor_get(x_119, 1); +lean_inc(x_131); +lean_dec(x_119); +x_132 = lean_unbox(x_120); +lean_dec(x_120); +x_51 = x_132; +x_52 = x_131; +goto block_116; } } else { -uint8_t x_81; +uint8_t x_133; +lean_dec(x_50); +lean_dec(x_49); +lean_dec(x_47); lean_dec(x_41); lean_dec(x_40); -lean_dec(x_32); -lean_dec(x_31); -lean_dec(x_30); -lean_dec(x_20); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_7); -lean_dec(x_5); -x_81 = !lean_is_exclusive(x_67); -if (x_81 == 0) +lean_dec(x_39); +x_133 = !lean_is_exclusive(x_119); +if (x_133 == 0) { -return x_67; +x_21 = x_119; +goto block_37; } else { -lean_object* x_82; lean_object* x_83; lean_object* x_84; -x_82 = lean_ctor_get(x_67, 0); -x_83 = lean_ctor_get(x_67, 1); -lean_inc(x_83); -lean_inc(x_82); -lean_dec(x_67); -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_object* x_134; lean_object* x_135; lean_object* x_136; +x_134 = lean_ctor_get(x_119, 0); +x_135 = lean_ctor_get(x_119, 1); +lean_inc(x_135); +lean_inc(x_134); +lean_dec(x_119); +x_136 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_136, 0, x_134); +lean_ctor_set(x_136, 1, x_135); +x_21 = x_136; +goto block_37; } } } else { -lean_object* x_85; -lean_dec(x_32); +uint8_t x_137; +lean_dec(x_41); +x_137 = lean_ctor_get_uint8(x_47, sizeof(void*)*1 + 2); +lean_dec(x_47); +if (x_137 == 0) +{ +lean_object* x_138; lean_object* x_139; +x_138 = lean_box(0); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); -x_85 = lean_is_expr_def_eq(x_40, x_41, x_11, x_12, x_13, x_14, x_15); -if (lean_obj_tag(x_85) == 0) -{ -lean_object* x_86; uint8_t x_87; -x_86 = lean_ctor_get(x_85, 0); -lean_inc(x_86); -x_87 = lean_unbox(x_86); -lean_dec(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; -x_88 = lean_ctor_get(x_85, 1); -lean_inc(x_88); -lean_dec(x_85); -x_89 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_89, 0, x_30); -lean_ctor_set(x_89, 1, x_31); -x_90 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___closed__1; -x_91 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_91, 0, x_90); -lean_ctor_set(x_91, 1, x_89); -x_92 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_92, 0, x_91); -x_21 = x_92; -x_22 = x_88; -goto block_28; +lean_inc(x_5); +x_139 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___lambda__2(x_49, x_50, x_39, x_40, x_5, x_138, x_11, x_12, x_13, x_14, x_15); +x_21 = x_139; +goto block_37; } else { -lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; -x_93 = lean_ctor_get(x_85, 1); -lean_inc(x_93); -lean_dec(x_85); -x_94 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_94, 0, x_30); -lean_ctor_set(x_94, 1, x_31); +lean_object* x_140; +lean_inc(x_49); +x_140 = l_Lean_Meta_isAbstractedUnassignedMVar(x_49, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_140) == 0) +{ +lean_object* x_141; uint8_t x_142; +x_141 = lean_ctor_get(x_140, 0); +lean_inc(x_141); +x_142 = lean_unbox(x_141); +lean_dec(x_141); +if (x_142 == 0) +{ +lean_object* x_143; lean_object* x_144; +x_143 = lean_ctor_get(x_140, 1); +lean_inc(x_143); +lean_dec(x_140); +lean_inc(x_50); +x_144 = l_Lean_Meta_isAbstractedUnassignedMVar(x_50, x_11, x_12, x_13, x_14, x_143); +if (lean_obj_tag(x_144) == 0) +{ +lean_object* x_145; uint8_t x_146; +x_145 = lean_ctor_get(x_144, 0); +lean_inc(x_145); +x_146 = lean_unbox(x_145); +lean_dec(x_145); +if (x_146 == 0) +{ +uint8_t x_147; +lean_dec(x_50); +lean_dec(x_49); +x_147 = !lean_is_exclusive(x_144); +if (x_147 == 0) +{ +lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; +x_148 = lean_ctor_get(x_144, 0); +lean_dec(x_148); +x_149 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_149, 0, x_39); +lean_ctor_set(x_149, 1, x_40); lean_inc(x_5); -x_95 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_95, 0, x_5); -lean_ctor_set(x_95, 1, x_94); -x_96 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_96, 0, x_95); -x_21 = x_96; -x_22 = x_93; -goto block_28; -} +x_150 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_150, 0, x_5); +lean_ctor_set(x_150, 1, x_149); +x_151 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_151, 0, x_150); +lean_ctor_set(x_144, 0, x_151); +x_21 = x_144; +goto block_37; } else { -uint8_t x_97; -lean_dec(x_31); -lean_dec(x_30); -lean_dec(x_20); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_7); -lean_dec(x_5); -x_97 = !lean_is_exclusive(x_85); -if (x_97 == 0) +lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; +x_152 = lean_ctor_get(x_144, 1); +lean_inc(x_152); +lean_dec(x_144); +x_153 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_153, 0, x_39); +lean_ctor_set(x_153, 1, x_40); +lean_inc(x_5); +x_154 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_154, 0, x_5); +lean_ctor_set(x_154, 1, x_153); +x_155 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_155, 0, x_154); +x_156 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_156, 0, x_155); +lean_ctor_set(x_156, 1, x_152); +x_21 = x_156; +goto block_37; +} +} +else { -return x_85; +lean_object* x_157; lean_object* x_158; lean_object* x_159; +x_157 = lean_ctor_get(x_144, 1); +lean_inc(x_157); +lean_dec(x_144); +x_158 = lean_box(0); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_5); +x_159 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___lambda__2(x_49, x_50, x_39, x_40, x_5, x_158, x_11, x_12, x_13, x_14, x_157); +x_21 = x_159; +goto block_37; +} } else { -lean_object* x_98; lean_object* x_99; lean_object* x_100; -x_98 = lean_ctor_get(x_85, 0); -x_99 = lean_ctor_get(x_85, 1); -lean_inc(x_99); -lean_inc(x_98); -lean_dec(x_85); -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; +uint8_t x_160; +lean_dec(x_50); +lean_dec(x_49); +lean_dec(x_40); +lean_dec(x_39); +x_160 = !lean_is_exclusive(x_144); +if (x_160 == 0) +{ +x_21 = x_144; +goto block_37; } +else +{ +lean_object* x_161; lean_object* x_162; lean_object* x_163; +x_161 = lean_ctor_get(x_144, 0); +x_162 = lean_ctor_get(x_144, 1); +lean_inc(x_162); +lean_inc(x_161); +lean_dec(x_144); +x_163 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_163, 0, x_161); +lean_ctor_set(x_163, 1, x_162); +x_21 = x_163; +goto block_37; } } } else { -lean_object* x_101; lean_object* x_102; lean_object* x_103; uint8_t x_104; -lean_dec(x_38); -lean_dec(x_32); -x_101 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___closed__2; -x_102 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_101, x_11, x_12, x_13, x_14, x_15); -x_103 = lean_ctor_get(x_102, 0); -lean_inc(x_103); -x_104 = lean_unbox(x_103); -lean_dec(x_103); -if (x_104 == 0) +lean_object* x_164; lean_object* x_165; lean_object* x_166; +x_164 = lean_ctor_get(x_140, 1); +lean_inc(x_164); +lean_dec(x_140); +x_165 = lean_box(0); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_5); +x_166 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___lambda__2(x_49, x_50, x_39, x_40, x_5, x_165, x_11, x_12, x_13, x_14, x_164); +x_21 = x_166; +goto block_37; +} +} +else { -lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; -lean_dec(x_41); +uint8_t x_167; +lean_dec(x_50); +lean_dec(x_49); lean_dec(x_40); -x_105 = lean_ctor_get(x_102, 1); -lean_inc(x_105); -lean_dec(x_102); -x_106 = lean_box(0); +lean_dec(x_39); +x_167 = !lean_is_exclusive(x_140); +if (x_167 == 0) +{ +x_21 = x_140; +goto block_37; +} +else +{ +lean_object* x_168; lean_object* x_169; lean_object* x_170; +x_168 = lean_ctor_get(x_140, 0); +x_169 = lean_ctor_get(x_140, 1); +lean_inc(x_169); +lean_inc(x_168); +lean_dec(x_140); +x_170 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_170, 0, x_168); +lean_ctor_set(x_170, 1, x_169); +x_21 = x_170; +goto block_37; +} +} +} +} +} +else +{ +lean_object* x_171; lean_object* x_172; lean_object* x_173; uint8_t x_174; +lean_dec(x_47); +lean_dec(x_41); +x_171 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___closed__1; +x_172 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_171, x_11, x_12, x_13, x_14, x_15); +x_173 = lean_ctor_get(x_172, 0); +lean_inc(x_173); +x_174 = lean_unbox(x_173); +lean_dec(x_173); +if (x_174 == 0) +{ +lean_object* x_175; lean_object* x_176; lean_object* x_177; +lean_dec(x_50); +lean_dec(x_49); +x_175 = lean_ctor_get(x_172, 1); +lean_inc(x_175); +lean_dec(x_172); +x_176 = lean_box(0); lean_inc(x_5); lean_inc(x_7); -x_107 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___lambda__1(x_7, x_31, x_5, x_30, x_106, x_11, x_12, x_13, x_14, 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_21 = x_108; -x_22 = x_109; -goto block_28; +x_177 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___lambda__3(x_7, x_40, x_5, x_39, x_176, x_11, x_12, x_13, x_14, x_175); +x_21 = x_177; +goto block_37; } 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; -x_110 = lean_ctor_get(x_102, 1); -lean_inc(x_110); -lean_dec(x_102); -x_111 = l_Lean_MessageData_ofExpr(x_40); -x_112 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___closed__4; -x_113 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_113, 0, x_112); -lean_ctor_set(x_113, 1, x_111); -x_114 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__7; -x_115 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_115, 0, x_113); -lean_ctor_set(x_115, 1, x_114); -x_116 = l_Lean_MessageData_ofExpr(x_41); -x_117 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_117, 0, x_115); -lean_ctor_set(x_117, 1, x_116); -x_118 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__2___closed__9; -x_119 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_119, 0, x_117); -lean_ctor_set(x_119, 1, x_118); -x_120 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_101, x_119, x_11, x_12, x_13, x_14, x_110); -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); +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_178 = lean_ctor_get(x_172, 1); +lean_inc(x_178); +lean_dec(x_172); +x_179 = l_Lean_MessageData_ofExpr(x_49); +x_180 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___closed__3; +x_181 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_181, 0, x_180); +lean_ctor_set(x_181, 1, x_179); +x_182 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__7; +x_183 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_183, 0, x_181); +lean_ctor_set(x_183, 1, x_182); +x_184 = l_Lean_MessageData_ofExpr(x_50); +x_185 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_185, 0, x_183); +lean_ctor_set(x_185, 1, x_184); +x_186 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__2___closed__9; +x_187 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_187, 0, x_185); +lean_ctor_set(x_187, 1, x_186); +x_188 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_171, x_187, x_11, x_12, x_13, x_14, x_178); +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); lean_inc(x_5); lean_inc(x_7); -x_123 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___lambda__1(x_7, x_31, x_5, x_30, x_121, x_11, x_12, x_13, x_14, x_122); -lean_dec(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_21 = x_124; -x_22 = x_125; -goto block_28; +x_191 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___lambda__3(x_7, x_40, x_5, x_39, x_189, x_11, x_12, x_13, x_14, x_190); +lean_dec(x_189); +x_21 = x_191; +goto block_37; } } } else { -lean_object* x_126; lean_object* x_127; lean_object* x_128; uint8_t x_129; -lean_dec(x_38); -lean_dec(x_32); -x_126 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___closed__2; -x_127 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_126, x_11, x_12, x_13, x_14, x_15); -x_128 = lean_ctor_get(x_127, 0); -lean_inc(x_128); -x_129 = lean_unbox(x_128); -lean_dec(x_128); -if (x_129 == 0) -{ -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_192; lean_object* x_193; lean_object* x_194; uint8_t x_195; +lean_dec(x_47); lean_dec(x_41); -lean_dec(x_40); -x_130 = lean_ctor_get(x_127, 1); -lean_inc(x_130); -lean_dec(x_127); -x_131 = lean_box(0); +x_192 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___closed__1; +x_193 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_192, x_11, x_12, x_13, x_14, x_15); +x_194 = lean_ctor_get(x_193, 0); +lean_inc(x_194); +x_195 = lean_unbox(x_194); +lean_dec(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_dec(x_50); +lean_dec(x_49); +x_196 = lean_ctor_get(x_193, 1); +lean_inc(x_196); +lean_dec(x_193); +x_197 = lean_box(0); lean_inc(x_5); lean_inc(x_7); -x_132 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___lambda__1(x_7, x_31, x_5, x_30, x_131, x_11, x_12, x_13, x_14, x_130); -x_133 = lean_ctor_get(x_132, 0); -lean_inc(x_133); -x_134 = lean_ctor_get(x_132, 1); -lean_inc(x_134); -lean_dec(x_132); -x_135 = lean_ctor_get(x_133, 0); -lean_inc(x_135); -lean_dec(x_133); -x_136 = lean_nat_add(x_7, x_9); +x_198 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___lambda__3(x_7, x_40, x_5, x_39, x_197, x_11, x_12, x_13, x_14, 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_nat_add(x_7, x_9); lean_dec(x_7); x_6 = x_20; -x_7 = x_136; -x_10 = x_135; -x_15 = x_134; +x_7 = x_202; +x_10 = x_201; +x_15 = x_200; goto _start; } 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; -x_138 = lean_ctor_get(x_127, 1); -lean_inc(x_138); -lean_dec(x_127); -x_139 = l_Lean_MessageData_ofExpr(x_40); -x_140 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___closed__4; -x_141 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_141, 0, x_140); -lean_ctor_set(x_141, 1, x_139); -x_142 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__7; -x_143 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_143, 0, x_141); -lean_ctor_set(x_143, 1, x_142); -x_144 = l_Lean_MessageData_ofExpr(x_41); -x_145 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_145, 0, x_143); -lean_ctor_set(x_145, 1, x_144); -x_146 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__2___closed__9; -x_147 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_147, 0, x_145); -lean_ctor_set(x_147, 1, x_146); -x_148 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_126, x_147, x_11, x_12, x_13, x_14, x_138); -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); +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_204 = lean_ctor_get(x_193, 1); +lean_inc(x_204); +lean_dec(x_193); +x_205 = l_Lean_MessageData_ofExpr(x_49); +x_206 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___closed__3; +x_207 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_207, 0, x_206); +lean_ctor_set(x_207, 1, x_205); +x_208 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__7; +x_209 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_209, 0, x_207); +lean_ctor_set(x_209, 1, x_208); +x_210 = l_Lean_MessageData_ofExpr(x_50); +x_211 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_211, 0, x_209); +lean_ctor_set(x_211, 1, x_210); +x_212 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__2___closed__9; +x_213 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_213, 0, x_211); +lean_ctor_set(x_213, 1, x_212); +x_214 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_192, x_213, x_11, x_12, x_13, x_14, x_204); +x_215 = lean_ctor_get(x_214, 0); +lean_inc(x_215); +x_216 = lean_ctor_get(x_214, 1); +lean_inc(x_216); +lean_dec(x_214); lean_inc(x_5); lean_inc(x_7); -x_151 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___lambda__1(x_7, x_31, x_5, x_30, x_149, x_11, x_12, x_13, x_14, x_150); -lean_dec(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_nat_add(x_7, x_9); +x_217 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___lambda__3(x_7, x_40, x_5, x_39, x_215, x_11, x_12, x_13, x_14, x_216); +lean_dec(x_215); +x_218 = lean_ctor_get(x_217, 0); +lean_inc(x_218); +x_219 = lean_ctor_get(x_217, 1); +lean_inc(x_219); +lean_dec(x_217); +x_220 = lean_ctor_get(x_218, 0); +lean_inc(x_220); +lean_dec(x_218); +x_221 = lean_nat_add(x_7, x_9); lean_dec(x_7); x_6 = x_20; -x_7 = x_155; -x_10 = x_154; -x_15 = x_153; +x_7 = x_221; +x_10 = x_220; +x_15 = x_219; goto _start; } } -block_64: +block_116: { -if (x_42 == 0) +if (x_51 == 0) { -lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; +uint8_t x_53; +x_53 = lean_ctor_get_uint8(x_47, sizeof(void*)*1 + 2); +lean_dec(x_47); +if (x_53 == 0) +{ +lean_object* x_54; lean_object* x_55; +lean_dec(x_50); +lean_dec(x_49); lean_dec(x_41); -lean_dec(x_40); +x_54 = lean_box(0); +lean_inc(x_5); lean_inc(x_7); -x_44 = lean_array_push(x_31, x_7); -if (lean_is_scalar(x_32)) { - x_45 = lean_alloc_ctor(0, 2, 0); +x_55 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___lambda__1(x_7, x_39, x_5, x_40, x_54, x_11, x_12, x_13, x_14, x_52); +x_21 = x_55; +goto block_37; +} +else +{ +lean_object* x_56; +x_56 = l_Lean_Meta_isAbstractedUnassignedMVar(x_49, x_11, x_12, x_13, x_14, x_52); +if (lean_obj_tag(x_56) == 0) +{ +lean_object* x_57; uint8_t x_58; +x_57 = lean_ctor_get(x_56, 0); +lean_inc(x_57); +x_58 = lean_unbox(x_57); +lean_dec(x_57); +if (x_58 == 0) +{ +lean_object* x_59; lean_object* x_60; +x_59 = lean_ctor_get(x_56, 1); +lean_inc(x_59); +lean_dec(x_56); +x_60 = l_Lean_Meta_isAbstractedUnassignedMVar(x_50, x_11, x_12, x_13, x_14, x_59); +if (lean_obj_tag(x_60) == 0) +{ +lean_object* x_61; uint8_t x_62; +x_61 = lean_ctor_get(x_60, 0); +lean_inc(x_61); +x_62 = lean_unbox(x_61); +lean_dec(x_61); +if (x_62 == 0) +{ +uint8_t x_63; +x_63 = !lean_is_exclusive(x_60); +if (x_63 == 0) +{ +lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; +x_64 = lean_ctor_get(x_60, 0); +lean_dec(x_64); +if (lean_is_scalar(x_41)) { + x_65 = lean_alloc_ctor(0, 2, 0); } else { - x_45 = x_32; + x_65 = x_41; } -lean_ctor_set(x_45, 0, x_30); -lean_ctor_set(x_45, 1, x_44); +lean_ctor_set(x_65, 0, x_39); +lean_ctor_set(x_65, 1, x_40); lean_inc(x_5); -x_46 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_46, 0, x_5); -lean_ctor_set(x_46, 1, x_45); -x_47 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_47, 0, x_46); -x_21 = x_47; -x_22 = x_43; -goto block_28; +x_66 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_66, 0, x_5); +lean_ctor_set(x_66, 1, x_65); +x_67 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_67, 0, x_66); +lean_ctor_set(x_60, 0, x_67); +x_21 = x_60; +goto block_37; } else { -lean_object* x_48; +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_60, 1); +lean_inc(x_68); +lean_dec(x_60); +if (lean_is_scalar(x_41)) { + x_69 = lean_alloc_ctor(0, 2, 0); +} else { + x_69 = x_41; +} +lean_ctor_set(x_69, 0, x_39); +lean_ctor_set(x_69, 1, x_40); +lean_inc(x_5); +x_70 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_70, 0, x_5); +lean_ctor_set(x_70, 1, x_69); +x_71 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_71, 0, x_70); +x_72 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_72, 0, x_71); +lean_ctor_set(x_72, 1, x_68); +x_21 = x_72; +goto block_37; +} +} +else +{ +lean_object* x_73; lean_object* x_74; lean_object* x_75; +lean_dec(x_41); +x_73 = lean_ctor_get(x_60, 1); +lean_inc(x_73); +lean_dec(x_60); +x_74 = lean_box(0); +lean_inc(x_5); +lean_inc(x_7); +x_75 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___lambda__1(x_7, x_39, x_5, x_40, x_74, x_11, x_12, x_13, x_14, x_73); +x_21 = x_75; +goto block_37; +} +} +else +{ +uint8_t x_76; +lean_dec(x_41); +lean_dec(x_40); +lean_dec(x_39); +x_76 = !lean_is_exclusive(x_60); +if (x_76 == 0) +{ +x_21 = x_60; +goto block_37; +} +else +{ +lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_77 = lean_ctor_get(x_60, 0); +x_78 = lean_ctor_get(x_60, 1); +lean_inc(x_78); +lean_inc(x_77); +lean_dec(x_60); +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_21 = x_79; +goto block_37; +} +} +} +else +{ +lean_object* x_80; lean_object* x_81; lean_object* x_82; +lean_dec(x_50); +lean_dec(x_41); +x_80 = lean_ctor_get(x_56, 1); +lean_inc(x_80); +lean_dec(x_56); +x_81 = lean_box(0); +lean_inc(x_5); +lean_inc(x_7); +x_82 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___lambda__1(x_7, x_39, x_5, x_40, x_81, x_11, x_12, x_13, x_14, x_80); +x_21 = x_82; +goto block_37; +} +} +else +{ +uint8_t x_83; +lean_dec(x_50); +lean_dec(x_41); +lean_dec(x_40); +lean_dec(x_39); +x_83 = !lean_is_exclusive(x_56); +if (x_83 == 0) +{ +x_21 = x_56; +goto block_37; +} +else +{ +lean_object* x_84; lean_object* x_85; lean_object* x_86; +x_84 = lean_ctor_get(x_56, 0); +x_85 = lean_ctor_get(x_56, 1); +lean_inc(x_85); +lean_inc(x_84); +lean_dec(x_56); +x_86 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_86, 0, x_84); +lean_ctor_set(x_86, 1, x_85); +x_21 = x_86; +goto block_37; +} +} +} +} +else +{ +lean_object* x_87; +lean_dec(x_47); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); -x_48 = lean_is_expr_def_eq(x_40, x_41, x_11, x_12, x_13, x_14, x_43); -if (lean_obj_tag(x_48) == 0) +x_87 = lean_is_expr_def_eq(x_49, x_50, x_11, x_12, x_13, x_14, x_52); +if (lean_obj_tag(x_87) == 0) { -lean_object* x_49; uint8_t x_50; -x_49 = lean_ctor_get(x_48, 0); -lean_inc(x_49); -x_50 = lean_unbox(x_49); -lean_dec(x_49); -if (x_50 == 0) +lean_object* x_88; uint8_t x_89; +x_88 = lean_ctor_get(x_87, 0); +lean_inc(x_88); +x_89 = lean_unbox(x_88); +lean_dec(x_88); +if (x_89 == 0) { -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_48, 1); -lean_inc(x_51); -lean_dec(x_48); -if (lean_is_scalar(x_32)) { - x_52 = lean_alloc_ctor(0, 2, 0); +uint8_t x_90; +x_90 = !lean_is_exclusive(x_87); +if (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_87, 0); +lean_dec(x_91); +if (lean_is_scalar(x_41)) { + x_92 = lean_alloc_ctor(0, 2, 0); } else { - x_52 = x_32; + x_92 = x_41; } -lean_ctor_set(x_52, 0, x_30); -lean_ctor_set(x_52, 1, x_31); -x_53 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___closed__1; -x_54 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_54, 0, x_53); -lean_ctor_set(x_54, 1, x_52); -x_55 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_55, 0, x_54); -x_21 = x_55; -x_22 = x_51; -goto block_28; +lean_ctor_set(x_92, 0, x_39); +lean_ctor_set(x_92, 1, x_40); +x_93 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___lambda__2___closed__1; +x_94 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_94, 0, x_93); +lean_ctor_set(x_94, 1, x_92); +x_95 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_95, 0, x_94); +lean_ctor_set(x_87, 0, x_95); +x_21 = x_87; +goto block_37; } else { -lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_56 = lean_ctor_get(x_48, 1); -lean_inc(x_56); -lean_dec(x_48); -if (lean_is_scalar(x_32)) { - x_57 = lean_alloc_ctor(0, 2, 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; +x_96 = lean_ctor_get(x_87, 1); +lean_inc(x_96); +lean_dec(x_87); +if (lean_is_scalar(x_41)) { + x_97 = lean_alloc_ctor(0, 2, 0); +} else { + x_97 = x_41; +} +lean_ctor_set(x_97, 0, x_39); +lean_ctor_set(x_97, 1, x_40); +x_98 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___lambda__2___closed__1; +x_99 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_99, 0, x_98); +lean_ctor_set(x_99, 1, x_97); +x_100 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_100, 0, 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); +x_21 = x_101; +goto block_37; +} +} +else +{ +uint8_t x_102; +x_102 = !lean_is_exclusive(x_87); +if (x_102 == 0) +{ +lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; +x_103 = lean_ctor_get(x_87, 0); +lean_dec(x_103); +if (lean_is_scalar(x_41)) { + x_104 = lean_alloc_ctor(0, 2, 0); } else { - x_57 = x_32; + x_104 = x_41; } -lean_ctor_set(x_57, 0, x_30); -lean_ctor_set(x_57, 1, x_31); +lean_ctor_set(x_104, 0, x_39); +lean_ctor_set(x_104, 1, x_40); lean_inc(x_5); -x_58 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_58, 0, x_5); -lean_ctor_set(x_58, 1, x_57); -x_59 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_59, 0, x_58); -x_21 = x_59; -x_22 = x_56; -goto block_28; +x_105 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_105, 0, x_5); +lean_ctor_set(x_105, 1, x_104); +x_106 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_106, 0, x_105); +lean_ctor_set(x_87, 0, x_106); +x_21 = x_87; +goto block_37; +} +else +{ +lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; +x_107 = lean_ctor_get(x_87, 1); +lean_inc(x_107); +lean_dec(x_87); +if (lean_is_scalar(x_41)) { + x_108 = lean_alloc_ctor(0, 2, 0); +} else { + x_108 = x_41; +} +lean_ctor_set(x_108, 0, x_39); +lean_ctor_set(x_108, 1, x_40); +lean_inc(x_5); +x_109 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_109, 0, x_5); +lean_ctor_set(x_109, 1, x_108); +x_110 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_110, 0, x_109); +x_111 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_111, 0, x_110); +lean_ctor_set(x_111, 1, x_107); +x_21 = x_111; +goto block_37; +} } } else { -uint8_t x_60; -lean_dec(x_32); -lean_dec(x_31); -lean_dec(x_30); -lean_dec(x_20); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_7); -lean_dec(x_5); -x_60 = !lean_is_exclusive(x_48); -if (x_60 == 0) +uint8_t x_112; +lean_dec(x_41); +lean_dec(x_40); +lean_dec(x_39); +x_112 = !lean_is_exclusive(x_87); +if (x_112 == 0) { -return x_48; +x_21 = x_87; +goto block_37; } else { -lean_object* x_61; lean_object* x_62; lean_object* x_63; -x_61 = lean_ctor_get(x_48, 0); -x_62 = lean_ctor_get(x_48, 1); -lean_inc(x_62); -lean_inc(x_61); -lean_dec(x_48); -x_63 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_63, 0, x_61); -lean_ctor_set(x_63, 1, x_62); -return x_63; +lean_object* x_113; lean_object* x_114; lean_object* x_115; +x_113 = lean_ctor_get(x_87, 0); +x_114 = lean_ctor_get(x_87, 1); +lean_inc(x_114); +lean_inc(x_113); +lean_dec(x_87); +x_115 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_115, 0, x_113); +lean_ctor_set(x_115, 1, x_114); +x_21 = x_115; +goto block_37; } } } @@ -6664,7 +7616,7 @@ return x_63; } else { -lean_object* x_169; +lean_object* x_235; lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); @@ -6672,15 +7624,15 @@ lean_dec(x_11); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_169 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_169, 0, x_10); -lean_ctor_set(x_169, 1, x_15); -return x_169; +x_235 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_235, 0, x_10); +lean_ctor_set(x_235, 1, x_15); +return x_235; } } else { -lean_object* x_170; +lean_object* x_236; lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); @@ -6688,10 +7640,10 @@ lean_dec(x_11); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_170 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_170, 0, x_10); -lean_ctor_set(x_170, 1, x_15); -return x_170; +x_236 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_236, 0, x_10); +lean_ctor_set(x_236, 1, x_15); +return x_236; } } } @@ -6865,6 +7817,28 @@ lean_dec(x_5); return x_11; } } +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__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, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___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_6); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___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: +{ +lean_object* x_11; +x_11 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__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_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +return x_11; +} +} LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___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) { _start: { @@ -10440,15 +11414,15 @@ LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_Exp _start: { lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_13 = l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Meta_processPostponed___spec__7(x_1, x_2, x_3, x_6, x_4, x_8, x_9, x_10, x_11, x_12); +x_13 = l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Meta_processPostponed___spec__5(x_1, x_2, x_3, x_6, x_4, x_8, x_9, x_10, x_11, x_12); x_14 = lean_ctor_get(x_13, 1); lean_inc(x_14); lean_dec(x_13); -x_15 = l_MonadExcept_ofExcept___at_Lean_Meta_processPostponed___spec__8(x_5, x_8, x_9, x_10, x_11, x_14); +x_15 = l_MonadExcept_ofExcept___at_Lean_Meta_processPostponed___spec__6(x_5, x_8, x_9, x_10, x_11, x_14); return x_15; } } -static lean_object* _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___closed__1() { +static lean_object* _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__1() { _start: { lean_object* x_1; @@ -10456,16 +11430,24 @@ x_1 = lean_mk_string_from_bytes(" ", 1); return x_1; } } -static lean_object* _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___closed__2() { +static lean_object* _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___closed__1; +x_1 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__1; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___closed__3() { +static lean_object* _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_profiler; +return x_1; +} +} +static lean_object* _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__4() { _start: { lean_object* x_1; @@ -10473,16 +11455,16 @@ x_1 = lean_mk_string_from_bytes("[", 1); return x_1; } } -static lean_object* _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___closed__4() { +static lean_object* _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___closed__3; +x_1 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__4; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___closed__5() { +static lean_object* _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__6() { _start: { lean_object* x_1; @@ -10490,251 +11472,574 @@ x_1 = lean_mk_string_from_bytes("s] ", 3); return x_1; } } -static lean_object* _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___closed__6() { +static lean_object* _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__7() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___closed__5; +x_1 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__6; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___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_EXPORT lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3(lean_object* 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, uint8_t x_8, double 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_10; lean_object* x_11; uint8_t x_12; -lean_inc(x_1); -x_10 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_1, 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_unbox(x_11); -lean_dec(x_11); -if (x_12 == 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; uint8_t x_25; +x_16 = l_Lean_exceptBoolEmoji___rarg(x_1); +x_17 = l_Lean_stringToMessageData(x_16); +lean_dec(x_16); +x_18 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__2___closed__9; +x_19 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_17); +x_20 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__2; +x_21 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); +x_22 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_2); +x_23 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_18); +x_24 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__3; +x_25 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_7, x_24); +lean_dec(x_7); +if (x_25 == 0) { -lean_object* x_13; lean_object* x_14; -lean_dec(x_2); +if (x_8 == 0) +{ +lean_object* x_26; lean_object* x_27; +x_26 = lean_box(0); +x_27 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2(x_3, x_4, x_5, x_6, x_1, x_23, x_26, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_12); lean_dec(x_1); -x_13 = lean_ctor_get(x_10, 1); -lean_inc(x_13); -lean_dec(x_10); -x_14 = lean_apply_5(x_3, x_5, x_6, x_7, x_8, x_13); -return x_14; +return x_27; } 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_31; lean_object* x_32; -x_15 = lean_ctor_get(x_10, 1); +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_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_28 = lean_float_to_string(x_9); +x_29 = lean_alloc_ctor(3, 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); +x_31 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__5; +x_32 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_30); +x_33 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__7; +x_34 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +x_35 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_23); +x_36 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_18); +x_37 = lean_box(0); +x_38 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2(x_3, x_4, x_5, x_6, x_1, x_36, x_37, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_12); +lean_dec(x_1); +return x_38; +} +} +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; +x_39 = lean_float_to_string(x_9); +x_40 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_40, 0, x_39); +x_41 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_41, 0, x_40); +x_42 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__5; +x_43 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_41); +x_44 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__7; +x_45 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +x_46 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_23); +x_47 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_18); +x_48 = lean_box(0); +x_49 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2(x_3, x_4, x_5, x_6, x_1, x_47, x_48, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_12); +lean_dec(x_1); +return x_49; +} +} +} +static lean_object* _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__4___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_trace_profiler; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t 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, 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_dec(x_7); +x_13 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_processPostponed___spec__2___rarg(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_10); -x_16 = lean_ctor_get(x_7, 0); +lean_dec(x_13); +x_16 = lean_ctor_get(x_10, 0); lean_inc(x_16); -x_17 = lean_ctor_get(x_7, 1); +x_17 = lean_ctor_get(x_10, 1); lean_inc(x_17); -x_18 = lean_ctor_get(x_7, 2); +x_18 = lean_ctor_get(x_10, 2); lean_inc(x_18); -x_19 = lean_ctor_get(x_7, 3); +x_19 = lean_ctor_get(x_10, 3); lean_inc(x_19); -x_20 = lean_ctor_get(x_7, 4); +x_20 = lean_ctor_get(x_10, 4); lean_inc(x_20); -x_21 = lean_ctor_get(x_7, 5); +x_21 = lean_ctor_get(x_10, 5); lean_inc(x_21); -x_22 = lean_ctor_get(x_7, 6); +x_22 = lean_ctor_get(x_10, 6); lean_inc(x_22); -x_23 = lean_ctor_get(x_7, 7); +x_23 = lean_ctor_get(x_10, 7); lean_inc(x_23); -x_24 = lean_ctor_get(x_7, 8); +x_24 = lean_ctor_get(x_10, 8); lean_inc(x_24); -x_25 = lean_ctor_get(x_7, 9); +x_25 = lean_ctor_get(x_10, 9); lean_inc(x_25); -x_26 = lean_ctor_get(x_7, 10); +x_26 = lean_ctor_get(x_10, 10); lean_inc(x_26); -x_27 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_processPostponed___spec__2___rarg(x_8, x_15); -x_28 = lean_ctor_get(x_27, 0); +x_27 = l_Lean_replaceRef(x_21, 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); +lean_inc(x_11); lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); -lean_inc(x_29); -lean_dec(x_27); -x_30 = l_Lean_replaceRef(x_21, x_21); -x_31 = lean_alloc_ctor(0, 11, 0); -lean_ctor_set(x_31, 0, x_16); -lean_ctor_set(x_31, 1, x_17); -lean_ctor_set(x_31, 2, x_18); -lean_ctor_set(x_31, 3, x_19); -lean_ctor_set(x_31, 4, x_20); -lean_ctor_set(x_31, 5, x_30); -lean_ctor_set(x_31, 6, x_22); -lean_ctor_set(x_31, 7, x_23); -lean_ctor_set(x_31, 8, x_24); -lean_ctor_set(x_31, 9, x_25); -lean_ctor_set(x_31, 10, x_26); +lean_inc(x_9); lean_inc(x_8); -lean_inc(x_31); -lean_inc(x_6); -lean_inc(x_5); -x_32 = lean_apply_5(x_2, x_5, x_6, x_31, x_8, x_29); -if (lean_obj_tag(x_32) == 0) +x_29 = lean_apply_5(x_1, x_8, x_9, x_28, x_11, x_15); +if (lean_obj_tag(x_29) == 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_30; 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_29, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_29, 1); +lean_inc(x_31); +lean_dec(x_29); +x_32 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_30, x_8, x_9, x_28, x_11, x_31); +lean_dec(x_28); 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_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_33, x_5, x_6, x_31, x_8, x_34); -lean_dec(x_31); -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_alloc_closure((void*)(l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__1), 6, 1); -lean_closure_set(x_38, 0, x_3); +x_35 = lean_alloc_closure((void*)(l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__1), 6, 1); +lean_closure_set(x_35, 0, x_2); +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_39 = l_Lean_withOptProfile___at_Lean_Meta_processPostponed___spec__3(x_38, x_5, x_6, x_7, x_8, x_37); -if (lean_obj_tag(x_39) == 0) +x_36 = l_Lean_withSeconds___at_Lean_Meta_processPostponed___spec__3(x_35, x_8, x_9, x_10, x_11, x_34); +if (lean_obj_tag(x_36) == 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; -x_40 = lean_ctor_get(x_39, 0); +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; lean_object* x_86; uint8_t x_87; +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_37, 0); +lean_inc(x_39); +x_40 = lean_ctor_get(x_37, 1); lean_inc(x_40); -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); -x_43 = lean_ctor_get(x_40, 1); -lean_inc(x_43); -lean_dec(x_40); -x_44 = l_Lean_exceptBoolEmoji___rarg(x_42); -x_45 = l_Lean_stringToMessageData(x_44); -lean_dec(x_44); -x_46 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__2___closed__9; -x_47 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_45); -x_48 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___closed__2; -x_49 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_49, 0, x_47); -lean_ctor_set(x_49, 1, x_48); -x_50 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_50, 0, x_49); -lean_ctor_set(x_50, 1, x_36); -x_51 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_51, 0, x_50); -lean_ctor_set(x_51, 1, x_46); -if (lean_obj_tag(x_43) == 0) +lean_dec(x_37); +x_86 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__4___closed__1; +x_87 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_5, x_86); +if (x_87 == 0) { -lean_object* x_52; lean_object* x_53; -x_52 = lean_box(0); -x_53 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2(x_28, x_1, x_21, x_4, x_42, x_51, x_52, x_5, x_6, x_7, x_8, x_41); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); +uint8_t x_88; +x_88 = 0; +x_41 = x_88; +goto block_85; +} +else +{ +double x_89; double x_90; uint8_t x_91; +x_89 = l_Lean_trace_profiler_threshold_getSecs(x_5); +x_90 = lean_unbox_float(x_40); +x_91 = lean_float_decLt(x_89, x_90); +x_41 = x_91; +goto block_85; +} +block_85: +{ +if (x_6 == 0) +{ +if (x_41 == 0) +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; +lean_dec(x_40); +lean_dec(x_33); +lean_dec(x_21); lean_dec(x_5); +lean_dec(x_3); +x_42 = lean_st_ref_take(x_11, x_38); +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); -return x_53; +x_45 = !lean_is_exclusive(x_43); +if (x_45 == 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, 3); +x_47 = l_Lean_PersistentArray_append___rarg(x_14, x_46); +lean_ctor_set(x_43, 3, x_47); +x_48 = lean_st_ref_set(x_11, x_43, x_44); +x_49 = lean_ctor_get(x_48, 1); +lean_inc(x_49); +lean_dec(x_48); +x_50 = l_MonadExcept_ofExcept___at_Lean_Meta_processPostponed___spec__6(x_39, x_8, x_9, x_10, x_11, x_49); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_39); +if (lean_obj_tag(x_50) == 0) +{ +uint8_t x_51; +x_51 = !lean_is_exclusive(x_50); +if (x_51 == 0) +{ +return x_50; } else { -lean_object* x_54; double 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_54 = lean_ctor_get(x_43, 0); -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_inc(x_52); +lean_dec(x_50); +x_54 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_53); +return x_54; +} +} +else +{ +uint8_t x_55; +x_55 = !lean_is_exclusive(x_50); +if (x_55 == 0) +{ +return x_50; +} +else +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_56 = lean_ctor_get(x_50, 0); +x_57 = lean_ctor_get(x_50, 1); +lean_inc(x_57); +lean_inc(x_56); +lean_dec(x_50); +x_58 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_58, 0, x_56); +lean_ctor_set(x_58, 1, x_57); +return x_58; +} +} +} +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_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_59 = lean_ctor_get(x_43, 0); +x_60 = lean_ctor_get(x_43, 1); +x_61 = lean_ctor_get(x_43, 2); +x_62 = lean_ctor_get(x_43, 3); +x_63 = lean_ctor_get(x_43, 4); +x_64 = lean_ctor_get(x_43, 5); +x_65 = lean_ctor_get(x_43, 6); +lean_inc(x_65); +lean_inc(x_64); +lean_inc(x_63); +lean_inc(x_62); +lean_inc(x_61); +lean_inc(x_60); +lean_inc(x_59); lean_dec(x_43); -x_55 = lean_unbox_float(x_54); -lean_dec(x_54); -x_56 = lean_float_to_string(x_55); -x_57 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_57, 0, x_56); -x_58 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_58, 0, x_57); -x_59 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___closed__4; -x_60 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_60, 0, x_59); -lean_ctor_set(x_60, 1, x_58); -x_61 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___closed__6; -x_62 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_62, 0, x_60); -lean_ctor_set(x_62, 1, x_61); -x_63 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_63, 0, x_62); -lean_ctor_set(x_63, 1, x_51); -x_64 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_64, 0, x_63); -lean_ctor_set(x_64, 1, x_46); -x_65 = lean_box(0); -x_66 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2(x_28, x_1, x_21, x_4, x_42, x_64, x_65, x_5, x_6, x_7, x_8, x_41); +x_66 = l_Lean_PersistentArray_append___rarg(x_14, x_62); +x_67 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_67, 0, x_59); +lean_ctor_set(x_67, 1, x_60); +lean_ctor_set(x_67, 2, x_61); +lean_ctor_set(x_67, 3, x_66); +lean_ctor_set(x_67, 4, x_63); +lean_ctor_set(x_67, 5, x_64); +lean_ctor_set(x_67, 6, x_65); +x_68 = lean_st_ref_set(x_11, x_67, x_44); +x_69 = lean_ctor_get(x_68, 1); +lean_inc(x_69); +lean_dec(x_68); +x_70 = l_MonadExcept_ofExcept___at_Lean_Meta_processPostponed___spec__6(x_39, x_8, x_9, x_10, x_11, x_69); +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_42); -return x_66; +lean_dec(x_39); +if (lean_obj_tag(x_70) == 0) +{ +lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_71 = lean_ctor_get(x_70, 0); +lean_inc(x_71); +x_72 = lean_ctor_get(x_70, 1); +lean_inc(x_72); +if (lean_is_exclusive(x_70)) { + lean_ctor_release(x_70, 0); + lean_ctor_release(x_70, 1); + x_73 = x_70; +} else { + lean_dec_ref(x_70); + x_73 = lean_box(0); } +if (lean_is_scalar(x_73)) { + x_74 = lean_alloc_ctor(0, 2, 0); +} else { + x_74 = x_73; +} +lean_ctor_set(x_74, 0, x_71); +lean_ctor_set(x_74, 1, x_72); +return x_74; } else { -uint8_t x_67; -lean_dec(x_36); -lean_dec(x_28); +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_75 = lean_ctor_get(x_70, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_70, 1); +lean_inc(x_76); +if (lean_is_exclusive(x_70)) { + lean_ctor_release(x_70, 0); + lean_ctor_release(x_70, 1); + x_77 = x_70; +} else { + lean_dec_ref(x_70); + x_77 = lean_box(0); +} +if (lean_is_scalar(x_77)) { + x_78 = lean_alloc_ctor(1, 2, 0); +} else { + x_78 = x_77; +} +lean_ctor_set(x_78, 0, x_75); +lean_ctor_set(x_78, 1, x_76); +return x_78; +} +} +} +else +{ +lean_object* x_79; double x_80; lean_object* x_81; +x_79 = lean_box(0); +x_80 = lean_unbox_float(x_40); +lean_dec(x_40); +x_81 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3(x_39, x_33, x_14, x_3, x_21, x_4, x_5, x_41, x_80, x_79, x_8, x_9, x_10, x_11, x_38); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_8); +return x_81; +} +} +else +{ +lean_object* x_82; double x_83; lean_object* x_84; +x_82 = lean_box(0); +x_83 = lean_unbox_float(x_40); +lean_dec(x_40); +x_84 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3(x_39, x_33, x_14, x_3, x_21, x_4, x_5, x_41, x_83, x_82, x_8, x_9, x_10, x_11, x_38); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_8); +return x_84; +} +} +} +else +{ +uint8_t x_92; +lean_dec(x_33); lean_dec(x_21); +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_5); -lean_dec(x_1); -x_67 = !lean_is_exclusive(x_39); -if (x_67 == 0) +lean_dec(x_3); +x_92 = !lean_is_exclusive(x_36); +if (x_92 == 0) { -return x_39; +return x_36; } else { -lean_object* x_68; lean_object* x_69; lean_object* x_70; -x_68 = lean_ctor_get(x_39, 0); -x_69 = lean_ctor_get(x_39, 1); -lean_inc(x_69); -lean_inc(x_68); -lean_dec(x_39); -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_object* x_93; lean_object* x_94; lean_object* x_95; +x_93 = lean_ctor_get(x_36, 0); +x_94 = lean_ctor_get(x_36, 1); +lean_inc(x_94); +lean_inc(x_93); +lean_dec(x_36); +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_71; -lean_dec(x_31); +uint8_t x_96; lean_dec(x_28); lean_dec(x_21); +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_5); lean_dec(x_3); +lean_dec(x_2); +x_96 = !lean_is_exclusive(x_29); +if (x_96 == 0) +{ +return x_29; +} +else +{ +lean_object* x_97; lean_object* x_98; lean_object* x_99; +x_97 = lean_ctor_get(x_29, 0); +x_98 = lean_ctor_get(x_29, 1); +lean_inc(x_98); +lean_inc(x_97); +lean_dec(x_29); +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_EXPORT lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___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) { +_start: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_10 = lean_ctor_get(x_7, 2); +lean_inc(x_10); +lean_inc(x_1); +x_11 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___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_unbox(x_12); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); +lean_dec(x_11); +x_15 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__4___closed__1; +x_16 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_10, x_15); +if (x_16 == 0) +{ +lean_object* x_17; +lean_dec(x_12); +lean_dec(x_10); +lean_dec(x_2); lean_dec(x_1); -x_71 = !lean_is_exclusive(x_32); -if (x_71 == 0) +x_17 = lean_apply_5(x_3, x_5, x_6, x_7, x_8, x_14); +if (lean_obj_tag(x_17) == 0) { -return x_32; +uint8_t x_18; +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) +{ +return x_17; } else { -lean_object* x_72; lean_object* x_73; lean_object* x_74; -x_72 = lean_ctor_get(x_32, 0); -x_73 = lean_ctor_get(x_32, 1); -lean_inc(x_73); -lean_inc(x_72); -lean_dec(x_32); -x_74 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_74, 0, x_72); -lean_ctor_set(x_74, 1, x_73); -return x_74; +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_17, 0); +x_20 = lean_ctor_get(x_17, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_17); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); +return x_21; +} } +else +{ +uint8_t x_22; +x_22 = !lean_is_exclusive(x_17); +if (x_22 == 0) +{ +return x_17; } +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_17, 0); +x_24 = lean_ctor_get(x_17, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_17); +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_26; uint8_t x_27; lean_object* x_28; +x_26 = lean_box(0); +x_27 = lean_unbox(x_12); +lean_dec(x_12); +x_28 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__4(x_2, x_3, x_1, x_4, x_10, x_27, x_26, x_5, x_6, x_7, x_8, x_14); +return x_28; +} +} +else +{ +lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; +x_29 = lean_ctor_get(x_11, 1); +lean_inc(x_29); +lean_dec(x_11); +x_30 = lean_box(0); +x_31 = lean_unbox(x_12); +lean_dec(x_12); +x_32 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__4(x_2, x_3, x_1, x_4, x_10, x_31, x_30, x_5, x_6, x_7, x_8, x_29); +return x_32; } } } @@ -11456,8 +12761,8 @@ static lean_object* _init_l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkType _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__1; -x_2 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__2; +x_1 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__1; +x_2 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__2; x_3 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___closed__1; x_4 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___closed__2; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); @@ -11515,6 +12820,36 @@ lean_dec(x_5); return x_14; } } +LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___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: +{ +uint8_t x_16; uint8_t x_17; double x_18; lean_object* x_19; +x_16 = lean_unbox(x_6); +lean_dec(x_6); +x_17 = lean_unbox(x_8); +lean_dec(x_8); +x_18 = lean_unbox_float(x_9); +lean_dec(x_9); +x_19 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3(x_1, x_2, x_3, x_4, x_5, x_16, x_7, x_17, x_18, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_11); +lean_dec(x_10); +return x_19; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___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) { +_start: +{ +uint8_t x_13; uint8_t x_14; lean_object* x_15; +x_13 = lean_unbox(x_4); +lean_dec(x_4); +x_14 = lean_unbox(x_6); +lean_dec(x_6); +x_15 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__4(x_1, x_2, x_3, x_13, x_5, x_14, x_7, x_8, x_9, x_10, x_11, x_12); +return x_15; +} +} LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___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) { _start: { @@ -15436,7 +16771,7 @@ lean_dec(x_5); return x_10; } } -static lean_object* _init_l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6051____closed__1() { +static lean_object* _init_l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6556____closed__1() { _start: { lean_object* x_1; @@ -15444,26 +16779,26 @@ x_1 = lean_mk_string_from_bytes("checkAssignment", 15); return x_1; } } -static lean_object* _init_l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6051____closed__2() { +static lean_object* _init_l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6556____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6051____closed__1; +x_2 = l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6556____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6051_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6556_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6051____closed__2; +x_2 = l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6556____closed__2; x_3 = l_Lean_registerInternalExceptionId(x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6071____closed__1() { +static lean_object* _init_l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6576____closed__1() { _start: { lean_object* x_1; @@ -15471,21 +16806,21 @@ x_1 = lean_mk_string_from_bytes("outOfScope", 10); return x_1; } } -static lean_object* _init_l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6071____closed__2() { +static lean_object* _init_l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6576____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6071____closed__1; +x_2 = l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6576____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6071_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6576_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6071____closed__2; +x_2 = l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6576____closed__2; x_3 = l_Lean_registerInternalExceptionId(x_2, x_1); return x_3; } @@ -15796,7 +17131,7 @@ x_15 = l_Lean_MessageData_ofExpr(x_14); x_16 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_16, 0, x_12); lean_ctor_set(x_16, 1, x_15); -x_17 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___closed__2; +x_17 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__2; x_18 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_18, 0, x_16); lean_ctor_set(x_18, 1, x_17); @@ -16348,8 +17683,8 @@ static lean_object* _init_l_Lean_Meta_CheckAssignment_checkFVar___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__1; -x_2 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__2; +x_1 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__1; +x_2 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__2; x_3 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___closed__1; x_4 = l_Lean_Meta_CheckAssignment_checkFVar___closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); @@ -39068,8 +40403,8 @@ static lean_object* _init_l_Lean_Meta_CheckAssignment_checkMVar___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__1; -x_2 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__2; +x_1 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__1; +x_2 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__2; x_3 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___closed__1; x_4 = l_Lean_Meta_CheckAssignment_checkMVar___closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); @@ -39088,8 +40423,8 @@ static lean_object* _init_l_Lean_Meta_CheckAssignment_checkMVar___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_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__1; -x_2 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__2; +x_1 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__1; +x_2 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__2; x_3 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___closed__1; x_4 = l_Lean_Meta_CheckAssignment_checkMVar___closed__3; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); @@ -43486,7 +44821,7 @@ if (x_26 == 0) { lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_free_object(x_19); -x_27 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_4, x_8, x_9, x_10, x_11, x_22); +x_27 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_4, x_8, x_9, x_10, x_11, x_22); x_28 = lean_ctor_get(x_27, 0); lean_inc(x_28); x_29 = lean_ctor_get(x_27, 1); @@ -43529,7 +44864,7 @@ x_37 = l_Lean_Meta_CheckAssignmentQuick_check_visit(x_36, x_34, x_35, x_15, x_1, if (x_37 == 0) { lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_38 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_4, x_8, x_9, x_10, x_11, x_33); +x_38 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_4, x_8, x_9, x_10, x_11, x_33); x_39 = lean_ctor_get(x_38, 0); lean_inc(x_39); x_40 = lean_ctor_get(x_38, 1); @@ -43581,7 +44916,7 @@ if (x_51 == 0) { lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_free_object(x_19); -x_52 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_4, x_8, x_9, x_10, x_11, x_47); +x_52 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_4, x_8, x_9, x_10, x_11, x_47); x_53 = lean_ctor_get(x_52, 0); lean_inc(x_53); x_54 = lean_ctor_get(x_52, 1); @@ -43624,7 +44959,7 @@ x_62 = l_Lean_Meta_CheckAssignmentQuick_check_visit(x_61, x_59, x_60, x_15, x_1, if (x_62 == 0) { lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_63 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_4, x_8, x_9, x_10, x_11, x_58); +x_63 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_4, x_8, x_9, x_10, x_11, x_58); x_64 = lean_ctor_get(x_63, 0); lean_inc(x_64); x_65 = lean_ctor_get(x_63, 1); @@ -43673,7 +45008,7 @@ if (x_75 == 0) { lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_free_object(x_19); -x_76 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_4, x_8, x_9, x_10, x_11, x_71); +x_76 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_4, x_8, x_9, x_10, x_11, x_71); x_77 = lean_ctor_get(x_76, 0); lean_inc(x_77); x_78 = lean_ctor_get(x_76, 1); @@ -43717,7 +45052,7 @@ x_86 = l_Lean_Meta_CheckAssignmentQuick_check_visit(x_83, x_84, x_85, x_15, x_1, if (x_86 == 0) { lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; -x_87 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_4, x_8, x_9, x_10, x_11, x_82); +x_87 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_4, x_8, x_9, x_10, x_11, x_82); x_88 = lean_ctor_get(x_87, 0); lean_inc(x_88); x_89 = lean_ctor_get(x_87, 1); @@ -45078,8 +46413,8 @@ static lean_object* _init_l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processAs _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___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__1; -x_2 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__2; +x_1 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__1; +x_2 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__2; x_3 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processAssignmentFOApprox_loop___closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; @@ -45163,7 +46498,7 @@ x_30 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_is x_31 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_31, 0, x_30); lean_ctor_set(x_31, 1, x_29); -x_32 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___closed__2; +x_32 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__2; x_33 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_33, 0, x_31); lean_ctor_set(x_33, 1, x_32); @@ -45357,7 +46692,7 @@ return x_10; else { lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_11 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_1, x_2, x_3, x_4, x_5, x_6); +x_11 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_1, x_2, x_3, x_4, x_5, x_6); x_12 = lean_ctor_get(x_11, 0); lean_inc(x_12); x_13 = lean_ctor_get(x_11, 1); @@ -45388,8 +46723,8 @@ static lean_object* _init_l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_assignCon _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___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__1; -x_2 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__2; +x_1 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__1; +x_2 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__2; x_3 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_assignConst___lambda__2___closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; @@ -46347,7 +47682,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processCon _start: { lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_13 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___closed__2; +x_13 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___closed__1; x_14 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_13, x_8, x_9, x_10, x_11, x_12); x_15 = lean_ctor_get(x_14, 0); lean_inc(x_15); @@ -46375,7 +47710,7 @@ x_22 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processConstApprox_go___clos x_23 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_23, 0, x_22); lean_ctor_set(x_23, 1, x_21); -x_24 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___closed__2; +x_24 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__2; x_25 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_25, 0, x_23); lean_ctor_set(x_25, 1, x_24); @@ -47031,7 +48366,7 @@ x_19 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_is x_20 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_20, 0, x_19); lean_ctor_set(x_20, 1, x_18); -x_21 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___closed__2; +x_21 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__2; x_22 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_22, 0, x_20); lean_ctor_set(x_22, 1, x_21); @@ -47170,8 +48505,8 @@ static lean_object* _init_l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processAs _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__1; -x_2 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__2; +x_1 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__1; +x_2 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__2; x_3 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___closed__1; x_4 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processAssignment_process___lambda__1___closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); @@ -47573,8 +48908,8 @@ static lean_object* _init_l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processAs _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__1; -x_2 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__2; +x_1 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__1; +x_2 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__2; x_3 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___closed__1; x_4 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processAssignment_process___closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); @@ -47598,7 +48933,7 @@ if (x_15 == 0) lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; lean_dec(x_12); lean_inc(x_5); -x_16 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_5, x_6, x_7, x_8, x_9, x_13); +x_16 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_5, x_6, x_7, x_8, x_9, x_13); x_17 = lean_ctor_get(x_16, 0); lean_inc(x_17); x_18 = lean_ctor_get(x_16, 1); @@ -47756,7 +49091,7 @@ x_49 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_is x_50 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_50, 0, x_49); lean_ctor_set(x_50, 1, x_48); -x_51 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___closed__2; +x_51 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__2; x_52 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_52, 0, x_50); lean_ctor_set(x_52, 1, x_51); @@ -48473,8 +49808,8 @@ static lean_object* _init_l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processAs _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___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__1; -x_2 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__2; +x_1 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__1; +x_2 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__2; x_3 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; @@ -49053,8 +50388,8 @@ static lean_object* _init_l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqLe _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__1; -x_2 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__2; +x_1 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__1; +x_2 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__2; x_3 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqLeft___closed__1; x_4 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqLeft___closed__2; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); @@ -49122,8 +50457,8 @@ static lean_object* _init_l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqRi _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__1; -x_2 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__2; +x_1 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__1; +x_2 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__2; x_3 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqLeft___closed__1; x_4 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqRight___closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); @@ -49182,8 +50517,8 @@ static lean_object* _init_l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqLe _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__1; -x_2 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__2; +x_1 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__1; +x_2 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__2; x_3 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqLeft___closed__1; x_4 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqLeftRight___closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); @@ -50140,8 +51475,8 @@ static lean_object* _init_l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_tryHeuris _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___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__1; -x_2 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__2; +x_1 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__1; +x_2 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__2; x_3 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqLeft___closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; @@ -50157,7 +51492,7 @@ x_12 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_is x_13 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_13, 0, x_12); lean_ctor_set(x_13, 1, x_11); -x_14 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__7; +x_14 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__7; x_15 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_15, 0, x_13); lean_ctor_set(x_15, 1, x_14); @@ -52855,7 +54190,7 @@ x_22 = lean_ctor_get(x_19, 1); lean_inc(x_22); lean_dec(x_19); lean_inc(x_2); -x_23 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_2, x_4, x_5, x_6, x_7, x_20); +x_23 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_2, x_4, x_5, x_6, x_7, x_20); x_24 = !lean_is_exclusive(x_23); if (x_24 == 0) { @@ -52946,7 +54281,7 @@ x_39 = lean_ctor_get(x_36, 1); lean_inc(x_39); lean_dec(x_36); lean_inc(x_2); -x_40 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_2, x_4, x_5, x_6, x_7, x_37); +x_40 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_2, x_4, x_5, 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); @@ -57224,8 +58559,8 @@ static lean_object* _init_l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqQu _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___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__1; -x_2 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__2; +x_1 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__1; +x_2 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__2; x_3 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqQuickOther___lambda__2___closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; @@ -57405,7 +58740,7 @@ x_32 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_is x_33 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_33, 0, x_32); lean_ctor_set(x_33, 1, x_31); -x_34 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__7; +x_34 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__7; x_35 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_35, 0, x_33); lean_ctor_set(x_35, 1, x_34); @@ -57887,7 +59222,7 @@ lean_inc(x_36); x_37 = lean_ctor_get(x_35, 1); lean_inc(x_37); lean_dec(x_35); -x_38 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___closed__2; +x_38 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___closed__1; x_39 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_38, x_3, x_4, x_5, x_6, x_37); x_40 = lean_ctor_get(x_39, 0); lean_inc(x_40); @@ -57920,7 +59255,7 @@ x_49 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_is x_50 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_50, 0, x_49); lean_ctor_set(x_50, 1, x_48); -x_51 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___closed__2; +x_51 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__2; x_52 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_52, 0, x_50); lean_ctor_set(x_52, 1, x_51); @@ -57934,7 +59269,7 @@ x_55 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqQuickOther___closed__ x_56 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_56, 0, x_52); lean_ctor_set(x_56, 1, x_55); -x_57 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__7; +x_57 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__7; x_58 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_58, 0, x_56); lean_ctor_set(x_58, 1, x_57); @@ -58002,7 +59337,7 @@ x_79 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqQuickOther___closed__ x_80 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_80, 0, x_52); lean_ctor_set(x_80, 1, x_79); -x_81 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__7; +x_81 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__7; x_82 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_82, 0, x_80); lean_ctor_set(x_82, 1, x_81); @@ -58151,7 +59486,7 @@ lean_dec(x_13); x_149 = lean_ctor_get(x_21, 1); lean_inc(x_149); lean_dec(x_21); -x_150 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_2, x_3, x_4, x_5, x_6, x_149); +x_150 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_2, x_3, x_4, x_5, x_6, x_149); x_151 = lean_ctor_get(x_150, 0); lean_inc(x_151); x_152 = lean_ctor_get(x_150, 1); @@ -58169,7 +59504,7 @@ lean_dec(x_13); x_154 = lean_ctor_get(x_17, 1); lean_inc(x_154); lean_dec(x_17); -x_155 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_1, x_3, x_4, x_5, x_6, x_154); +x_155 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_1, x_3, x_4, x_5, x_6, x_154); x_156 = lean_ctor_get(x_155, 0); lean_inc(x_156); x_157 = lean_ctor_get(x_155, 1); @@ -58531,8 +59866,8 @@ static lean_object* _init_l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_unstuckMV _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___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__1; -x_2 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__2; +x_1 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__1; +x_2 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__2; x_3 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_unstuckMVar___closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; @@ -59081,7 +60416,7 @@ lean_dec(x_3); x_28 = lean_ctor_get(x_11, 1); lean_inc(x_28); lean_dec(x_11); -x_29 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_4, x_6, x_7, x_8, x_9, x_28); +x_29 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_4, x_6, x_7, x_8, x_9, x_28); x_30 = lean_ctor_get(x_29, 0); lean_inc(x_30); x_31 = lean_ctor_get(x_29, 1); @@ -59400,7 +60735,7 @@ lean_dec(x_2); x_16 = lean_ctor_get(x_11, 1); lean_inc(x_16); lean_dec(x_11); -x_17 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_4, x_6, x_7, x_8, x_9, x_16); +x_17 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_4, x_6, x_7, x_8, x_9, x_16); x_18 = lean_ctor_get(x_17, 0); lean_inc(x_18); x_19 = lean_ctor_get(x_17, 1); @@ -59617,8 +60952,8 @@ static lean_object* _init_l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqOn _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___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__1; -x_2 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__2; +x_1 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__1; +x_2 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__2; x_3 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqOnFailure___closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; @@ -59634,7 +60969,7 @@ x_9 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isD x_10 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_10, 0, x_9); lean_ctor_set(x_10, 1, x_8); -x_11 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__7; +x_11 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__7; x_12 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_12, 0, x_10); lean_ctor_set(x_12, 1, x_11); @@ -64750,7 +66085,7 @@ x_11 = lean_ctor_get(x_8, 1); x_12 = 3; x_13 = lean_unbox(x_10); lean_dec(x_10); -x_14 = l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_12336_(x_13, x_12); +x_14 = l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_12510_(x_13, x_12); if (x_14 == 0) { uint8_t x_15; lean_object* x_16; @@ -64785,7 +66120,7 @@ lean_dec(x_8); x_21 = 3; x_22 = lean_unbox(x_19); lean_dec(x_19); -x_23 = l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_12336_(x_22, x_21); +x_23 = l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_12510_(x_22, x_21); if (x_23 == 0) { uint8_t x_24; lean_object* x_25; lean_object* x_26; @@ -64824,180 +66159,190 @@ return x_9; LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isExprDefEqExpensive___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: { -uint8_t x_9; lean_object* x_10; uint8_t x_11; -x_9 = l_Lean_Expr_isConst(x_1); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_2); -lean_inc(x_1); -x_10 = l_Lean_Meta_isDefEqNative(x_1, x_2, x_4, x_5, x_6, x_7, x_8); -if (x_9 == 0) +lean_object* x_9; uint8_t x_92; +x_92 = l_Lean_Expr_isConst(x_1); +if (x_92 == 0) { -uint8_t x_214; -x_214 = 0; -x_11 = x_214; -goto block_213; +lean_object* x_93; +x_93 = lean_box(0); +x_9 = x_93; +goto block_91; } else { -uint8_t x_215; -x_215 = l_Lean_Expr_isConst(x_2); -x_11 = x_215; -goto block_213; -} -block_213: -{ -uint8_t x_12; uint8_t x_13; lean_object* x_14; -if (x_11 == 0) -{ -if (lean_obj_tag(x_10) == 0) +uint8_t x_94; +x_94 = l_Lean_Expr_isConst(x_2); +if (x_94 == 0) { -lean_object* x_197; lean_object* x_198; uint8_t x_199; uint8_t x_200; -x_197 = lean_ctor_get(x_10, 0); -lean_inc(x_197); -x_198 = lean_ctor_get(x_10, 1); -lean_inc(x_198); -lean_dec(x_10); -x_199 = 0; -x_200 = lean_unbox(x_197); -lean_dec(x_197); -x_12 = x_199; -x_13 = x_200; -x_14 = x_198; -goto block_196; +lean_object* x_95; +x_95 = lean_box(0); +x_9 = x_95; +goto block_91; } else { -uint8_t x_201; +lean_object* x_96; lean_object* x_97; uint8_t x_98; +x_96 = l_Lean_Expr_constName_x21(x_1); +x_97 = l_Lean_Expr_constName_x21(x_2); +x_98 = lean_name_eq(x_96, x_97); +lean_dec(x_97); +lean_dec(x_96); +if (x_98 == 0) +{ +uint8_t x_99; lean_object* x_100; lean_object* x_101; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_201 = !lean_is_exclusive(x_10); -if (x_201 == 0) -{ -return x_10; +x_99 = 0; +x_100 = lean_box(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_8); +return x_101; } else { -lean_object* x_202; lean_object* x_203; lean_object* x_204; -x_202 = lean_ctor_get(x_10, 0); -x_203 = lean_ctor_get(x_10, 1); -lean_inc(x_203); -lean_inc(x_202); -lean_dec(x_10); -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_102; lean_object* x_103; lean_object* x_104; +x_102 = l_Lean_Expr_constLevels_x21(x_1); +x_103 = l_Lean_Expr_constLevels_x21(x_2); +x_104 = l_Lean_Meta_isListLevelDefEqAux(x_102, x_103, x_4, x_5, x_6, x_7, x_8); +return x_104; +} } } +block_91: +{ +uint8_t x_10; uint8_t x_11; lean_object* x_12; +lean_dec(x_9); +x_10 = l_Lean_Expr_isApp(x_1); +if (x_10 == 0) +{ +x_11 = x_10; +x_12 = x_8; +goto block_81; +} +else +{ +uint8_t x_82; +x_82 = l_Lean_Expr_isApp(x_2); +if (x_82 == 0) +{ +x_11 = x_82; +x_12 = x_8; +goto block_81; } else { -if (lean_obj_tag(x_10) == 0) +lean_object* x_83; +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_2); +lean_inc(x_1); +x_83 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqApp(x_1, x_2, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_83) == 0) { -lean_object* x_205; lean_object* x_206; uint8_t x_207; uint8_t x_208; -x_205 = lean_ctor_get(x_10, 0); -lean_inc(x_205); -x_206 = lean_ctor_get(x_10, 1); -lean_inc(x_206); -lean_dec(x_10); -x_207 = 1; -x_208 = lean_unbox(x_205); -lean_dec(x_205); -x_12 = x_207; -x_13 = x_208; -x_14 = x_206; -goto block_196; +lean_object* x_84; lean_object* x_85; uint8_t x_86; +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_unbox(x_84); +lean_dec(x_84); +x_11 = x_86; +x_12 = x_85; +goto block_81; } else { -uint8_t x_209; +uint8_t x_87; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_209 = !lean_is_exclusive(x_10); -if (x_209 == 0) +x_87 = !lean_is_exclusive(x_83); +if (x_87 == 0) { -return x_10; +return x_83; } else { -lean_object* x_210; lean_object* x_211; lean_object* x_212; -x_210 = lean_ctor_get(x_10, 0); -x_211 = lean_ctor_get(x_10, 1); -lean_inc(x_211); -lean_inc(x_210); -lean_dec(x_10); -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; +lean_object* x_88; lean_object* x_89; lean_object* x_90; +x_88 = lean_ctor_get(x_83, 0); +x_89 = lean_ctor_get(x_83, 1); +lean_inc(x_89); +lean_inc(x_88); +lean_dec(x_83); +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; } } } -block_196: +} +block_81: +{ +if (x_11 == 0) +{ +lean_object* x_13; +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_2); +lean_inc(x_1); +x_13 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqProjInst(x_1, x_2, x_4, x_5, x_6, x_7, x_12); +if (lean_obj_tag(x_13) == 0) { -switch (x_13) { +lean_object* x_14; uint8_t x_15; +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_unbox(x_14); +lean_dec(x_14); +switch (x_15) { case 0: { -uint8_t x_15; lean_object* x_16; lean_object* x_17; +uint8_t x_16; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_15 = 0; -x_16 = lean_box(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_14); -return x_17; -} -case 1: +x_16 = !lean_is_exclusive(x_13); +if (x_16 == 0) { -uint8_t x_18; lean_object* x_19; lean_object* x_20; -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -x_18 = 1; +lean_object* x_17; uint8_t x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_13, 0); +lean_dec(x_17); +x_18 = 0; x_19 = lean_box(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_14); -return x_20; +lean_ctor_set(x_13, 0, x_19); +return x_13; } -default: -{ -lean_object* x_21; -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_2); -lean_inc(x_1); -x_21 = l_Lean_Meta_isDefEqNat(x_1, x_2, x_4, x_5, x_6, x_7, x_14); -if (lean_obj_tag(x_21) == 0) +else { -lean_object* x_22; uint8_t x_23; -x_22 = lean_ctor_get(x_21, 0); -lean_inc(x_22); -x_23 = lean_unbox(x_22); -lean_dec(x_22); -switch (x_23) { -case 0: +lean_object* x_20; uint8_t x_21; lean_object* x_22; lean_object* x_23; +x_20 = lean_ctor_get(x_13, 1); +lean_inc(x_20); +lean_dec(x_13); +x_21 = 0; +x_22 = lean_box(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_20); +return x_23; +} +} +case 1: { uint8_t x_24; lean_dec(x_7); @@ -65006,24 +66351,24 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_24 = !lean_is_exclusive(x_21); +x_24 = !lean_is_exclusive(x_13); if (x_24 == 0) { lean_object* x_25; uint8_t x_26; lean_object* x_27; -x_25 = lean_ctor_get(x_21, 0); +x_25 = lean_ctor_get(x_13, 0); lean_dec(x_25); -x_26 = 0; +x_26 = 1; x_27 = lean_box(x_26); -lean_ctor_set(x_21, 0, x_27); -return x_21; +lean_ctor_set(x_13, 0, x_27); +return x_13; } else { lean_object* x_28; uint8_t x_29; lean_object* x_30; lean_object* x_31; -x_28 = lean_ctor_get(x_21, 1); +x_28 = lean_ctor_get(x_13, 1); lean_inc(x_28); -lean_dec(x_21); -x_29 = 0; +lean_dec(x_13); +x_29 = 1; x_30 = lean_box(x_29); x_31 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_31, 0, x_30); @@ -65031,62 +66376,62 @@ lean_ctor_set(x_31, 1, x_28); return x_31; } } -case 1: +default: { -uint8_t x_32; +lean_object* x_32; lean_object* x_33; +x_32 = lean_ctor_get(x_13, 1); +lean_inc(x_32); +lean_dec(x_13); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_2); +lean_inc(x_1); +x_33 = l_Lean_Meta_isDefEqStringLit(x_1, x_2, x_4, x_5, x_6, x_7, x_32); +if (lean_obj_tag(x_33) == 0) +{ +lean_object* x_34; uint8_t x_35; +x_34 = lean_ctor_get(x_33, 0); +lean_inc(x_34); +x_35 = lean_unbox(x_34); +lean_dec(x_34); +switch (x_35) { +case 0: +{ +uint8_t x_36; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_32 = !lean_is_exclusive(x_21); -if (x_32 == 0) +x_36 = !lean_is_exclusive(x_33); +if (x_36 == 0) { -lean_object* x_33; uint8_t x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_21, 0); -lean_dec(x_33); -x_34 = 1; -x_35 = lean_box(x_34); -lean_ctor_set(x_21, 0, x_35); -return x_21; +lean_object* x_37; uint8_t x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_33, 0); +lean_dec(x_37); +x_38 = 0; +x_39 = lean_box(x_38); +lean_ctor_set(x_33, 0, x_39); +return x_33; } else { -lean_object* x_36; uint8_t x_37; lean_object* x_38; lean_object* x_39; -x_36 = lean_ctor_get(x_21, 1); -lean_inc(x_36); -lean_dec(x_21); -x_37 = 1; -x_38 = lean_box(x_37); -x_39 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_39, 1, x_36); -return x_39; +lean_object* x_40; uint8_t x_41; lean_object* x_42; lean_object* x_43; +x_40 = lean_ctor_get(x_33, 1); +lean_inc(x_40); +lean_dec(x_33); +x_41 = 0; +x_42 = lean_box(x_41); +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; } } -default: -{ -lean_object* x_40; lean_object* x_41; -x_40 = lean_ctor_get(x_21, 1); -lean_inc(x_40); -lean_dec(x_21); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_2); -lean_inc(x_1); -x_41 = l_Lean_Meta_isDefEqOffset(x_1, x_2, x_4, x_5, x_6, x_7, x_40); -if (lean_obj_tag(x_41) == 0) -{ -lean_object* x_42; uint8_t x_43; -x_42 = lean_ctor_get(x_41, 0); -lean_inc(x_42); -x_43 = lean_unbox(x_42); -lean_dec(x_42); -switch (x_43) { -case 0: +case 1: { uint8_t x_44; lean_dec(x_7); @@ -65095,24 +66440,24 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_44 = !lean_is_exclusive(x_41); +x_44 = !lean_is_exclusive(x_33); if (x_44 == 0) { lean_object* x_45; uint8_t x_46; lean_object* x_47; -x_45 = lean_ctor_get(x_41, 0); +x_45 = lean_ctor_get(x_33, 0); lean_dec(x_45); -x_46 = 0; +x_46 = 1; x_47 = lean_box(x_46); -lean_ctor_set(x_41, 0, x_47); -return x_41; +lean_ctor_set(x_33, 0, x_47); +return x_33; } else { lean_object* x_48; uint8_t x_49; lean_object* x_50; lean_object* x_51; -x_48 = lean_ctor_get(x_41, 1); +x_48 = lean_ctor_get(x_33, 1); lean_inc(x_48); -lean_dec(x_41); -x_49 = 0; +lean_dec(x_33); +x_49 = 1; x_50 = lean_box(x_49); x_51 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_51, 0, x_50); @@ -65120,645 +66465,651 @@ lean_ctor_set(x_51, 1, x_48); return x_51; } } -case 1: +default: { -uint8_t x_52; +lean_object* x_52; lean_object* x_53; +x_52 = lean_ctor_get(x_33, 1); +lean_inc(x_52); +lean_dec(x_33); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_2); +lean_inc(x_1); +x_53 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqUnitLike(x_1, x_2, x_4, x_5, x_6, x_7, x_52); +if (lean_obj_tag(x_53) == 0) +{ +lean_object* x_54; uint8_t x_55; +x_54 = lean_ctor_get(x_53, 0); +lean_inc(x_54); +x_55 = lean_unbox(x_54); +lean_dec(x_54); +if (x_55 == 0) +{ +lean_object* x_56; lean_object* x_57; +x_56 = lean_ctor_get(x_53, 1); +lean_inc(x_56); +lean_dec(x_53); +x_57 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqOnFailure(x_1, x_2, x_4, x_5, x_6, x_7, x_56); +return x_57; +} +else +{ +uint8_t x_58; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_52 = !lean_is_exclusive(x_41); -if (x_52 == 0) +x_58 = !lean_is_exclusive(x_53); +if (x_58 == 0) { -lean_object* x_53; uint8_t x_54; lean_object* x_55; -x_53 = lean_ctor_get(x_41, 0); -lean_dec(x_53); -x_54 = 1; -x_55 = lean_box(x_54); -lean_ctor_set(x_41, 0, x_55); -return x_41; +lean_object* x_59; uint8_t x_60; lean_object* x_61; +x_59 = lean_ctor_get(x_53, 0); +lean_dec(x_59); +x_60 = 1; +x_61 = lean_box(x_60); +lean_ctor_set(x_53, 0, x_61); +return x_53; } else { -lean_object* x_56; uint8_t x_57; lean_object* x_58; lean_object* x_59; -x_56 = lean_ctor_get(x_41, 1); -lean_inc(x_56); -lean_dec(x_41); -x_57 = 1; -x_58 = lean_box(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_56); -return x_59; +lean_object* x_62; uint8_t x_63; lean_object* x_64; lean_object* x_65; +x_62 = lean_ctor_get(x_53, 1); +lean_inc(x_62); +lean_dec(x_53); +x_63 = 1; +x_64 = lean_box(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_62); +return x_65; } } -default: -{ -lean_object* x_60; lean_object* x_61; -x_60 = lean_ctor_get(x_41, 1); -lean_inc(x_60); -lean_dec(x_41); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_2); -lean_inc(x_1); -x_61 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqDelta(x_1, x_2, x_4, x_5, x_6, x_7, x_60); -if (lean_obj_tag(x_61) == 0) -{ -lean_object* x_62; uint8_t x_63; -x_62 = lean_ctor_get(x_61, 0); -lean_inc(x_62); -x_63 = lean_unbox(x_62); -lean_dec(x_62); -switch (x_63) { -case 0: +} +else { -uint8_t x_64; +uint8_t x_66; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_64 = !lean_is_exclusive(x_61); -if (x_64 == 0) +x_66 = !lean_is_exclusive(x_53); +if (x_66 == 0) { -lean_object* x_65; uint8_t x_66; lean_object* x_67; -x_65 = lean_ctor_get(x_61, 0); -lean_dec(x_65); -x_66 = 0; -x_67 = lean_box(x_66); -lean_ctor_set(x_61, 0, x_67); -return x_61; +return x_53; } else { -lean_object* x_68; uint8_t x_69; lean_object* x_70; lean_object* x_71; -x_68 = lean_ctor_get(x_61, 1); +lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_67 = lean_ctor_get(x_53, 0); +x_68 = lean_ctor_get(x_53, 1); lean_inc(x_68); -lean_dec(x_61); -x_69 = 0; -x_70 = lean_box(x_69); -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; +lean_inc(x_67); +lean_dec(x_53); +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; } } -case 1: +} +} +} +else { -uint8_t x_72; +uint8_t x_70; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_72 = !lean_is_exclusive(x_61); -if (x_72 == 0) +x_70 = !lean_is_exclusive(x_33); +if (x_70 == 0) { -lean_object* x_73; uint8_t x_74; lean_object* x_75; -x_73 = lean_ctor_get(x_61, 0); -lean_dec(x_73); -x_74 = 1; -x_75 = lean_box(x_74); -lean_ctor_set(x_61, 0, x_75); -return x_61; +return x_33; } else { -lean_object* x_76; uint8_t x_77; lean_object* x_78; lean_object* x_79; -x_76 = lean_ctor_get(x_61, 1); -lean_inc(x_76); -lean_dec(x_61); -x_77 = 1; -x_78 = lean_box(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_76); -return x_79; +lean_object* x_71; lean_object* x_72; lean_object* x_73; +x_71 = lean_ctor_get(x_33, 0); +x_72 = lean_ctor_get(x_33, 1); +lean_inc(x_72); +lean_inc(x_71); +lean_dec(x_33); +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; +} } } -default: -{ -if (x_12 == 0) -{ -lean_object* x_80; lean_object* x_81; uint8_t x_82; uint8_t x_83; lean_object* x_84; -x_80 = lean_ctor_get(x_61, 1); -lean_inc(x_80); -if (lean_is_exclusive(x_61)) { - lean_ctor_release(x_61, 0); - lean_ctor_release(x_61, 1); - x_81 = x_61; -} else { - lean_dec_ref(x_61); - x_81 = lean_box(0); } -x_82 = l_Lean_Expr_isApp(x_1); -if (x_82 == 0) -{ -x_83 = x_82; -x_84 = x_80; -goto block_153; } else { -uint8_t x_154; -x_154 = l_Lean_Expr_isApp(x_2); -if (x_154 == 0) +uint8_t x_74; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_74 = !lean_is_exclusive(x_13); +if (x_74 == 0) { -x_83 = x_154; -x_84 = x_80; -goto block_153; +return x_13; } else { -lean_object* x_155; -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_2); -lean_inc(x_1); -x_155 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqApp(x_1, x_2, x_4, x_5, x_6, x_7, x_80); -if (lean_obj_tag(x_155) == 0) -{ -lean_object* x_156; lean_object* x_157; uint8_t x_158; -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_unbox(x_156); -lean_dec(x_156); -x_83 = x_158; -x_84 = x_157; -goto block_153; +lean_object* x_75; lean_object* x_76; lean_object* x_77; +x_75 = lean_ctor_get(x_13, 0); +x_76 = lean_ctor_get(x_13, 1); +lean_inc(x_76); +lean_inc(x_75); +lean_dec(x_13); +x_77 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_77, 0, x_75); +lean_ctor_set(x_77, 1, x_76); +return x_77; +} +} } else { -uint8_t x_159; -lean_dec(x_81); +uint8_t x_78; lean_object* x_79; lean_object* x_80; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_159 = !lean_is_exclusive(x_155); -if (x_159 == 0) -{ -return x_155; +x_78 = 1; +x_79 = lean_box(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_12); +return x_80; } -else -{ -lean_object* x_160; lean_object* x_161; lean_object* x_162; -x_160 = lean_ctor_get(x_155, 0); -x_161 = lean_ctor_get(x_155, 1); -lean_inc(x_161); -lean_inc(x_160); -lean_dec(x_155); -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; } } } } -block_153: -{ -if (x_83 == 0) +LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isExprDefEqExpensive___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_85; -lean_dec(x_81); +uint8_t x_9; lean_object* x_10; lean_object* x_17; +lean_dec(x_3); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_2); lean_inc(x_1); -x_85 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqProjInst(x_1, x_2, x_4, x_5, x_6, x_7, x_84); -if (lean_obj_tag(x_85) == 0) +x_17 = l_Lean_Meta_isDefEqNative(x_1, x_2, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_17) == 0) { -lean_object* x_86; uint8_t x_87; -x_86 = lean_ctor_get(x_85, 0); -lean_inc(x_86); -x_87 = lean_unbox(x_86); -lean_dec(x_86); -switch (x_87) { +lean_object* x_18; uint8_t x_19; +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_unbox(x_18); +lean_dec(x_18); +switch (x_19) { case 0: { -uint8_t x_88; +uint8_t x_20; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_88 = !lean_is_exclusive(x_85); -if (x_88 == 0) +x_20 = !lean_is_exclusive(x_17); +if (x_20 == 0) { -lean_object* x_89; uint8_t x_90; lean_object* x_91; -x_89 = lean_ctor_get(x_85, 0); -lean_dec(x_89); -x_90 = 0; -x_91 = lean_box(x_90); -lean_ctor_set(x_85, 0, x_91); -return x_85; +lean_object* x_21; uint8_t x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_17, 0); +lean_dec(x_21); +x_22 = 0; +x_23 = lean_box(x_22); +lean_ctor_set(x_17, 0, x_23); +return x_17; } else { -lean_object* x_92; uint8_t x_93; lean_object* x_94; lean_object* x_95; -x_92 = lean_ctor_get(x_85, 1); -lean_inc(x_92); -lean_dec(x_85); -x_93 = 0; -x_94 = lean_box(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; +lean_object* x_24; uint8_t x_25; lean_object* x_26; lean_object* x_27; +x_24 = lean_ctor_get(x_17, 1); +lean_inc(x_24); +lean_dec(x_17); +x_25 = 0; +x_26 = lean_box(x_25); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_24); +return x_27; } } case 1: { -uint8_t x_96; +uint8_t x_28; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_96 = !lean_is_exclusive(x_85); -if (x_96 == 0) +x_28 = !lean_is_exclusive(x_17); +if (x_28 == 0) { -lean_object* x_97; uint8_t x_98; lean_object* x_99; -x_97 = lean_ctor_get(x_85, 0); -lean_dec(x_97); -x_98 = 1; -x_99 = lean_box(x_98); -lean_ctor_set(x_85, 0, x_99); -return x_85; +lean_object* x_29; uint8_t x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_17, 0); +lean_dec(x_29); +x_30 = 1; +x_31 = lean_box(x_30); +lean_ctor_set(x_17, 0, x_31); +return x_17; } else { -lean_object* x_100; uint8_t x_101; lean_object* x_102; lean_object* x_103; -x_100 = lean_ctor_get(x_85, 1); -lean_inc(x_100); -lean_dec(x_85); -x_101 = 1; -x_102 = lean_box(x_101); -x_103 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_103, 0, x_102); -lean_ctor_set(x_103, 1, x_100); -return x_103; +lean_object* x_32; uint8_t x_33; lean_object* x_34; lean_object* x_35; +x_32 = lean_ctor_get(x_17, 1); +lean_inc(x_32); +lean_dec(x_17); +x_33 = 1; +x_34 = lean_box(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_32); +return x_35; } } default: { -lean_object* x_104; lean_object* x_105; -x_104 = lean_ctor_get(x_85, 1); -lean_inc(x_104); -lean_dec(x_85); +lean_object* x_36; lean_object* x_37; +x_36 = lean_ctor_get(x_17, 1); +lean_inc(x_36); +lean_dec(x_17); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_2); lean_inc(x_1); -x_105 = l_Lean_Meta_isDefEqStringLit(x_1, x_2, x_4, x_5, x_6, x_7, x_104); -if (lean_obj_tag(x_105) == 0) +x_37 = l_Lean_Meta_isDefEqNat(x_1, x_2, x_4, x_5, x_6, x_7, x_36); +if (lean_obj_tag(x_37) == 0) { -lean_object* x_106; uint8_t x_107; -x_106 = lean_ctor_get(x_105, 0); -lean_inc(x_106); -x_107 = lean_unbox(x_106); -lean_dec(x_106); -switch (x_107) { +lean_object* x_38; uint8_t x_39; +x_38 = lean_ctor_get(x_37, 0); +lean_inc(x_38); +x_39 = lean_unbox(x_38); +lean_dec(x_38); +switch (x_39) { case 0: { -uint8_t x_108; +uint8_t x_40; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_108 = !lean_is_exclusive(x_105); -if (x_108 == 0) +x_40 = !lean_is_exclusive(x_37); +if (x_40 == 0) { -lean_object* x_109; uint8_t x_110; lean_object* x_111; -x_109 = lean_ctor_get(x_105, 0); -lean_dec(x_109); -x_110 = 0; -x_111 = lean_box(x_110); -lean_ctor_set(x_105, 0, x_111); -return x_105; +lean_object* x_41; uint8_t x_42; lean_object* x_43; +x_41 = lean_ctor_get(x_37, 0); +lean_dec(x_41); +x_42 = 0; +x_43 = lean_box(x_42); +lean_ctor_set(x_37, 0, x_43); +return x_37; } else { -lean_object* x_112; uint8_t x_113; lean_object* x_114; lean_object* x_115; -x_112 = lean_ctor_get(x_105, 1); -lean_inc(x_112); -lean_dec(x_105); -x_113 = 0; -x_114 = lean_box(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); -return x_115; +lean_object* x_44; uint8_t x_45; lean_object* x_46; lean_object* x_47; +x_44 = lean_ctor_get(x_37, 1); +lean_inc(x_44); +lean_dec(x_37); +x_45 = 0; +x_46 = lean_box(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_44); +return x_47; } } case 1: { -uint8_t x_116; +uint8_t x_48; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_116 = !lean_is_exclusive(x_105); -if (x_116 == 0) +x_48 = !lean_is_exclusive(x_37); +if (x_48 == 0) { -lean_object* x_117; uint8_t x_118; lean_object* x_119; -x_117 = lean_ctor_get(x_105, 0); -lean_dec(x_117); -x_118 = 1; -x_119 = lean_box(x_118); -lean_ctor_set(x_105, 0, x_119); -return x_105; +lean_object* x_49; uint8_t x_50; lean_object* x_51; +x_49 = lean_ctor_get(x_37, 0); +lean_dec(x_49); +x_50 = 1; +x_51 = lean_box(x_50); +lean_ctor_set(x_37, 0, x_51); +return x_37; } else { -lean_object* x_120; uint8_t x_121; lean_object* x_122; lean_object* x_123; -x_120 = lean_ctor_get(x_105, 1); -lean_inc(x_120); -lean_dec(x_105); -x_121 = 1; -x_122 = lean_box(x_121); -x_123 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_123, 0, x_122); -lean_ctor_set(x_123, 1, x_120); -return x_123; +lean_object* x_52; uint8_t x_53; lean_object* x_54; lean_object* x_55; +x_52 = lean_ctor_get(x_37, 1); +lean_inc(x_52); +lean_dec(x_37); +x_53 = 1; +x_54 = lean_box(x_53); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_54); +lean_ctor_set(x_55, 1, x_52); +return x_55; } } default: { -lean_object* x_124; lean_object* x_125; -x_124 = lean_ctor_get(x_105, 1); -lean_inc(x_124); -lean_dec(x_105); +lean_object* x_56; lean_object* x_57; +x_56 = lean_ctor_get(x_37, 1); +lean_inc(x_56); +lean_dec(x_37); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_2); lean_inc(x_1); -x_125 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqUnitLike(x_1, x_2, x_4, x_5, x_6, x_7, x_124); -if (lean_obj_tag(x_125) == 0) -{ -lean_object* x_126; uint8_t x_127; -x_126 = lean_ctor_get(x_125, 0); -lean_inc(x_126); -x_127 = lean_unbox(x_126); -lean_dec(x_126); -if (x_127 == 0) +x_57 = l_Lean_Meta_isDefEqOffset(x_1, x_2, x_4, x_5, x_6, x_7, x_56); +if (lean_obj_tag(x_57) == 0) { -lean_object* x_128; lean_object* x_129; -x_128 = lean_ctor_get(x_125, 1); -lean_inc(x_128); -lean_dec(x_125); -x_129 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqOnFailure(x_1, x_2, x_4, x_5, x_6, x_7, x_128); -return x_129; -} -else +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); +switch (x_59) { +case 0: { -uint8_t x_130; +uint8_t x_60; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_130 = !lean_is_exclusive(x_125); -if (x_130 == 0) +x_60 = !lean_is_exclusive(x_57); +if (x_60 == 0) { -lean_object* x_131; uint8_t x_132; lean_object* x_133; -x_131 = lean_ctor_get(x_125, 0); -lean_dec(x_131); -x_132 = 1; -x_133 = lean_box(x_132); -lean_ctor_set(x_125, 0, x_133); -return x_125; +lean_object* x_61; uint8_t x_62; lean_object* x_63; +x_61 = lean_ctor_get(x_57, 0); +lean_dec(x_61); +x_62 = 0; +x_63 = lean_box(x_62); +lean_ctor_set(x_57, 0, x_63); +return x_57; } else { -lean_object* x_134; uint8_t x_135; lean_object* x_136; lean_object* x_137; -x_134 = lean_ctor_get(x_125, 1); -lean_inc(x_134); -lean_dec(x_125); -x_135 = 1; -x_136 = lean_box(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_134); -return x_137; -} +lean_object* x_64; uint8_t x_65; lean_object* x_66; lean_object* x_67; +x_64 = lean_ctor_get(x_57, 1); +lean_inc(x_64); +lean_dec(x_57); +x_65 = 0; +x_66 = lean_box(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_64); +return x_67; } } -else +case 1: { -uint8_t x_138; +uint8_t x_68; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_138 = !lean_is_exclusive(x_125); -if (x_138 == 0) +x_68 = !lean_is_exclusive(x_57); +if (x_68 == 0) { -return x_125; +lean_object* x_69; uint8_t x_70; lean_object* x_71; +x_69 = lean_ctor_get(x_57, 0); +lean_dec(x_69); +x_70 = 1; +x_71 = lean_box(x_70); +lean_ctor_set(x_57, 0, x_71); +return x_57; } else { -lean_object* x_139; lean_object* x_140; lean_object* x_141; -x_139 = lean_ctor_get(x_125, 0); -x_140 = lean_ctor_get(x_125, 1); -lean_inc(x_140); -lean_inc(x_139); -lean_dec(x_125); -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_object* x_72; uint8_t x_73; lean_object* x_74; lean_object* x_75; +x_72 = lean_ctor_get(x_57, 1); +lean_inc(x_72); +lean_dec(x_57); +x_73 = 1; +x_74 = lean_box(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_72); +return x_75; } } -else +default: { -uint8_t x_142; +lean_object* x_76; lean_object* x_77; +x_76 = lean_ctor_get(x_57, 1); +lean_inc(x_76); +lean_dec(x_57); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_2); +lean_inc(x_1); +x_77 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqDelta(x_1, x_2, x_4, x_5, x_6, x_7, x_76); +if (lean_obj_tag(x_77) == 0) +{ +lean_object* x_78; uint8_t x_79; +x_78 = lean_ctor_get(x_77, 0); +lean_inc(x_78); +x_79 = lean_unbox(x_78); +lean_dec(x_78); +switch (x_79) { +case 0: +{ +uint8_t x_80; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_142 = !lean_is_exclusive(x_105); -if (x_142 == 0) +x_80 = !lean_is_exclusive(x_77); +if (x_80 == 0) { -return x_105; +lean_object* x_81; uint8_t x_82; lean_object* x_83; +x_81 = lean_ctor_get(x_77, 0); +lean_dec(x_81); +x_82 = 0; +x_83 = lean_box(x_82); +lean_ctor_set(x_77, 0, x_83); +return x_77; } else { -lean_object* x_143; lean_object* x_144; lean_object* x_145; -x_143 = lean_ctor_get(x_105, 0); -x_144 = lean_ctor_get(x_105, 1); -lean_inc(x_144); -lean_inc(x_143); -lean_dec(x_105); -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_object* x_84; uint8_t x_85; lean_object* x_86; lean_object* x_87; +x_84 = lean_ctor_get(x_77, 1); +lean_inc(x_84); +lean_dec(x_77); +x_85 = 0; +x_86 = lean_box(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_84); +return x_87; } } -else +case 1: { -uint8_t x_146; +uint8_t x_88; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_146 = !lean_is_exclusive(x_85); -if (x_146 == 0) +x_88 = !lean_is_exclusive(x_77); +if (x_88 == 0) { -return x_85; +lean_object* x_89; uint8_t x_90; lean_object* x_91; +x_89 = lean_ctor_get(x_77, 0); +lean_dec(x_89); +x_90 = 1; +x_91 = lean_box(x_90); +lean_ctor_set(x_77, 0, x_91); +return x_77; } else { -lean_object* x_147; lean_object* x_148; lean_object* x_149; -x_147 = lean_ctor_get(x_85, 0); -x_148 = lean_ctor_get(x_85, 1); -lean_inc(x_148); -lean_inc(x_147); -lean_dec(x_85); -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_object* x_92; uint8_t x_93; lean_object* x_94; lean_object* x_95; +x_92 = lean_ctor_get(x_77, 1); +lean_inc(x_92); +lean_dec(x_77); +x_93 = 1; +x_94 = lean_box(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; } } +default: +{ +lean_object* x_96; lean_object* x_97; +x_96 = lean_ctor_get(x_77, 1); +lean_inc(x_96); +lean_dec(x_77); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_2); +lean_inc(x_1); +x_97 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct(x_1, x_2, x_4, x_5, x_6, x_7, x_96); +if (lean_obj_tag(x_97) == 0) +{ +lean_object* x_98; uint8_t x_99; +x_98 = lean_ctor_get(x_97, 0); +lean_inc(x_98); +x_99 = lean_unbox(x_98); +if (x_99 == 0) +{ +lean_object* x_100; lean_object* x_101; +lean_dec(x_98); +x_100 = lean_ctor_get(x_97, 1); +lean_inc(x_100); +lean_dec(x_97); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_1); +lean_inc(x_2); +x_101 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct(x_2, x_1, x_4, x_5, x_6, x_7, x_100); +if (lean_obj_tag(x_101) == 0) +{ +lean_object* x_102; lean_object* x_103; uint8_t 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 = lean_unbox(x_102); +lean_dec(x_102); +x_9 = x_104; +x_10 = x_103; +goto block_16; } else { -uint8_t x_150; lean_object* x_151; lean_object* x_152; +uint8_t x_105; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_150 = 1; -x_151 = lean_box(x_150); -if (lean_is_scalar(x_81)) { - x_152 = lean_alloc_ctor(0, 2, 0); -} else { - x_152 = x_81; -} -lean_ctor_set(x_152, 0, x_151); -lean_ctor_set(x_152, 1, x_84); -return x_152; -} -} +x_105 = !lean_is_exclusive(x_101); +if (x_105 == 0) +{ +return x_101; } else { -uint8_t x_163; -x_163 = !lean_is_exclusive(x_61); -if (x_163 == 0) -{ -lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; uint8_t x_168; -x_164 = lean_ctor_get(x_61, 1); -x_165 = lean_ctor_get(x_61, 0); -lean_dec(x_165); -x_166 = l_Lean_Expr_constName_x21(x_1); -x_167 = l_Lean_Expr_constName_x21(x_2); -x_168 = lean_name_eq(x_166, x_167); -lean_dec(x_167); -lean_dec(x_166); -if (x_168 == 0) -{ -uint8_t x_169; lean_object* x_170; -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -x_169 = 0; -x_170 = lean_box(x_169); -lean_ctor_set(x_61, 0, x_170); -return x_61; +lean_object* x_106; lean_object* x_107; lean_object* x_108; +x_106 = lean_ctor_get(x_101, 0); +x_107 = lean_ctor_get(x_101, 1); +lean_inc(x_107); +lean_inc(x_106); +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; +} +} } else { -lean_object* x_171; lean_object* x_172; lean_object* x_173; -lean_free_object(x_61); -x_171 = l_Lean_Expr_constLevels_x21(x_1); -x_172 = l_Lean_Expr_constLevels_x21(x_2); -x_173 = l_Lean_Meta_isListLevelDefEqAux(x_171, x_172, x_4, x_5, x_6, x_7, x_164); -return x_173; +lean_object* x_109; uint8_t x_110; +x_109 = lean_ctor_get(x_97, 1); +lean_inc(x_109); +lean_dec(x_97); +x_110 = lean_unbox(x_98); +lean_dec(x_98); +x_9 = x_110; +x_10 = x_109; +goto block_16; } } else { -lean_object* x_174; lean_object* x_175; lean_object* x_176; uint8_t x_177; -x_174 = lean_ctor_get(x_61, 1); -lean_inc(x_174); -lean_dec(x_61); -x_175 = l_Lean_Expr_constName_x21(x_1); -x_176 = l_Lean_Expr_constName_x21(x_2); -x_177 = lean_name_eq(x_175, x_176); -lean_dec(x_176); -lean_dec(x_175); -if (x_177 == 0) -{ -uint8_t x_178; lean_object* x_179; lean_object* x_180; +uint8_t x_111; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_178 = 0; -x_179 = lean_box(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_174); -return x_180; +x_111 = !lean_is_exclusive(x_97); +if (x_111 == 0) +{ +return x_97; } else { -lean_object* x_181; lean_object* x_182; lean_object* x_183; -x_181 = l_Lean_Expr_constLevels_x21(x_1); -x_182 = l_Lean_Expr_constLevels_x21(x_2); -x_183 = l_Lean_Meta_isListLevelDefEqAux(x_181, x_182, x_4, x_5, x_6, x_7, x_174); -return x_183; -} +lean_object* x_112; lean_object* x_113; lean_object* x_114; +x_112 = lean_ctor_get(x_97, 0); +x_113 = lean_ctor_get(x_97, 1); +lean_inc(x_113); +lean_inc(x_112); +lean_dec(x_97); +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; } } } @@ -65766,30 +67117,30 @@ return x_183; } else { -uint8_t x_184; +uint8_t x_115; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_184 = !lean_is_exclusive(x_61); -if (x_184 == 0) +x_115 = !lean_is_exclusive(x_77); +if (x_115 == 0) { -return x_61; +return x_77; } else { -lean_object* x_185; lean_object* x_186; lean_object* x_187; -x_185 = lean_ctor_get(x_61, 0); -x_186 = lean_ctor_get(x_61, 1); -lean_inc(x_186); -lean_inc(x_185); -lean_dec(x_61); -x_187 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_187, 0, x_185); -lean_ctor_set(x_187, 1, x_186); -return x_187; +lean_object* x_116; lean_object* x_117; lean_object* x_118; +x_116 = lean_ctor_get(x_77, 0); +x_117 = lean_ctor_get(x_77, 1); +lean_inc(x_117); +lean_inc(x_116); +lean_dec(x_77); +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; } } } @@ -65797,30 +67148,30 @@ return x_187; } else { -uint8_t x_188; +uint8_t x_119; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_188 = !lean_is_exclusive(x_41); -if (x_188 == 0) +x_119 = !lean_is_exclusive(x_57); +if (x_119 == 0) { -return x_41; +return x_57; } else { -lean_object* x_189; lean_object* x_190; lean_object* x_191; -x_189 = lean_ctor_get(x_41, 0); -x_190 = lean_ctor_get(x_41, 1); -lean_inc(x_190); -lean_inc(x_189); -lean_dec(x_41); -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_120; lean_object* x_121; lean_object* x_122; +x_120 = lean_ctor_get(x_57, 0); +x_121 = lean_ctor_get(x_57, 1); +lean_inc(x_121); +lean_inc(x_120); +lean_dec(x_57); +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; } } } @@ -65828,128 +67179,87 @@ return x_191; } else { -uint8_t x_192; +uint8_t x_123; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_192 = !lean_is_exclusive(x_21); -if (x_192 == 0) +x_123 = !lean_is_exclusive(x_37); +if (x_123 == 0) { -return x_21; +return x_37; } else { -lean_object* x_193; lean_object* x_194; lean_object* x_195; -x_193 = lean_ctor_get(x_21, 0); -x_194 = lean_ctor_get(x_21, 1); -lean_inc(x_194); -lean_inc(x_193); -lean_dec(x_21); -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_124; lean_object* x_125; lean_object* x_126; +x_124 = lean_ctor_get(x_37, 0); +x_125 = lean_ctor_get(x_37, 1); +lean_inc(x_125); +lean_inc(x_124); +lean_dec(x_37); +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; } } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isExprDefEqExpensive___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_9; -lean_dec(x_3); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_2); -lean_inc(x_1); -x_9 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqProj(x_1, x_2, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_9) == 0) -{ -lean_object* x_10; uint8_t x_11; -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -x_11 = lean_unbox(x_10); -lean_dec(x_10); -if (x_11 == 0) -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_12 = lean_ctor_get(x_9, 1); -lean_inc(x_12); -lean_dec(x_9); -x_13 = lean_box(0); -x_14 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isExprDefEqExpensive___lambda__1(x_1, x_2, x_13, x_4, x_5, x_6, x_7, x_12); -return x_14; } else { -uint8_t x_15; +uint8_t x_127; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_15 = !lean_is_exclusive(x_9); -if (x_15 == 0) +x_127 = !lean_is_exclusive(x_17); +if (x_127 == 0) { -lean_object* x_16; uint8_t x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_9, 0); -lean_dec(x_16); -x_17 = 1; -x_18 = lean_box(x_17); -lean_ctor_set(x_9, 0, x_18); -return x_9; +return x_17; } else { -lean_object* x_19; uint8_t x_20; lean_object* x_21; lean_object* x_22; -x_19 = lean_ctor_get(x_9, 1); -lean_inc(x_19); -lean_dec(x_9); -x_20 = 1; -x_21 = lean_box(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_19); -return x_22; +lean_object* x_128; lean_object* x_129; lean_object* x_130; +x_128 = lean_ctor_get(x_17, 0); +x_129 = lean_ctor_get(x_17, 1); +lean_inc(x_129); +lean_inc(x_128); +lean_dec(x_17); +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; } } +block_16: +{ +if (x_9 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_box(0); +x_12 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isExprDefEqExpensive___lambda__1(x_1, x_2, x_11, x_4, x_5, x_6, x_7, x_10); +return x_12; } else { -uint8_t x_23; +uint8_t x_13; lean_object* x_14; lean_object* x_15; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_23 = !lean_is_exclusive(x_9); -if (x_23 == 0) -{ -return x_9; -} -else -{ -lean_object* x_24; lean_object* x_25; lean_object* x_26; -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_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; +x_13 = 1; +x_14 = lean_box(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_10); +return x_15; } } } @@ -65957,294 +67267,270 @@ return x_26; LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isExprDefEqExpensive(lean_object* x_1, 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_9; lean_object* x_16; +lean_object* x_8; 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_16 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEta(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_16) == 0) +x_8 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEta(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_8) == 0) { -lean_object* x_17; uint8_t x_18; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -x_18 = lean_unbox(x_17); -lean_dec(x_17); -switch (x_18) { +lean_object* x_9; uint8_t x_10; +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_unbox(x_9); +lean_dec(x_9); +switch (x_10) { case 0: { -uint8_t x_19; +uint8_t x_11; 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_19 = !lean_is_exclusive(x_16); -if (x_19 == 0) +x_11 = !lean_is_exclusive(x_8); +if (x_11 == 0) { -lean_object* x_20; uint8_t x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_16, 0); -lean_dec(x_20); -x_21 = 0; -x_22 = lean_box(x_21); -lean_ctor_set(x_16, 0, x_22); -return x_16; +lean_object* x_12; uint8_t x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_8, 0); +lean_dec(x_12); +x_13 = 0; +x_14 = lean_box(x_13); +lean_ctor_set(x_8, 0, x_14); +return x_8; } else { -lean_object* x_23; uint8_t x_24; lean_object* x_25; lean_object* x_26; -x_23 = lean_ctor_get(x_16, 1); -lean_inc(x_23); -lean_dec(x_16); -x_24 = 0; -x_25 = lean_box(x_24); -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_23); -return x_26; +lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; +x_15 = lean_ctor_get(x_8, 1); +lean_inc(x_15); +lean_dec(x_8); +x_16 = 0; +x_17 = lean_box(x_16); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_15); +return x_18; } } case 1: { -uint8_t x_27; +uint8_t x_19; 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_27 = !lean_is_exclusive(x_16); -if (x_27 == 0) +x_19 = !lean_is_exclusive(x_8); +if (x_19 == 0) { -lean_object* x_28; uint8_t x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_16, 0); -lean_dec(x_28); -x_29 = 1; -x_30 = lean_box(x_29); -lean_ctor_set(x_16, 0, x_30); -return x_16; +lean_object* x_20; uint8_t x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_8, 0); +lean_dec(x_20); +x_21 = 1; +x_22 = lean_box(x_21); +lean_ctor_set(x_8, 0, x_22); +return x_8; } else { -lean_object* x_31; uint8_t x_32; lean_object* x_33; lean_object* x_34; -x_31 = lean_ctor_get(x_16, 1); -lean_inc(x_31); -lean_dec(x_16); -x_32 = 1; -x_33 = lean_box(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_31); -return x_34; +lean_object* x_23; uint8_t x_24; lean_object* x_25; lean_object* x_26; +x_23 = lean_ctor_get(x_8, 1); +lean_inc(x_23); +lean_dec(x_8); +x_24 = 1; +x_25 = lean_box(x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_23); +return x_26; } } default: { -lean_object* x_35; lean_object* x_36; -x_35 = lean_ctor_get(x_16, 1); -lean_inc(x_35); -lean_dec(x_16); +lean_object* x_27; lean_object* x_28; +x_27 = lean_ctor_get(x_8, 1); +lean_inc(x_27); +lean_dec(x_8); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_1); lean_inc(x_2); -x_36 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEta(x_2, x_1, x_3, x_4, x_5, x_6, x_35); -if (lean_obj_tag(x_36) == 0) +x_28 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEta(x_2, x_1, x_3, x_4, x_5, x_6, x_27); +if (lean_obj_tag(x_28) == 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); -switch (x_38) { +lean_object* x_29; uint8_t x_30; +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +x_30 = lean_unbox(x_29); +lean_dec(x_29); +switch (x_30) { case 0: { -uint8_t x_39; +uint8_t x_31; 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_39 = !lean_is_exclusive(x_36); -if (x_39 == 0) +x_31 = !lean_is_exclusive(x_28); +if (x_31 == 0) { -lean_object* x_40; uint8_t x_41; lean_object* x_42; -x_40 = lean_ctor_get(x_36, 0); -lean_dec(x_40); -x_41 = 0; -x_42 = lean_box(x_41); -lean_ctor_set(x_36, 0, x_42); -return x_36; +lean_object* x_32; uint8_t x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_28, 0); +lean_dec(x_32); +x_33 = 0; +x_34 = lean_box(x_33); +lean_ctor_set(x_28, 0, x_34); +return x_28; } else { -lean_object* x_43; uint8_t x_44; lean_object* x_45; lean_object* x_46; -x_43 = lean_ctor_get(x_36, 1); -lean_inc(x_43); -lean_dec(x_36); -x_44 = 0; -x_45 = lean_box(x_44); -x_46 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_46, 0, x_45); -lean_ctor_set(x_46, 1, x_43); -return x_46; +lean_object* x_35; uint8_t x_36; lean_object* x_37; lean_object* x_38; +x_35 = lean_ctor_get(x_28, 1); +lean_inc(x_35); +lean_dec(x_28); +x_36 = 0; +x_37 = lean_box(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_35); +return x_38; } } case 1: { -uint8_t x_47; +uint8_t x_39; 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_47 = !lean_is_exclusive(x_36); -if (x_47 == 0) +x_39 = !lean_is_exclusive(x_28); +if (x_39 == 0) { -lean_object* x_48; uint8_t x_49; lean_object* x_50; -x_48 = lean_ctor_get(x_36, 0); -lean_dec(x_48); -x_49 = 1; -x_50 = lean_box(x_49); -lean_ctor_set(x_36, 0, x_50); -return x_36; +lean_object* x_40; uint8_t x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_28, 0); +lean_dec(x_40); +x_41 = 1; +x_42 = lean_box(x_41); +lean_ctor_set(x_28, 0, x_42); +return x_28; } else { -lean_object* x_51; uint8_t x_52; lean_object* x_53; lean_object* x_54; -x_51 = lean_ctor_get(x_36, 1); -lean_inc(x_51); -lean_dec(x_36); -x_52 = 1; -x_53 = lean_box(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; +lean_object* x_43; uint8_t x_44; lean_object* x_45; lean_object* x_46; +x_43 = lean_ctor_get(x_28, 1); +lean_inc(x_43); +lean_dec(x_28); +x_44 = 1; +x_45 = lean_box(x_44); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_43); +return x_46; } } default: { -lean_object* x_55; lean_object* x_56; -x_55 = lean_ctor_get(x_36, 1); -lean_inc(x_55); -lean_dec(x_36); +lean_object* x_47; lean_object* x_48; +x_47 = lean_ctor_get(x_28, 1); +lean_inc(x_47); +lean_dec(x_28); 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_56 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct(x_1, x_2, x_3, x_4, x_5, x_6, x_55); -if (lean_obj_tag(x_56) == 0) -{ -lean_object* x_57; uint8_t x_58; -x_57 = lean_ctor_get(x_56, 0); -lean_inc(x_57); -x_58 = lean_unbox(x_57); -if (x_58 == 0) +x_48 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqProj(x_1, x_2, x_3, x_4, x_5, x_6, x_47); +if (lean_obj_tag(x_48) == 0) { -lean_object* x_59; lean_object* x_60; -lean_dec(x_57); -x_59 = lean_ctor_get(x_56, 1); -lean_inc(x_59); -lean_dec(x_56); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -lean_inc(x_1); -lean_inc(x_2); -x_60 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct(x_2, x_1, x_3, x_4, x_5, x_6, x_59); -if (lean_obj_tag(x_60) == 0) +lean_object* x_49; uint8_t x_50; +x_49 = lean_ctor_get(x_48, 0); +lean_inc(x_49); +x_50 = lean_unbox(x_49); +lean_dec(x_49); +if (x_50 == 0) { -lean_object* x_61; lean_object* x_62; uint8_t x_63; -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_63 = lean_unbox(x_61); -lean_dec(x_61); -x_8 = x_63; -x_9 = x_62; -goto block_15; +lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_51 = lean_ctor_get(x_48, 1); +lean_inc(x_51); +lean_dec(x_48); +x_52 = lean_box(0); +x_53 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isExprDefEqExpensive___lambda__2(x_1, x_2, x_52, x_3, x_4, x_5, x_6, x_51); +return x_53; } else { -uint8_t x_64; +uint8_t x_54; 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_64 = !lean_is_exclusive(x_60); -if (x_64 == 0) +x_54 = !lean_is_exclusive(x_48); +if (x_54 == 0) { -return x_60; +lean_object* x_55; uint8_t x_56; lean_object* x_57; +x_55 = lean_ctor_get(x_48, 0); +lean_dec(x_55); +x_56 = 1; +x_57 = lean_box(x_56); +lean_ctor_set(x_48, 0, x_57); +return x_48; } else { -lean_object* x_65; lean_object* x_66; lean_object* x_67; -x_65 = lean_ctor_get(x_60, 0); -x_66 = lean_ctor_get(x_60, 1); -lean_inc(x_66); -lean_inc(x_65); -lean_dec(x_60); -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; -} -} +lean_object* x_58; uint8_t x_59; lean_object* x_60; lean_object* x_61; +x_58 = lean_ctor_get(x_48, 1); +lean_inc(x_58); +lean_dec(x_48); +x_59 = 1; +x_60 = lean_box(x_59); +x_61 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_61, 0, x_60); +lean_ctor_set(x_61, 1, x_58); +return x_61; } -else -{ -lean_object* x_68; uint8_t x_69; -x_68 = lean_ctor_get(x_56, 1); -lean_inc(x_68); -lean_dec(x_56); -x_69 = lean_unbox(x_57); -lean_dec(x_57); -x_8 = x_69; -x_9 = x_68; -goto block_15; } } else { -uint8_t x_70; +uint8_t x_62; 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_70 = !lean_is_exclusive(x_56); -if (x_70 == 0) +x_62 = !lean_is_exclusive(x_48); +if (x_62 == 0) { -return x_56; +return x_48; } else { -lean_object* x_71; lean_object* x_72; lean_object* x_73; -x_71 = lean_ctor_get(x_56, 0); -x_72 = lean_ctor_get(x_56, 1); -lean_inc(x_72); -lean_inc(x_71); -lean_dec(x_56); -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_object* x_63; lean_object* x_64; lean_object* x_65; +x_63 = lean_ctor_get(x_48, 0); +x_64 = lean_ctor_get(x_48, 1); +lean_inc(x_64); +lean_inc(x_63); +lean_dec(x_48); +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; } } } @@ -66252,30 +67538,30 @@ return x_73; } else { -uint8_t x_74; +uint8_t x_66; 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_74 = !lean_is_exclusive(x_36); -if (x_74 == 0) +x_66 = !lean_is_exclusive(x_28); +if (x_66 == 0) { -return x_36; +return x_28; } else { -lean_object* x_75; lean_object* x_76; lean_object* x_77; -x_75 = lean_ctor_get(x_36, 0); -x_76 = lean_ctor_get(x_36, 1); -lean_inc(x_76); -lean_inc(x_75); -lean_dec(x_36); -x_77 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_77, 0, x_75); -lean_ctor_set(x_77, 1, x_76); -return x_77; +lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_67 = lean_ctor_get(x_28, 0); +x_68 = lean_ctor_get(x_28, 1); +lean_inc(x_68); +lean_inc(x_67); +lean_dec(x_28); +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; } } } @@ -66283,56 +67569,30 @@ return x_77; } else { -uint8_t x_78; +uint8_t x_70; 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_78 = !lean_is_exclusive(x_16); -if (x_78 == 0) -{ -return x_16; -} -else -{ -lean_object* x_79; lean_object* x_80; lean_object* x_81; -x_79 = lean_ctor_get(x_16, 0); -x_80 = lean_ctor_get(x_16, 1); -lean_inc(x_80); -lean_inc(x_79); -lean_dec(x_16); -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; -} -} -block_15: -{ -if (x_8 == 0) +x_70 = !lean_is_exclusive(x_8); +if (x_70 == 0) { -lean_object* x_10; lean_object* x_11; -x_10 = lean_box(0); -x_11 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isExprDefEqExpensive___lambda__2(x_1, x_2, x_10, x_3, x_4, x_5, x_6, x_9); -return x_11; +return x_8; } else { -uint8_t x_12; lean_object* x_13; lean_object* 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_12 = 1; -x_13 = lean_box(x_12); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_13); -lean_ctor_set(x_14, 1, x_9); -return x_14; +lean_object* x_71; lean_object* x_72; lean_object* x_73; +x_71 = lean_ctor_get(x_8, 0); +x_72 = lean_ctor_get(x_8, 1); +lean_inc(x_72); +lean_inc(x_71); +lean_dec(x_8); +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; } } } @@ -67405,7 +68665,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_cacheResul 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; uint8_t x_21; x_8 = lean_ctor_get(x_1, 0); lean_inc(x_8); -x_9 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_8, x_3, x_4, x_5, x_6, x_7); +x_9 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_8, x_3, x_4, x_5, x_6, x_7); x_10 = lean_ctor_get(x_9, 0); lean_inc(x_10); x_11 = lean_ctor_get(x_9, 1); @@ -67414,7 +68674,7 @@ lean_dec(x_9); x_12 = lean_ctor_get(x_1, 1); lean_inc(x_12); lean_dec(x_1); -x_13 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_12, x_3, x_4, x_5, x_6, x_11); +x_13 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_12, x_3, x_4, x_5, x_6, x_11); x_14 = lean_ctor_get(x_13, 0); lean_inc(x_14); x_15 = lean_ctor_get(x_13, 1); @@ -67686,8 +68946,8 @@ static lean_object* _init_l_Lean_Meta_isExprDefEqAuxImpl___lambda__3___closed__2 _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___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__1; -x_2 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__2; +x_1 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__1; +x_2 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__2; x_3 = l_Lean_Meta_isExprDefEqAuxImpl___lambda__3___closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; @@ -67869,7 +69129,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_isExprDefEqAuxImpl___lambda__3(lean_object* _start: { lean_object* x_10; lean_object* x_11; -x_10 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__2; +x_10 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__2; x_11 = l_Lean_Core_checkMaxHeartbeats(x_10, x_7, x_8, x_9); if (lean_obj_tag(x_11) == 0) { @@ -68123,13 +69383,13 @@ 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; lean_object* x_74; lean_object* x_75; lean_object* x_76; uint8_t x_77; lean_dec(x_59); lean_dec(x_56); -x_65 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_1, x_5, x_6, x_7, x_8, x_60); +x_65 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_1, x_5, x_6, x_7, x_8, x_60); 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_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_2, x_5, x_6, x_7, x_8, x_67); +x_68 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_2, x_5, x_6, x_7, x_8, x_67); x_69 = lean_ctor_get(x_68, 0); lean_inc(x_69); x_70 = lean_ctor_get(x_68, 1); @@ -68824,7 +70084,7 @@ x_9 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isD x_10 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_10, 0, x_9); lean_ctor_set(x_10, 1, x_8); -x_11 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__7; +x_11 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__7; x_12 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_12, 0, x_10); lean_ctor_set(x_12, 1, x_11); @@ -68899,7 +70159,7 @@ x_42 = lean_unsigned_to_nat(1u); x_43 = lean_nat_add(x_21, x_42); lean_dec(x_21); lean_ctor_set(x_5, 3, x_43); -x_44 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___closed__2; +x_44 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___closed__1; x_45 = 1; x_46 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1(x_44, x_16, x_17, x_45, x_3, x_4, x_5, x_6, x_7); return x_46; @@ -68923,7 +70183,7 @@ lean_ctor_set(x_49, 7, x_25); lean_ctor_set(x_49, 8, x_26); lean_ctor_set(x_49, 9, x_27); lean_ctor_set(x_49, 10, x_28); -x_50 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___closed__2; +x_50 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___closed__1; x_51 = 1; x_52 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1(x_50, x_16, x_17, x_51, x_3, x_4, x_49, x_6, x_7); return x_52; @@ -68979,7 +70239,7 @@ lean_dec(x_4); return x_11; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__1() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__1() { _start: { lean_object* x_1; @@ -68987,27 +70247,27 @@ x_1 = lean_mk_string_from_bytes("Lean", 4); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__2() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__1; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__3() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__2; -x_2 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__1; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__2; +x_2 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__4() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__4() { _start: { lean_object* x_1; @@ -69015,17 +70275,17 @@ x_1 = lean_mk_string_from_bytes("initFn", 6); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__5() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__3; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__4; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__3; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__6() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__6() { _start: { lean_object* x_1; @@ -69033,37 +70293,37 @@ x_1 = lean_mk_string_from_bytes("_@", 2); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__7() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__5; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__6; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__5; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__8() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__7; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__1; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__7; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__9() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__8; -x_2 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__1; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__8; +x_2 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__10() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__10() { _start: { lean_object* x_1; @@ -69071,17 +70331,17 @@ x_1 = lean_mk_string_from_bytes("ExprDefEq", 9); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__11() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__9; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__10; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__9; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__10; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__12() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__12() { _start: { lean_object* x_1; @@ -69089,33 +70349,33 @@ x_1 = lean_mk_string_from_bytes("_hyg", 4); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__13() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__11; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__12; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__11; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__12; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__14() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__13; -x_2 = lean_unsigned_to_nat(19595u); +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__13; +x_2 = lean_unsigned_to_nat(20099u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___closed__2; +x_2 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___closed__1; x_3 = 0; -x_4 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__14; +x_4 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__14; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); if (lean_obj_tag(x_5) == 0) { @@ -69260,7 +70520,7 @@ 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_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__5; +x_59 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__5; x_60 = l_Lean_registerTraceClass(x_59, x_3, x_4, x_58); return x_60; } @@ -69724,28 +70984,28 @@ l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEt lean_mark_persistent(l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__2___closed__8); l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__2___closed__9 = _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__2___closed__9(); lean_mark_persistent(l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__2___closed__9); -l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__1 = _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__1(); -lean_mark_persistent(l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__1); -l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__2 = _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__2(); -lean_mark_persistent(l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__2); -l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__3 = _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__3(); -lean_mark_persistent(l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__3); -l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__4 = _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__4(); -lean_mark_persistent(l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__4); -l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__5 = _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__5(); -lean_mark_persistent(l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__5); -l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__6 = _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__6(); -lean_mark_persistent(l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__6); -l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__7 = _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__7(); -lean_mark_persistent(l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__7); -l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__8 = _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__8(); -lean_mark_persistent(l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__8); -l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__9 = _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__9(); -lean_mark_persistent(l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__9); -l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__10 = _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__10(); -lean_mark_persistent(l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__10); -l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__11 = _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__11(); -lean_mark_persistent(l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___closed__11); +l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__1 = _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__1(); +lean_mark_persistent(l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__1); +l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__2 = _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__2(); +lean_mark_persistent(l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__2); +l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__3 = _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__3(); +lean_mark_persistent(l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__3); +l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__4 = _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__4(); +lean_mark_persistent(l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__4); +l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__5 = _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__5(); +lean_mark_persistent(l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__5); +l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__6 = _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__6(); +lean_mark_persistent(l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__6); +l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__7 = _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__7(); +lean_mark_persistent(l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__7); +l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__8 = _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__8(); +lean_mark_persistent(l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__8); +l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__9 = _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__9(); +lean_mark_persistent(l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__9); +l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__10 = _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__10(); +lean_mark_persistent(l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__10); +l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__11 = _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__11(); +lean_mark_persistent(l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__11); l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__2___closed__1 = _init_l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__2___closed__1(); lean_mark_persistent(l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__2___closed__1); l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__2___closed__2 = _init_l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__2___closed__2(); @@ -69780,32 +71040,36 @@ l_Lean_Meta_isDefEqStringLit___closed__2 = _init_l_Lean_Meta_isDefEqStringLit___ lean_mark_persistent(l_Lean_Meta_isDefEqStringLit___closed__2); l_Lean_Meta_isDefEqStringLit___closed__3 = _init_l_Lean_Meta_isDefEqStringLit___closed__3(); lean_mark_persistent(l_Lean_Meta_isDefEqStringLit___closed__3); +l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___lambda__2___closed__1 = _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___lambda__2___closed__1(); +lean_mark_persistent(l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___lambda__2___closed__1); l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___closed__1 = _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___closed__1(); lean_mark_persistent(l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___closed__1); l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___closed__2 = _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___closed__2(); lean_mark_persistent(l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___closed__2); l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___closed__3 = _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___closed__3(); lean_mark_persistent(l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___closed__3); -l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___closed__4 = _init_l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___closed__4(); -lean_mark_persistent(l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___closed__4); l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___closed__1 = _init_l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___closed__1(); lean_mark_persistent(l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___closed__1); l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___closed__2 = _init_l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___closed__2(); lean_mark_persistent(l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___closed__2); l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___closed__3 = _init_l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___closed__3(); lean_mark_persistent(l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___closed__3); -l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___closed__1 = _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___closed__1(); -lean_mark_persistent(l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___closed__1); -l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___closed__2 = _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___closed__2(); -lean_mark_persistent(l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___closed__2); -l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___closed__3 = _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___closed__3(); -lean_mark_persistent(l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___closed__3); -l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___closed__4 = _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___closed__4(); -lean_mark_persistent(l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___closed__4); -l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___closed__5 = _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___closed__5(); -lean_mark_persistent(l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___closed__5); -l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___closed__6 = _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___closed__6(); -lean_mark_persistent(l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___closed__6); +l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__1 = _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__1); +l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__2 = _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__2); +l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__3 = _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__3(); +lean_mark_persistent(l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__3); +l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__4 = _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__4(); +lean_mark_persistent(l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__4); +l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__5 = _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__5(); +lean_mark_persistent(l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__5); +l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__6 = _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__6(); +lean_mark_persistent(l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__6); +l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__7 = _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__7(); +lean_mark_persistent(l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__7); +l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__4___closed__1 = _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__4___closed__1); l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___lambda__1___closed__1 = _init_l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___lambda__1___closed__1(); lean_mark_persistent(l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___lambda__1___closed__1); l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___lambda__1___closed__2 = _init_l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___lambda__1___closed__2(); @@ -69840,20 +71104,20 @@ l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_hasLetDecl lean_mark_persistent(l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_hasLetDeclsInBetween___closed__1); l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_hasLetDeclsInBetween___closed__2 = _init_l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_hasLetDeclsInBetween___closed__2(); lean_mark_persistent(l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_hasLetDeclsInBetween___closed__2); -l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6051____closed__1 = _init_l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6051____closed__1(); -lean_mark_persistent(l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6051____closed__1); -l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6051____closed__2 = _init_l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6051____closed__2(); -lean_mark_persistent(l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6051____closed__2); -if (builtin) {res = l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6051_(lean_io_mk_world()); +l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6556____closed__1 = _init_l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6556____closed__1(); +lean_mark_persistent(l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6556____closed__1); +l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6556____closed__2 = _init_l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6556____closed__2(); +lean_mark_persistent(l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6556____closed__2); +if (builtin) {res = l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6556_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Meta_CheckAssignment_checkAssignmentExceptionId = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Meta_CheckAssignment_checkAssignmentExceptionId); lean_dec_ref(res); -}l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6071____closed__1 = _init_l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6071____closed__1(); -lean_mark_persistent(l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6071____closed__1); -l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6071____closed__2 = _init_l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6071____closed__2(); -lean_mark_persistent(l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6071____closed__2); -if (builtin) {res = l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6071_(lean_io_mk_world()); +}l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6576____closed__1 = _init_l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6576____closed__1(); +lean_mark_persistent(l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6576____closed__1); +l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6576____closed__2 = _init_l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6576____closed__2(); +lean_mark_persistent(l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6576____closed__2); +if (builtin) {res = l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_6576_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Meta_CheckAssignment_outOfScopeExceptionId = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Meta_CheckAssignment_outOfScopeExceptionId); @@ -70064,35 +71328,35 @@ l_Lean_Meta_isExprDefEqAuxImpl___lambda__3___closed__19 = _init_l_Lean_Meta_isEx lean_mark_persistent(l_Lean_Meta_isExprDefEqAuxImpl___lambda__3___closed__19); l_Lean_Meta_isExprDefEqAuxImpl___lambda__3___closed__20 = _init_l_Lean_Meta_isExprDefEqAuxImpl___lambda__3___closed__20(); lean_mark_persistent(l_Lean_Meta_isExprDefEqAuxImpl___lambda__3___closed__20); -l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__1 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__1(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__1); -l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__2 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__2(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__2); -l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__3 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__3(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__3); -l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__4 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__4(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__4); -l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__5 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__5(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__5); -l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__6 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__6(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__6); -l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__7 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__7(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__7); -l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__8 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__8(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__8); -l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__9 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__9(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__9); -l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__10 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__10(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__10); -l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__11 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__11(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__11); -l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__12 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__12(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__12); -l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__13 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__13(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__13); -l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__14 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__14(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595____closed__14); -res = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19595_(lean_io_mk_world()); +l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__1 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__1(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__1); +l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__2 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__2(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__2); +l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__3 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__3(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__3); +l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__4 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__4(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__4); +l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__5 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__5(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__5); +l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__6 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__6(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__6); +l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__7 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__7(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__7); +l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__8 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__8(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__8); +l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__9 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__9(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__9); +l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__10 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__10(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__10); +l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__11 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__11(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__11); +l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__12 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__12(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__12); +l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__13 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__13(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__13); +l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__14 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__14(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099____closed__14); +res = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_20099_(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/Meta/ExprLens.c b/stage0/stdlib/Lean/Meta/ExprLens.c index e37dd1f9030d..058b38b73132 100644 --- a/stage0/stdlib/Lean/Meta/ExprLens.c +++ b/stage0/stdlib/Lean/Meta/ExprLens.c @@ -416,7 +416,7 @@ static lean_object* _init_l___private_Lean_Meta_ExprLens_0__Lean_Meta_lensCoord_ 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_Meta_ExprLens_0__Lean_Meta_lensCoord___rarg___lambda__3___closed__1; x_2 = l___private_Lean_Meta_ExprLens_0__Lean_Meta_lensCoord___rarg___lambda__3___closed__2; -x_3 = lean_unsigned_to_nat(1465u); +x_3 = lean_unsigned_to_nat(1469u); x_4 = lean_unsigned_to_nat(18u); x_5 = l___private_Lean_Meta_ExprLens_0__Lean_Meta_lensCoord___rarg___lambda__3___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -594,7 +594,7 @@ static lean_object* _init_l___private_Lean_Meta_ExprLens_0__Lean_Meta_lensCoord_ 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_Meta_ExprLens_0__Lean_Meta_lensCoord___rarg___lambda__3___closed__1; x_2 = l___private_Lean_Meta_ExprLens_0__Lean_Meta_lensCoord___rarg___lambda__8___closed__1; -x_3 = lean_unsigned_to_nat(1569u); +x_3 = lean_unsigned_to_nat(1573u); x_4 = lean_unsigned_to_nat(22u); x_5 = l___private_Lean_Meta_ExprLens_0__Lean_Meta_lensCoord___rarg___lambda__8___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -795,7 +795,7 @@ static lean_object* _init_l___private_Lean_Meta_ExprLens_0__Lean_Meta_lensCoord_ 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_Meta_ExprLens_0__Lean_Meta_lensCoord___rarg___lambda__3___closed__1; x_2 = l___private_Lean_Meta_ExprLens_0__Lean_Meta_lensCoord___rarg___lambda__10___closed__1; -x_3 = lean_unsigned_to_nat(1560u); +x_3 = lean_unsigned_to_nat(1564u); x_4 = lean_unsigned_to_nat(20u); x_5 = l___private_Lean_Meta_ExprLens_0__Lean_Meta_lensCoord___rarg___lambda__10___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -919,7 +919,7 @@ static lean_object* _init_l___private_Lean_Meta_ExprLens_0__Lean_Meta_lensCoord_ 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_Meta_ExprLens_0__Lean_Meta_lensCoord___rarg___lambda__3___closed__1; x_2 = l___private_Lean_Meta_ExprLens_0__Lean_Meta_lensCoord___rarg___lambda__11___closed__1; -x_3 = lean_unsigned_to_nat(1540u); +x_3 = lean_unsigned_to_nat(1544u); x_4 = lean_unsigned_to_nat(24u); x_5 = l___private_Lean_Meta_ExprLens_0__Lean_Meta_lensCoord___rarg___lambda__11___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); diff --git a/stage0/stdlib/Lean/Meta/ForEachExpr.c b/stage0/stdlib/Lean/Meta/ForEachExpr.c index eab4181f1f76..4780a7a7d138 100644 --- a/stage0/stdlib/Lean/Meta/ForEachExpr.c +++ b/stage0/stdlib/Lean/Meta/ForEachExpr.c @@ -116,7 +116,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_forEachExpr___rarg(lean_object*, lean_objec LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at_Lean_Meta_forEachExpr_x27_visit___spec__11___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_st_mk_ref(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_visitLambda_visit___at_Lean_Meta_forEachExpr_x27_visit___spec__4___rarg___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_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_visitLambda_visit___at_Lean_Meta_setMVarUserNamesAt___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_Meta_visitForall___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_ST_Prim_Ref_modifyGetUnsafe___rarg___boxed(lean_object*, lean_object*, lean_object*); @@ -177,6 +176,7 @@ lean_object* l_Lean_Meta_getFVarLocalDecl(lean_object*, lean_object*, lean_objec uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Meta_setMVarUserNamesAt___spec__10___rarg___lambda__1(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_forEachExpr_x27_visit___spec__8___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*); +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_setMVarUserNamesAt(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forEachExpr_x27_visit(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_visitForall(lean_object*); @@ -5166,7 +5166,7 @@ lean_inc(x_10); x_11 = lean_ctor_get(x_9, 1); lean_inc(x_11); lean_dec(x_9); -x_12 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_1, x_3, x_4, x_5, x_6, x_11); +x_12 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_1, x_3, x_4, x_5, x_6, x_11); x_13 = lean_ctor_get(x_12, 0); lean_inc(x_13); x_14 = lean_ctor_get(x_12, 1); diff --git a/stage0/stdlib/Lean/Meta/GeneralizeTelescope.c b/stage0/stdlib/Lean/Meta/GeneralizeTelescope.c index 8051eb9f9890..2a15124e1442 100644 --- a/stage0/stdlib/Lean/Meta/GeneralizeTelescope.c +++ b/stage0/stdlib/Lean/Meta/GeneralizeTelescope.c @@ -40,7 +40,6 @@ lean_object* l_Lean_FVarId_getDecl(lean_object*, lean_object*, lean_object*, lea LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_GeneralizeTelescope_generalizeTelescopeAux___spec__2(size_t, size_t, lean_object*); static lean_object* l_Lean_Meta_generalizeTelescope___rarg___closed__1; lean_object* lean_array_to_list(lean_object*, lean_object*); -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_hasLooseBVars(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Meta_GeneralizeTelescope_generalizeTelescopeAux___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* l_Lean_Name_str___override(lean_object*, lean_object*); @@ -49,6 +48,7 @@ lean_object* l_Lean_LocalDecl_userName(lean_object*); static lean_object* l_Lean_Meta_GeneralizeTelescope_generalizeTelescopeAux___rarg___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Meta_GeneralizeTelescope_generalizeTelescopeAux___spec__1(lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_GeneralizeTelescope_generalizeTelescopeAux___rarg___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_GeneralizeTelescope_generalizeTelescopeAux___rarg(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); @@ -859,7 +859,7 @@ lean_inc(x_15); x_16 = lean_ctor_get(x_14, 1); lean_inc(x_16); lean_dec(x_14); -x_17 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_15, x_4, x_5, x_6, x_7, x_16); +x_17 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_15, x_4, x_5, x_6, x_7, x_16); x_18 = lean_ctor_get(x_17, 0); lean_inc(x_18); x_19 = lean_ctor_get(x_17, 1); diff --git a/stage0/stdlib/Lean/Meta/GeneralizeVars.c b/stage0/stdlib/Lean/Meta/GeneralizeVars.c index e7e290a758df..c50f5d1880db 100644 --- a/stage0/stdlib/Lean/Meta/GeneralizeVars.c +++ b/stage0/stdlib/Lean/Meta/GeneralizeVars.c @@ -102,7 +102,6 @@ lean_object* lean_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_getFVarSetToGeneralize___spec__70___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_array_to_list(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_Meta_getFVarSetToGeneralize___spec__50___at_Lean_Meta_getFVarSetToGeneralize___spec__58(lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_RBNode_findCore___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at_Lean_Meta_getFVarSetToGeneralize___spec__53___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_Meta_getFVarSetToGeneralize___spec__72___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -145,6 +144,7 @@ LEAN_EXPORT lean_object* l_Lean_RBNode_fold___at_Lean_Meta_getFVarsToGeneralize_ LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_Meta_getFVarSetToGeneralize___spec__62(lean_object*, uint8_t, 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___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_Meta_getFVarSetToGeneralize___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_Meta_getFVarSetToGeneralize___spec__66___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_Lean_Meta_getFVarSetToGeneralize___spec__74(lean_object*, uint8_t, 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_Array_anyMUnsafe_any___at_Lean_Meta_getFVarSetToGeneralize___spec__19___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -567,7 +567,7 @@ x_11 = lean_ctor_get(x_9, 1); lean_inc(x_11); lean_dec(x_9); x_12 = l_Lean_LocalDecl_type(x_10); -x_13 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_12, x_4, x_5, x_6, x_7, x_11); +x_13 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_12, x_4, x_5, x_6, x_7, x_11); x_14 = lean_ctor_get(x_13, 0); lean_inc(x_14); x_15 = lean_ctor_get(x_13, 1); @@ -594,7 +594,7 @@ lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean x_21 = lean_ctor_get(x_18, 0); lean_inc(x_21); lean_dec(x_18); -x_22 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_21, x_4, x_5, x_6, x_7, x_15); +x_22 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_21, x_4, x_5, x_6, x_7, x_15); x_23 = lean_ctor_get(x_22, 0); lean_inc(x_23); x_24 = lean_ctor_get(x_22, 1); @@ -822,7 +822,7 @@ lean_inc(x_18); x_19 = lean_ctor_get(x_17, 1); lean_inc(x_19); lean_dec(x_17); -x_20 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_18, x_5, x_6, x_7, x_8, x_19); +x_20 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_18, x_5, x_6, x_7, x_8, x_19); x_21 = lean_ctor_get(x_20, 0); lean_inc(x_21); x_22 = lean_ctor_get(x_20, 1); @@ -903,7 +903,7 @@ lean_inc(x_40); x_41 = lean_ctor_get(x_39, 1); lean_inc(x_41); lean_dec(x_39); -x_42 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_40, x_5, x_6, x_7, x_8, x_41); +x_42 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_40, x_5, x_6, x_7, x_8, x_41); x_43 = lean_ctor_get(x_42, 0); lean_inc(x_43); x_44 = lean_ctor_get(x_42, 1); diff --git a/stage0/stdlib/Lean/Meta/GetConst.c b/stage0/stdlib/Lean/Meta/GetConst.c index 5905c442e25a..b78e76a8309e 100644 --- a/stage0/stdlib/Lean/Meta/GetConst.c +++ b/stage0/stdlib/Lean/Meta/GetConst.c @@ -21,6 +21,7 @@ LEAN_EXPORT lean_object* l_Lean_isIrreducible___at___private_Lean_Meta_GetConst_ lean_object* lean_environment_find(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_canUnfold___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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*); +uint8_t l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_12510_(uint8_t, uint8_t); uint8_t lean_get_reducibility_status(lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getReducibilityStatus___at___private_Lean_Meta_GetConst_0__Lean_Meta_canUnfoldDefault___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -32,7 +33,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_GetConst_0__Lean_Meta_canUnfoldDe LEAN_EXPORT lean_object* l_Lean_isIrreducible___at___private_Lean_Meta_GetConst_0__Lean_Meta_canUnfoldDefault___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_throwUnknownConstant___at___private_Lean_Meta_Basic_0__Lean_Meta_getConstTemp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_GetConst_0__Lean_Meta_canUnfoldDefault___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_12336_(uint8_t, uint8_t); lean_object* l_Lean_Meta_getTheoremInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getConst_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getReducibilityStatus___at___private_Lean_Meta_GetConst_0__Lean_Meta_canUnfoldDefault___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { @@ -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_12336_(x_6, x_40); +x_41 = l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_12510_(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_12336_(x_6, x_52); +x_53 = l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_12510_(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 50888baf88f5..37e5693b4a62 100644 --- a/stage0/stdlib/Lean/Meta/IndPredBelow.c +++ b/stage0/stdlib/Lean/Meta/IndPredBelow.c @@ -105,7 +105,6 @@ static lean_object* l_Lean_Meta_IndPredBelow_mkCtorType_checkCount___lambda__2__ lean_object* l_Lean_Expr_mdata___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_IndPredBelow_mkCtorType_copyVarName___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_IndPredBelow_mkCtorType_mkMotiveBinder___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* l_Lean_Option_get___at_Std_Format_pretty_x27___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit___at_Lean_Meta_IndPredBelow_mkCtorType_checkCount___spec__2___lambda__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_IndPredBelow_mkBelowMatcher_transformFields_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* l_Lean_Expr_proj___override(lean_object*, lean_object*, lean_object*); @@ -217,6 +216,7 @@ lean_object* l_List_mapTR_loop___at_Lean_MessageData_instCoeListExprMessageData_ LEAN_EXPORT lean_object* l_Lean_Meta_IndPredBelow_mkContext_mkIndValConst(lean_object*); lean_object* l_Lean_HashMapImp_find_x3f___at_Lean_instantiateExprMVars___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Meta_IndPredBelow_mkContext_motiveName___closed__2; +lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(lean_object*, lean_object*); lean_object* l_Lean_ConstantInfo_numLevelParams(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_IndPredBelow_mkBelowMatcher_convertToBelow___spec__1(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_IndPredBelow_mkCtorType_mkMotiveBinder(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -261,7 +261,6 @@ LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_IndPredBelow_find LEAN_EXPORT lean_object* l_Lean_Meta_IndPredBelow_mkCtorType_addHeaderVars(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*); LEAN_EXPORT lean_object* l_Lean_Meta_IndPredBelow_mkCtorType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_IndPredBelow_mkContext___spec__2(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_IndPredBelow_mkCtorType_replaceTempVars___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_IndPredBelow_mkBelow___closed__1; @@ -397,6 +396,7 @@ lean_object* l_Lean_MVarId_isAssigned___at___private_Lean_Meta_SynthInstance_0__ 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*); LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit___at_Lean_Meta_IndPredBelow_mkCtorType_checkCount___spec__2___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*); lean_object* l_Lean_throwError___at_Lean_Meta_setInlineAttribute___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(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*); LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_IndPredBelow_0__Lean_Meta_IndPredBelow_belowType___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_Meta_IndPredBelow_proveBrecOn_closeGoal___closed__2; @@ -13689,7 +13689,7 @@ lean_dec(x_30); x_33 = lean_ctor_get(x_6, 2); lean_inc(x_33); x_34 = l_Lean_Meta_IndPredBelow_proveBrecOn___closed__3; -x_35 = l_Lean_Option_get___at_Std_Format_pretty_x27___spec__1(x_33, x_34); +x_35 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_33, x_34); lean_dec(x_33); lean_inc(x_7); lean_inc(x_6); @@ -13702,7 +13702,7 @@ lean_object* x_37; lean_object* x_38; x_37 = lean_ctor_get(x_36, 1); lean_inc(x_37); lean_dec(x_36); -x_38 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_11, x_4, x_5, x_6, x_7, x_37); +x_38 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_11, x_4, x_5, x_6, x_7, x_37); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); diff --git a/stage0/stdlib/Lean/Meta/Injective.c b/stage0/stdlib/Lean/Meta/Injective.c index 44f65c702534..99c55be2efdb 100644 --- a/stage0/stdlib/Lean/Meta/Injective.c +++ b/stage0/stdlib/Lean/Meta/Injective.c @@ -14,7 +14,6 @@ extern "C" { #endif lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); -uint8_t l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveEqTheoremValue___lambda__2___closed__6; lean_object* l_Lean_Meta_mkEqHEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_assumptionCore(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -41,6 +40,7 @@ lean_object* l_Lean_mkAppB(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_solveEqOfCtorEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isAppOfArity(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveTheoremTypeCore_x3f_mkArgs2___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_mkInjectiveTheorems___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); lean_object* l_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveTheoremTypeCore_x3f_mkArgs2___closed__4; @@ -49,6 +49,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiv lean_object* l_Lean_Expr_fvarId_x21(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveEqTheoremType_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_refl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_mkInjectiveTheorems___closed__5; 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___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveEqTheorem___closed__1; lean_object* l_Lean_Meta_forallBoundedTelescope___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -72,6 +73,7 @@ lean_object* l_Lean_Expr_getRevArg_x21(lean_object*, lean_object*); static lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Meta_Injective_0__Lean_Meta_mkAnd_x3f___spec__1___closed__1; uint8_t lean_expr_eqv(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveTheoremTypeCore_x3f___spec__2___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_mkInjectiveTheorems___closed__2; static lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_injTheoremFailureHeader___closed__1; static lean_object* l_Lean_Meta_elimOptParam___lambda__1___closed__2; static lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_injTheoremFailureHeader___closed__2; @@ -84,6 +86,7 @@ LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___ static lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveEqTheoremValue___lambda__2___closed__7; static lean_object* l_Lean_Meta_elimOptParam___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveTheoremTypeCore_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_mkInjectiveTheorems___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveTheoremTypeCore_x3f_mkArgs2___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_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_2126____closed__3; LEAN_EXPORT lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveTheoremType_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -100,7 +103,6 @@ static lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveTheo static lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_solveEqOfCtorEq___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveTheoremTypeCore_x3f___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_Meta_mkEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(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___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveEqTheoremValue___lambda__2___closed__2; extern lean_object* l_Lean_instInhabitedExpr; @@ -114,6 +116,7 @@ lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); lean_object* l_Lean_Meta_injection(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_elimOptParam___lambda__1___closed__1; static lean_object* l_Lean_Meta_mkInjectiveTheoremNameFor___closed__1; +uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); extern lean_object* l_Lean_Meta_simpExtension; lean_object* l_Lean_Meta_forallTelescope___at___private_Lean_Meta_InferType_0__Lean_Meta_inferForallType___spec__2___rarg(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*); @@ -121,12 +124,12 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiv LEAN_EXPORT lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_injTheoremFailureHeader(lean_object*); lean_object* l_Lean_Meta_instInhabitedMetaM___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Environment_contains(lean_object*, lean_object*); -lean_object* l_Lean_Option_register___at_Std_Format_initFn____x40_Lean_Data_Format___hyg_95____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_substEqs(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_throwInjectiveTheoremFailure___rarg___closed__1; 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*); static lean_object* l_Lean_Meta_mkInjectiveEqTheoremNameFor___closed__2; 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_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2___rarg(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_Injective___hyg_2126____closed__7; LEAN_EXPORT lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveEqTheoremValue(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveTheoremTypeCore_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -136,15 +139,19 @@ static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_2126____ static lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_solveEqOfCtorEq___closed__4; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_2126____closed__2; +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_solveEqOfCtorEq___closed__2; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); static lean_object* l_Lean_Meta_mkInjectiveTheoremNameFor___closed__2; static lean_object* l_panic___at___private_Lean_Meta_Injective_0__Lean_Meta_solveEqOfCtorEq___spec__1___closed__1; 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_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_mkInjectiveTheorems___closed__6; lean_object* lean_panic_fn(lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); +lean_object* l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkArrow(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Injective_0__Lean_Meta_throwInjectiveTheoremFailure___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_apply(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*); @@ -158,11 +165,13 @@ LEAN_EXPORT lean_object* l_Lean_Meta_elimOptParam___lambda__1___boxed(lean_objec size_t lean_usize_add(size_t, size_t); static lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Meta_Injective_0__Lean_Meta_mkAnd_x3f___spec__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_mkInjectiveEqTheoremNameFor(lean_object*); +static lean_object* l_Lean_Meta_mkInjectiveTheorems___closed__4; lean_object* lean_array_uget(lean_object*, size_t); LEAN_EXPORT lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveTheoremTypeCore_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_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_2126____closed__6; static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_2126____closed__5; 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_Meta_mkInjectiveTheorems___closed__8; lean_object* lean_array_get_size(lean_object*); static lean_object* l_Lean_Meta_elimOptParam___closed__1; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveTheoremTypeCore_x3f_mkArgs2___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -185,6 +194,7 @@ lean_object* l_Lean_addDecl(lean_object*, lean_object*, lean_object*, lean_objec LEAN_EXPORT lean_object* l_Lean_Meta_elimOptParam(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_withNewBinderInfos___at_Lean_Meta_withInstImplicitAsImplict___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Meta_Injective_0__Lean_Meta_mkAnd_x3f___spec__1___closed__2; +static lean_object* l_Lean_Meta_mkInjectiveTheorems___closed__7; LEAN_EXPORT lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveTheoremTypeCore_x3f(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Array_isEmpty___rarg(lean_object*); static lean_object* _init_l_Subarray_forInUnsafe_loop___at___private_Lean_Meta_Injective_0__Lean_Meta_mkAnd_x3f___spec__1___closed__1() { @@ -2771,13 +2781,13 @@ x_25 = lean_ctor_get(x_23, 1); lean_inc(x_25); lean_dec(x_23); x_26 = l_Lean_Meta_mkInjectiveTheoremNameFor(x_20); -x_27 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_18, x_2, x_3, x_4, x_5, x_25); +x_27 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_18, x_2, x_3, x_4, x_5, 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 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_24, x_2, x_3, x_4, x_5, x_29); +x_30 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_24, x_2, x_3, x_4, x_5, x_29); lean_dec(x_3); lean_dec(x_2); x_31 = lean_ctor_get(x_30, 0); @@ -2856,13 +2866,13 @@ x_46 = lean_ctor_get(x_44, 1); lean_inc(x_46); lean_dec(x_44); x_47 = l_Lean_Meta_mkInjectiveTheoremNameFor(x_42); -x_48 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_18, x_2, x_3, x_4, x_5, x_46); +x_48 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_18, x_2, x_3, x_4, x_5, 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); lean_dec(x_48); -x_51 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_45, x_2, x_3, x_4, x_5, x_50); +x_51 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_45, x_2, x_3, x_4, x_5, x_50); lean_dec(x_3); lean_dec(x_2); x_52 = lean_ctor_get(x_51, 0); @@ -3641,13 +3651,13 @@ x_25 = lean_ctor_get(x_23, 1); lean_inc(x_25); lean_dec(x_23); x_26 = l_Lean_Meta_mkInjectiveEqTheoremNameFor(x_20); -x_27 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_18, x_2, x_3, x_4, x_5, x_25); +x_27 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_18, x_2, x_3, x_4, x_5, 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 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_24, x_2, x_3, x_4, x_5, x_29); +x_30 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_24, x_2, x_3, x_4, x_5, x_29); x_31 = lean_ctor_get(x_30, 0); lean_inc(x_31); x_32 = lean_ctor_get(x_30, 1); @@ -3766,13 +3776,13 @@ x_56 = lean_ctor_get(x_54, 1); lean_inc(x_56); lean_dec(x_54); x_57 = l_Lean_Meta_mkInjectiveEqTheoremNameFor(x_52); -x_58 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_18, x_2, x_3, x_4, x_5, x_56); +x_58 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_18, x_2, x_3, x_4, x_5, x_56); x_59 = lean_ctor_get(x_58, 0); lean_inc(x_59); x_60 = lean_ctor_get(x_58, 1); lean_inc(x_60); lean_dec(x_58); -x_61 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_55, x_2, x_3, x_4, x_5, x_60); +x_61 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_55, x_2, x_3, x_4, x_5, x_60); x_62 = lean_ctor_get(x_61, 0); lean_inc(x_62); x_63 = lean_ctor_get(x_61, 1); @@ -3978,7 +3988,7 @@ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_2126____closed__2; x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_2126____closed__4; x_4 = l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_2126____closed__7; -x_5 = l_Lean_Option_register___at_Std_Format_initFn____x40_Lean_Data_Format___hyg_95____spec__1(x_2, x_3, x_4, x_1); +x_5 = l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__1(x_2, x_3, x_4, x_1); return x_5; } } @@ -4150,6 +4160,60 @@ return x_36; } } } +LEAN_EXPORT lean_object* l_Lean_Meta_mkInjectiveTheorems___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; +x_7 = lean_box(0); +x_8 = l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__1(x_1, x_7, x_2, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_8) == 0) +{ +uint8_t x_9; +x_9 = !lean_is_exclusive(x_8); +if (x_9 == 0) +{ +lean_object* x_10; +x_10 = lean_ctor_get(x_8, 0); +lean_dec(x_10); +lean_ctor_set(x_8, 0, x_7); +return x_8; +} +else +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_ctor_get(x_8, 1); +lean_inc(x_11); +lean_dec(x_8); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_7); +lean_ctor_set(x_12, 1, x_11); +return x_12; +} +} +else +{ +uint8_t x_13; +x_13 = !lean_is_exclusive(x_8); +if (x_13 == 0) +{ +return x_8; +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_8, 0); +x_15 = lean_ctor_get(x_8, 1); +lean_inc(x_15); +lean_inc(x_14); +lean_dec(x_8); +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; +} +} +} +} static lean_object* _init_l_Lean_Meta_mkInjectiveTheorems___closed__1() { _start: { @@ -4158,6 +4222,84 @@ x_1 = l_Lean_Meta_genInjectivity; return x_1; } } +static lean_object* _init_l_Lean_Meta_mkInjectiveTheorems___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_box(0), lean_box(0)); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_mkInjectiveTheorems___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_mkInjectiveTheorems___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_Meta_mkInjectiveTheorems___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_mkInjectiveTheorems___closed__3; +x_2 = lean_unsigned_to_nat(0u); +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_Meta_mkInjectiveTheorems___closed__5() { +_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_Meta_mkInjectiveTheorems___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_mkInjectiveTheorems___closed__5; +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_mkInjectiveTheorems___closed__7() { +_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_Meta_mkInjectiveTheorems___closed__6; +x_3 = l_Lean_Meta_mkInjectiveTheorems___closed__5; +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_Meta_mkInjectiveTheorems___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_mkInjectiveTheorems___closed__4; +x_2 = l_Lean_Meta_mkInjectiveTheorems___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; +} +} LEAN_EXPORT lean_object* l_Lean_Meta_mkInjectiveTheorems(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: { @@ -4208,7 +4350,7 @@ else { lean_object* x_19; uint8_t x_20; x_19 = l_Lean_Meta_mkInjectiveTheorems___closed__1; -x_20 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_11, x_19); +x_20 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_11, x_19); lean_dec(x_11); if (x_20 == 0) { @@ -4241,131 +4383,89 @@ lean_inc(x_24); x_25 = lean_ctor_get_uint8(x_24, sizeof(void*)*5 + 1); if (x_25 == 0) { -lean_object* x_26; 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_object* x_29; lean_object* x_30; lean_object* x_31; x_26 = lean_ctor_get(x_23, 1); lean_inc(x_26); lean_dec(x_23); x_27 = lean_ctor_get(x_24, 4); lean_inc(x_27); lean_dec(x_24); -x_28 = lean_box(0); -x_29 = l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__1(x_27, x_28, x_2, x_3, x_4, x_5, x_26); -if (lean_obj_tag(x_29) == 0) -{ -uint8_t x_30; -x_30 = !lean_is_exclusive(x_29); -if (x_30 == 0) -{ -lean_object* x_31; -x_31 = lean_ctor_get(x_29, 0); -lean_dec(x_31); -lean_ctor_set(x_29, 0, x_28); -return x_29; -} -else -{ -lean_object* x_32; lean_object* x_33; -x_32 = lean_ctor_get(x_29, 1); -lean_inc(x_32); -lean_dec(x_29); -x_33 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_33, 0, x_28); -lean_ctor_set(x_33, 1, x_32); -return x_33; -} -} -else -{ -uint8_t x_34; -x_34 = !lean_is_exclusive(x_29); -if (x_34 == 0) -{ -return x_29; -} -else -{ -lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_35 = lean_ctor_get(x_29, 0); -x_36 = lean_ctor_get(x_29, 1); -lean_inc(x_36); -lean_inc(x_35); -lean_dec(x_29); -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; -} -} +x_28 = lean_alloc_closure((void*)(l_Lean_Meta_mkInjectiveTheorems___lambda__1), 6, 1); +lean_closure_set(x_28, 0, x_27); +x_29 = l_Lean_Meta_mkInjectiveTheorems___closed__8; +x_30 = l___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveTheoremTypeCore_x3f___lambda__2___closed__1; +x_31 = l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2___rarg(x_29, x_30, x_28, x_2, x_3, x_4, x_5, x_26); +return x_31; } else { -uint8_t x_38; +uint8_t x_32; lean_dec(x_24); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_38 = !lean_is_exclusive(x_23); -if (x_38 == 0) +x_32 = !lean_is_exclusive(x_23); +if (x_32 == 0) { -lean_object* x_39; lean_object* x_40; -x_39 = lean_ctor_get(x_23, 0); -lean_dec(x_39); -x_40 = lean_box(0); -lean_ctor_set(x_23, 0, x_40); +lean_object* x_33; lean_object* x_34; +x_33 = lean_ctor_get(x_23, 0); +lean_dec(x_33); +x_34 = lean_box(0); +lean_ctor_set(x_23, 0, x_34); return x_23; } else { -lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_41 = lean_ctor_get(x_23, 1); -lean_inc(x_41); +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_23, 1); +lean_inc(x_35); lean_dec(x_23); -x_42 = lean_box(0); -x_43 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_43, 0, x_42); -lean_ctor_set(x_43, 1, x_41); -return x_43; +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_35); +return x_37; } } } else { -uint8_t x_44; +uint8_t x_38; lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_44 = !lean_is_exclusive(x_23); -if (x_44 == 0) +x_38 = !lean_is_exclusive(x_23); +if (x_38 == 0) { return x_23; } else { -lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_45 = lean_ctor_get(x_23, 0); -x_46 = lean_ctor_get(x_23, 1); -lean_inc(x_46); -lean_inc(x_45); +lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_23, 0); +x_40 = lean_ctor_get(x_23, 1); +lean_inc(x_40); +lean_inc(x_39); lean_dec(x_23); -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; +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 { -lean_object* x_48; +lean_object* x_42; lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_48 = lean_box(0); -lean_ctor_set(x_12, 0, x_48); +x_42 = lean_box(0); +lean_ctor_set(x_12, 0, x_42); return x_12; } } @@ -4373,195 +4473,153 @@ return x_12; } else { -lean_object* x_49; lean_object* x_50; lean_object* x_51; uint8_t x_52; -x_49 = lean_ctor_get(x_12, 0); -x_50 = lean_ctor_get(x_12, 1); -lean_inc(x_50); -lean_inc(x_49); +lean_object* x_43; lean_object* x_44; lean_object* x_45; uint8_t x_46; +x_43 = lean_ctor_get(x_12, 0); +x_44 = lean_ctor_get(x_12, 1); +lean_inc(x_44); +lean_inc(x_43); lean_dec(x_12); -x_51 = l___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveEqTheoremValue___lambda__2___closed__3; -x_52 = l_Lean_Environment_contains(x_10, x_51); -if (x_52 == 0) +x_45 = l___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveEqTheoremValue___lambda__2___closed__3; +x_46 = l_Lean_Environment_contains(x_10, x_45); +if (x_46 == 0) { -lean_object* x_53; lean_object* x_54; -lean_dec(x_49); +lean_object* x_47; lean_object* x_48; +lean_dec(x_43); lean_dec(x_11); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -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_50); -return x_54; +x_47 = lean_box(0); +x_48 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set(x_48, 1, x_44); +return x_48; } else { -lean_object* x_55; uint8_t x_56; -x_55 = l_Lean_Meta_mkInjectiveTheorems___closed__1; -x_56 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_11, x_55); +lean_object* x_49; uint8_t x_50; +x_49 = l_Lean_Meta_mkInjectiveTheorems___closed__1; +x_50 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_11, x_49); lean_dec(x_11); -if (x_56 == 0) +if (x_50 == 0) { -lean_object* x_57; lean_object* x_58; -lean_dec(x_49); +lean_object* x_51; lean_object* x_52; +lean_dec(x_43); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_57 = lean_box(0); -x_58 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_58, 0, x_57); -lean_ctor_set(x_58, 1, x_50); -return x_58; +x_51 = lean_box(0); +x_52 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_52, 0, x_51); +lean_ctor_set(x_52, 1, x_44); +return x_52; } else { -uint8_t x_59; -x_59 = lean_unbox(x_49); -lean_dec(x_49); -if (x_59 == 0) -{ -lean_object* x_60; -x_60 = l_Lean_getConstInfoInduct___at___private_Lean_Meta_RecursorInfo_0__Lean_Meta_mkRecursorInfoForKernelRec___spec__1(x_1, x_2, x_3, x_4, x_5, x_50); -if (lean_obj_tag(x_60) == 0) -{ -lean_object* x_61; uint8_t x_62; -x_61 = lean_ctor_get(x_60, 0); -lean_inc(x_61); -x_62 = lean_ctor_get_uint8(x_61, sizeof(void*)*5 + 1); -if (x_62 == 0) +uint8_t x_53; +x_53 = lean_unbox(x_43); +lean_dec(x_43); +if (x_53 == 0) { -lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_63 = lean_ctor_get(x_60, 1); -lean_inc(x_63); -lean_dec(x_60); -x_64 = lean_ctor_get(x_61, 4); -lean_inc(x_64); -lean_dec(x_61); -x_65 = lean_box(0); -x_66 = l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__1(x_64, x_65, x_2, x_3, x_4, x_5, x_63); -if (lean_obj_tag(x_66) == 0) +lean_object* x_54; +x_54 = l_Lean_getConstInfoInduct___at___private_Lean_Meta_RecursorInfo_0__Lean_Meta_mkRecursorInfoForKernelRec___spec__1(x_1, x_2, x_3, x_4, x_5, x_44); +if (lean_obj_tag(x_54) == 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); -if (lean_is_exclusive(x_66)) { - lean_ctor_release(x_66, 0); - lean_ctor_release(x_66, 1); - x_68 = x_66; -} else { - lean_dec_ref(x_66); - x_68 = lean_box(0); -} -if (lean_is_scalar(x_68)) { - x_69 = lean_alloc_ctor(0, 2, 0); -} else { - x_69 = x_68; -} -lean_ctor_set(x_69, 0, x_65); -lean_ctor_set(x_69, 1, x_67); -return x_69; -} -else +lean_object* x_55; uint8_t x_56; +x_55 = lean_ctor_get(x_54, 0); +lean_inc(x_55); +x_56 = lean_ctor_get_uint8(x_55, sizeof(void*)*5 + 1); +if (x_56 == 0) { -lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; -x_70 = lean_ctor_get(x_66, 0); -lean_inc(x_70); -x_71 = lean_ctor_get(x_66, 1); -lean_inc(x_71); -if (lean_is_exclusive(x_66)) { - lean_ctor_release(x_66, 0); - lean_ctor_release(x_66, 1); - x_72 = x_66; -} else { - lean_dec_ref(x_66); - x_72 = lean_box(0); -} -if (lean_is_scalar(x_72)) { - x_73 = lean_alloc_ctor(1, 2, 0); -} else { - x_73 = x_72; -} -lean_ctor_set(x_73, 0, x_70); -lean_ctor_set(x_73, 1, x_71); -return 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; +x_57 = lean_ctor_get(x_54, 1); +lean_inc(x_57); +lean_dec(x_54); +x_58 = lean_ctor_get(x_55, 4); +lean_inc(x_58); +lean_dec(x_55); +x_59 = lean_alloc_closure((void*)(l_Lean_Meta_mkInjectiveTheorems___lambda__1), 6, 1); +lean_closure_set(x_59, 0, x_58); +x_60 = l_Lean_Meta_mkInjectiveTheorems___closed__8; +x_61 = l___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveTheoremTypeCore_x3f___lambda__2___closed__1; +x_62 = l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2___rarg(x_60, x_61, x_59, x_2, x_3, x_4, x_5, x_57); +return x_62; } else { -lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; -lean_dec(x_61); +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +lean_dec(x_55); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_74 = lean_ctor_get(x_60, 1); -lean_inc(x_74); -if (lean_is_exclusive(x_60)) { - lean_ctor_release(x_60, 0); - lean_ctor_release(x_60, 1); - x_75 = x_60; +x_63 = lean_ctor_get(x_54, 1); +lean_inc(x_63); +if (lean_is_exclusive(x_54)) { + lean_ctor_release(x_54, 0); + lean_ctor_release(x_54, 1); + x_64 = x_54; } else { - lean_dec_ref(x_60); - x_75 = lean_box(0); + lean_dec_ref(x_54); + x_64 = lean_box(0); } -x_76 = lean_box(0); -if (lean_is_scalar(x_75)) { - x_77 = lean_alloc_ctor(0, 2, 0); +x_65 = lean_box(0); +if (lean_is_scalar(x_64)) { + x_66 = lean_alloc_ctor(0, 2, 0); } else { - x_77 = x_75; + x_66 = x_64; } -lean_ctor_set(x_77, 0, x_76); -lean_ctor_set(x_77, 1, x_74); -return x_77; +lean_ctor_set(x_66, 0, x_65); +lean_ctor_set(x_66, 1, x_63); +return x_66; } } else { -lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_78 = lean_ctor_get(x_60, 0); -lean_inc(x_78); -x_79 = lean_ctor_get(x_60, 1); -lean_inc(x_79); -if (lean_is_exclusive(x_60)) { - lean_ctor_release(x_60, 0); - lean_ctor_release(x_60, 1); - x_80 = x_60; +x_67 = lean_ctor_get(x_54, 0); +lean_inc(x_67); +x_68 = lean_ctor_get(x_54, 1); +lean_inc(x_68); +if (lean_is_exclusive(x_54)) { + lean_ctor_release(x_54, 0); + lean_ctor_release(x_54, 1); + x_69 = x_54; } else { - lean_dec_ref(x_60); - x_80 = lean_box(0); + lean_dec_ref(x_54); + x_69 = lean_box(0); } -if (lean_is_scalar(x_80)) { - x_81 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_69)) { + x_70 = lean_alloc_ctor(1, 2, 0); } else { - x_81 = x_80; + x_70 = x_69; } -lean_ctor_set(x_81, 0, x_78); -lean_ctor_set(x_81, 1, x_79); -return x_81; +lean_ctor_set(x_70, 0, x_67); +lean_ctor_set(x_70, 1, x_68); +return x_70; } } else { -lean_object* x_82; lean_object* x_83; +lean_object* x_71; lean_object* x_72; lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_82 = lean_box(0); -x_83 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_83, 0, x_82); -lean_ctor_set(x_83, 1, x_50); -return x_83; +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_44); +return x_72; } } } @@ -4569,7 +4627,7 @@ return x_83; } else { -uint8_t x_84; +uint8_t x_73; lean_dec(x_11); lean_dec(x_10); lean_dec(x_5); @@ -4577,23 +4635,23 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_84 = !lean_is_exclusive(x_12); -if (x_84 == 0) +x_73 = !lean_is_exclusive(x_12); +if (x_73 == 0) { return x_12; } else { -lean_object* x_85; lean_object* x_86; lean_object* x_87; -x_85 = lean_ctor_get(x_12, 0); -x_86 = lean_ctor_get(x_12, 1); -lean_inc(x_86); -lean_inc(x_85); +lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_74 = lean_ctor_get(x_12, 0); +x_75 = lean_ctor_get(x_12, 1); +lean_inc(x_75); +lean_inc(x_74); lean_dec(x_12); -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; +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; } } } @@ -4732,6 +4790,20 @@ lean_mark_persistent(l_Lean_Meta_genInjectivity); lean_dec_ref(res); }l_Lean_Meta_mkInjectiveTheorems___closed__1 = _init_l_Lean_Meta_mkInjectiveTheorems___closed__1(); lean_mark_persistent(l_Lean_Meta_mkInjectiveTheorems___closed__1); +l_Lean_Meta_mkInjectiveTheorems___closed__2 = _init_l_Lean_Meta_mkInjectiveTheorems___closed__2(); +lean_mark_persistent(l_Lean_Meta_mkInjectiveTheorems___closed__2); +l_Lean_Meta_mkInjectiveTheorems___closed__3 = _init_l_Lean_Meta_mkInjectiveTheorems___closed__3(); +lean_mark_persistent(l_Lean_Meta_mkInjectiveTheorems___closed__3); +l_Lean_Meta_mkInjectiveTheorems___closed__4 = _init_l_Lean_Meta_mkInjectiveTheorems___closed__4(); +lean_mark_persistent(l_Lean_Meta_mkInjectiveTheorems___closed__4); +l_Lean_Meta_mkInjectiveTheorems___closed__5 = _init_l_Lean_Meta_mkInjectiveTheorems___closed__5(); +lean_mark_persistent(l_Lean_Meta_mkInjectiveTheorems___closed__5); +l_Lean_Meta_mkInjectiveTheorems___closed__6 = _init_l_Lean_Meta_mkInjectiveTheorems___closed__6(); +lean_mark_persistent(l_Lean_Meta_mkInjectiveTheorems___closed__6); +l_Lean_Meta_mkInjectiveTheorems___closed__7 = _init_l_Lean_Meta_mkInjectiveTheorems___closed__7(); +lean_mark_persistent(l_Lean_Meta_mkInjectiveTheorems___closed__7); +l_Lean_Meta_mkInjectiveTheorems___closed__8 = _init_l_Lean_Meta_mkInjectiveTheorems___closed__8(); +lean_mark_persistent(l_Lean_Meta_mkInjectiveTheorems___closed__8); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lean/Meta/Instances.c b/stage0/stdlib/Lean/Meta/Instances.c index 8251e6777873..5c5d83399337 100644 --- a/stage0/stdlib/Lean/Meta/Instances.c +++ b/stage0/stdlib/Lean/Meta/Instances.c @@ -15,7 +15,6 @@ extern "C" { #endif LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_addInstanceEntry___spec__3(uint8_t, lean_object*, size_t, lean_object*); lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); -uint8_t l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Instances___hyg_3105____closed__9; LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at_Lean_Meta_computeSynthOrder___spec__9___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* l___private_Init_Util_0__outOfBounds___rarg(lean_object*); @@ -234,7 +233,6 @@ lean_object* lean_st_mk_ref(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_addInstanceEntry(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_computeSynthOrder___spec__8___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_to_list(lean_object*, lean_object*); -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Instances___hyg_6____closed__7; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_addInstanceEntry___spec__19___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Instances___hyg_3105____closed__12; @@ -286,6 +284,7 @@ extern lean_object* l_Lean_warningAsError; static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Instances___hyg_3105____lambda__1___closed__7; LEAN_EXPORT lean_object* l_Lean_Meta_getErasedInstances___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Instances___hyg_4333_(lean_object*); +uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addInstanceEntry___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Instances___hyg_537_(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_getInstancePriority_x3f___spec__1(lean_object*, lean_object*); @@ -325,7 +324,6 @@ lean_object* lean_usize_to_nat(size_t); lean_object* l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Meta_WHNF_0__Lean_Meta_toCtorWhenK___spec__1___rarg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_instInhabitedInstanceEntry___closed__2; static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_addInstance___spec__2___closed__7; -lean_object* l_Lean_Option_register___at_Std_Format_initFn____x40_Lean_Data_Format___hyg_95____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofExpr(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertVal___at_Lean_Meta_addInstanceEntry___spec__10(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_isInstance(lean_object*, lean_object*, lean_object*, lean_object*); @@ -373,6 +371,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_getDefaultInstances(lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Instances___hyg_4333____closed__1; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getDefaultInstancesPriorities___rarg___lambda__1___boxed(lean_object*, lean_object*); +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Instances___hyg_3105____lambda__1___closed__11; static lean_object* l_Lean_Meta_Instances_instanceNames___default___closed__2; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_addInstanceEntry___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -402,6 +401,7 @@ lean_object* lean_panic_fn(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_computeSynthOrder___spec__11___at_Lean_Meta_computeSynthOrder___spec__12(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_computeSynthOrder___lambda__1___closed__4; lean_object* lean_nat_sub(lean_object*, lean_object*); +lean_object* l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Instances___hyg_3929____closed__3; lean_object* l_Lean_ScopedEnvExtension_getState___rarg(lean_object*, lean_object*, lean_object*); lean_object* l_Array_insertAt_x21___rarg(lean_object*, lean_object*, lean_object*); @@ -578,7 +578,7 @@ static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Instances___hyg_6_ _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("check instances do not introduce metavariable in non-out-params", 63); +x_1 = lean_mk_string_from_bytes("check that instances do not introduce metavariable in non-out-params", 68); return x_1; } } @@ -632,7 +632,7 @@ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Instances___hyg_6____closed__3; x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_Instances___hyg_6____closed__6; x_4 = l_Lean_Meta_initFn____x40_Lean_Meta_Instances___hyg_6____closed__9; -x_5 = l_Lean_Option_register___at_Std_Format_initFn____x40_Lean_Data_Format___hyg_95____spec__1(x_2, x_3, x_4, x_1); +x_5 = l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__1(x_2, x_3, x_4, x_1); return x_5; } } @@ -3848,7 +3848,7 @@ static lean_object* _init_l_Lean_Meta_DiscrTree_insertCore___at_Lean_Meta_addIns 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_Meta_DiscrTree_insertCore___at_Lean_Meta_addInstanceEntry___spec__1___closed__1; x_2 = l_Lean_Meta_DiscrTree_insertCore___at_Lean_Meta_addInstanceEntry___spec__1___closed__2; -x_3 = lean_unsigned_to_nat(366u); +x_3 = lean_unsigned_to_nat(396u); x_4 = lean_unsigned_to_nat(23u); x_5 = l_Lean_Meta_DiscrTree_insertCore___at_Lean_Meta_addInstanceEntry___spec__1___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -8002,7 +8002,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_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_20, x_6, x_7, x_8, x_9, x_21); +x_22 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_20, x_6, x_7, x_8, x_9, x_21); x_23 = lean_ctor_get(x_22, 0); lean_inc(x_23); x_24 = lean_ctor_get(x_22, 1); @@ -8065,7 +8065,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_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_38, x_6, x_7, x_8, x_9, x_39); +x_40 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_38, x_6, x_7, x_8, x_9, x_39); x_41 = lean_ctor_get(x_40, 0); lean_inc(x_41); x_42 = lean_ctor_get(x_40, 1); @@ -8175,7 +8175,7 @@ else { lean_object* x_263; uint8_t x_264; x_263 = l_Lean_logAt___at_Lean_Meta_computeSynthOrder___spec__7___closed__1; -x_264 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_260, x_263); +x_264 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_260, x_263); if (x_264 == 0) { x_10 = x_3; @@ -9253,7 +9253,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_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_19, x_9, x_10, x_11, x_12, x_20); +x_21 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_19, x_9, x_10, x_11, x_12, x_20); x_22 = lean_ctor_get(x_21, 0); lean_inc(x_22); x_23 = lean_ctor_get(x_21, 1); @@ -10018,7 +10018,7 @@ lean_closure_set(x_25, 4, x_24); x_26 = lean_ctor_get(x_13, 2); lean_inc(x_26); x_27 = l_Lean_Loop_forIn_loop___at_Lean_Meta_computeSynthOrder___spec__9___closed__2; -x_28 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_26, x_27); +x_28 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_26, x_27); lean_dec(x_26); if (x_28 == 0) { @@ -10572,7 +10572,7 @@ lean_closure_set(x_123, 4, x_122); x_124 = lean_ctor_get(x_13, 2); lean_inc(x_124); x_125 = l_Lean_Loop_forIn_loop___at_Lean_Meta_computeSynthOrder___spec__9___closed__2; -x_126 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_124, x_125); +x_126 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_124, x_125); lean_dec(x_124); if (x_126 == 0) { @@ -11657,7 +11657,7 @@ lean_dec(x_40); x_43 = lean_ctor_get(x_11, 2); lean_inc(x_43); x_44 = l_Lean_Loop_forIn_loop___at_Lean_Meta_computeSynthOrder___spec__9___closed__2; -x_45 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_43, x_44); +x_45 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_43, x_44); lean_dec(x_43); if (x_45 == 0) { @@ -11671,7 +11671,7 @@ return x_47; else { lean_object* x_48; lean_object* x_49; lean_object* x_50; uint8_t x_51; -x_48 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_1, x_9, x_10, x_11, x_12, x_41); +x_48 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_1, x_9, x_10, x_11, x_12, x_41); x_49 = lean_ctor_get(x_48, 0); lean_inc(x_49); x_50 = lean_ctor_get(x_48, 1); diff --git a/stage0/stdlib/Lean/Meta/KAbstract.c b/stage0/stdlib/Lean/Meta/KAbstract.c index 5a7893ad5ad0..e8ed9e1146ea 100644 --- a/stage0/stdlib/Lean/Meta/KAbstract.c +++ b/stage0/stdlib/Lean/Meta/KAbstract.c @@ -30,11 +30,11 @@ LEAN_EXPORT lean_object* l_Lean_Meta_kabstract_visit___boxed(lean_object*, lean_ lean_object* l_Lean_Expr_forallE___override(lean_object*, lean_object*, lean_object*, uint8_t); lean_object* lean_st_ref_get(lean_object*, lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_hasLooseBVars(lean_object*); lean_object* l_Lean_Expr_toHeadIndex(lean_object*); lean_object* l_Lean_Expr_letE___override(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); lean_object* l_Lean_Expr_app___override(lean_object*, lean_object*); +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l___private_Lean_HeadIndex_0__Lean_beqHeadIndex____x40_Lean_HeadIndex___hyg_67_(lean_object*, lean_object*); uint8_t l___private_Lean_Expr_0__Lean_beqBinderInfo____x40_Lean_Expr___hyg_473_(uint8_t, uint8_t); static lean_object* l_Lean_Meta_kabstract___closed__1; @@ -5801,7 +5801,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_kabstract(lean_object* x_1, lean_object* x_ _start: { lean_object* x_9; uint8_t x_10; -x_9 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_1, x_4, x_5, x_6, x_7, x_8); +x_9 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_1, x_4, x_5, x_6, x_7, x_8); x_10 = !lean_is_exclusive(x_9); if (x_10 == 0) { diff --git a/stage0/stdlib/Lean/Meta/Match/Basic.c b/stage0/stdlib/Lean/Meta/Match/Basic.c index 88daa19da669..60a37518b332 100644 --- a/stage0/stdlib/Lean/Meta/Match/Basic.c +++ b/stage0/stdlib/Lean/Meta/Match/Basic.c @@ -152,7 +152,6 @@ static lean_object* l_Lean_Meta_Match_toPattern___closed__3; lean_object* l_Lean_Meta_mkAppM(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forM___at_Lean_Meta_Match_Pattern_collectFVars___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_to_list(lean_object*, lean_object*); -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Match_mkNamedPattern(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Match_Pattern_toMessageData___closed__10; LEAN_EXPORT lean_object* l_Lean_Meta_Match_Pattern_applyFVarSubst___boxed(lean_object*, lean_object*); @@ -209,6 +208,7 @@ static lean_object* l_Lean_Meta_Match_Pattern_toMessageData___closed__13; LEAN_EXPORT lean_object* l_List_forM___at_Lean_Meta_Match_AltLHS_collectFVars___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Match_Example_toMessageData___closed__3; static lean_object* l_Lean_Meta_Match_Problem_toMessageData___lambda__1___closed__1; +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Match_instantiatePatternMVars___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Match_Alt_checkAndReplaceFVarId___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_indentExpr(lean_object*); @@ -2278,7 +2278,7 @@ if (x_10 == 0) lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; x_11 = lean_ctor_get(x_1, 0); x_12 = lean_ctor_get(x_1, 1); -x_13 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_11, x_3, x_4, x_5, x_6, x_7); +x_13 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_11, x_3, x_4, x_5, x_6, x_7); x_14 = lean_ctor_get(x_13, 0); lean_inc(x_14); x_15 = lean_ctor_get(x_13, 1); @@ -2304,7 +2304,7 @@ x_18 = lean_ctor_get(x_1, 1); lean_inc(x_18); lean_inc(x_17); lean_dec(x_1); -x_19 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_17, x_3, x_4, x_5, x_6, x_7); +x_19 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_17, x_3, x_4, x_5, x_6, x_7); x_20 = lean_ctor_get(x_19, 0); lean_inc(x_20); x_21 = lean_ctor_get(x_19, 1); @@ -2397,7 +2397,7 @@ if (x_7 == 0) { lean_object* x_8; lean_object* x_9; uint8_t x_10; x_8 = lean_ctor_get(x_1, 0); -x_9 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_8, x_2, x_3, x_4, x_5, x_6); +x_9 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_8, x_2, x_3, x_4, x_5, x_6); x_10 = !lean_is_exclusive(x_9); if (x_10 == 0) { @@ -2428,7 +2428,7 @@ lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean x_15 = lean_ctor_get(x_1, 0); lean_inc(x_15); lean_dec(x_1); -x_16 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_15, x_2, x_3, x_4, x_5, x_6); +x_16 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_15, x_2, x_3, x_4, x_5, x_6); x_17 = lean_ctor_get(x_16, 0); lean_inc(x_17); x_18 = lean_ctor_get(x_16, 1); @@ -2559,7 +2559,7 @@ if (x_50 == 0) { lean_object* x_51; lean_object* x_52; uint8_t x_53; x_51 = lean_ctor_get(x_1, 0); -x_52 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_51, x_2, x_3, x_4, x_5, x_6); +x_52 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_51, x_2, x_3, x_4, x_5, x_6); x_53 = !lean_is_exclusive(x_52); if (x_53 == 0) { @@ -2590,7 +2590,7 @@ lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean x_58 = lean_ctor_get(x_1, 0); lean_inc(x_58); lean_dec(x_1); -x_59 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_58, x_2, x_3, x_4, x_5, x_6); +x_59 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_58, x_2, x_3, x_4, x_5, x_6); x_60 = lean_ctor_get(x_59, 0); lean_inc(x_60); x_61 = lean_ctor_get(x_59, 1); @@ -2624,7 +2624,7 @@ if (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; uint8_t x_73; x_66 = lean_ctor_get(x_1, 0); x_67 = lean_ctor_get(x_1, 1); -x_68 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_66, x_2, x_3, x_4, x_5, x_6); +x_68 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_66, x_2, x_3, x_4, x_5, x_6); x_69 = lean_ctor_get(x_68, 0); lean_inc(x_69); x_70 = lean_ctor_get(x_68, 1); @@ -2666,7 +2666,7 @@ x_79 = lean_ctor_get(x_1, 1); lean_inc(x_79); lean_inc(x_78); lean_dec(x_1); -x_80 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_78, x_2, x_3, x_4, x_5, x_6); +x_80 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_78, x_2, x_3, x_4, x_5, x_6); x_81 = lean_ctor_get(x_80, 0); lean_inc(x_81); x_82 = lean_ctor_get(x_80, 1); @@ -2935,7 +2935,7 @@ if (x_7 == 0) { lean_object* x_8; lean_object* x_9; uint8_t x_10; x_8 = lean_ctor_get(x_1, 3); -x_9 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_8, x_2, x_3, x_4, x_5, x_6); +x_9 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_8, x_2, x_3, x_4, x_5, x_6); x_10 = !lean_is_exclusive(x_9); if (x_10 == 0) { @@ -2974,7 +2974,7 @@ lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_dec(x_1); -x_21 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_18, x_2, x_3, x_4, x_5, x_6); +x_21 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_18, x_2, x_3, x_4, x_5, x_6); x_22 = lean_ctor_get(x_21, 0); lean_inc(x_22); x_23 = lean_ctor_get(x_21, 1); @@ -3013,13 +3013,13 @@ 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; uint8_t x_34; x_28 = lean_ctor_get(x_1, 3); x_29 = lean_ctor_get(x_1, 4); -x_30 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_28, x_2, x_3, x_4, x_5, x_6); +x_30 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_28, x_2, x_3, x_4, x_5, x_6); 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_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_29, x_2, x_3, x_4, x_5, x_32); +x_33 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_29, x_2, x_3, x_4, x_5, x_32); x_34 = !lean_is_exclusive(x_33); if (x_34 == 0) { @@ -3062,13 +3062,13 @@ lean_inc(x_41); lean_inc(x_40); lean_inc(x_39); lean_dec(x_1); -x_46 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_42, x_2, x_3, x_4, x_5, x_6); +x_46 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_42, x_2, x_3, x_4, x_5, x_6); 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_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_43, x_2, x_3, x_4, x_5, x_48); +x_49 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_43, x_2, x_3, x_4, x_5, x_48); x_50 = lean_ctor_get(x_49, 0); lean_inc(x_50); x_51 = lean_ctor_get(x_49, 1); diff --git a/stage0/stdlib/Lean/Meta/Match/Match.c b/stage0/stdlib/Lean/Meta/Match/Match.c index 436453e671c4..60d917ae6400 100644 --- a/stage0/stdlib/Lean/Meta/Match/Match.c +++ b/stage0/stdlib/Lean/Meta/Match/Match.c @@ -18,7 +18,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_pr LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Meta_Match_mkMatcher___spec__9___at_Lean_Meta_Match_mkMatcher___spec__10(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); lean_object* l_Lean_Meta_FVarSubst_apply(lean_object*, lean_object*); -uint8_t l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_withAlts_loop___rarg___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_Match_0__Lean_Meta_updateAlts___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_Meta_mkEqHEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -285,7 +284,6 @@ extern lean_object* l_instInhabitedNat; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_Match_mkMatcher___spec__6(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Meta_Match_withCleanLCtxFor___spec__24___boxed(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_hasNonTrivialExample___spec__1___boxed(lean_object*, lean_object*); -lean_object* l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Match_MkMatcherInput_collectDependencies(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Match_mkMatcherAuxDefinition___lambda__2___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_Match_mkMatcher___lambda__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_object*); @@ -548,6 +546,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_so LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Match_withCleanLCtxFor___spec__25___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Match_mkMatcher___lambda__4(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_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_expandVarIntoCtor_x3f___lambda__5(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___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Meta_Match_withCleanLCtxFor___spec__17(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processAsPattern(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__24(lean_object*, lean_object*, size_t, size_t, lean_object*); @@ -619,7 +618,6 @@ lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__2___rarg(lean_obj LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_solveCnstrs_go___spec__3___boxed(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_panic___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processConstructor___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Option_register___at_Std_Format_initFn____x40_Lean_Data_Format___hyg_95____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_inheritedTraceOptions; lean_object* l_Lean_MessageData_ofExpr(lean_object*); static lean_object* l_Lean_Meta_Match_getMkMatcherInputInContext___lambda__3___closed__1; @@ -656,6 +654,7 @@ LEAN_EXPORT lean_object* l_List_filterTR_loop___at___private_Lean_Meta_Match_Mat lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_instantiateLambdaAux(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_panic___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processSkipInaccessible___spec__2___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_hasAsPattern___boxed(lean_object*); +lean_object* l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Match_mkMatcher___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_Meta_Match_getMkMatcherInputInContext___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Match_mkMatcher___lambda__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -758,6 +757,7 @@ LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_Match_processInaccessibleAsCtor_ LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_processValue___spec__9(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashSetImp_contains___at_Lean_Meta_Match_mkMatcher___spec__3___boxed(lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); +lean_object* l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Match_MkMatcherInput_collectDependencies___boxed(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_processNonVariable(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_withAlts_loop___rarg___closed__12; @@ -21807,7 +21807,7 @@ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_2 = l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_Match___hyg_10753____closed__3; x_3 = l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_Match___hyg_10753____closed__5; x_4 = l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_Match___hyg_10753____closed__7; -x_5 = l_Lean_Option_register___at_Std_Format_initFn____x40_Lean_Data_Format___hyg_95____spec__1(x_2, x_3, x_4, x_1); +x_5 = l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__1(x_2, x_3, x_4, x_1); return x_5; } } @@ -23772,7 +23772,7 @@ lean_dec(x_5); x_11 = lean_ctor_get(x_8, 2); lean_inc(x_11); x_12 = l_Lean_Meta_Match_mkMatcherAuxDefinition___lambda__4___closed__1; -x_13 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_11, x_12); +x_13 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_11, x_12); lean_dec(x_11); x_14 = 0; lean_inc(x_9); @@ -25759,7 +25759,7 @@ lean_dec(x_16); lean_dec(x_14); lean_dec(x_9); x_19 = l___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_withEqs___rarg___closed__1; -x_20 = l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__3___rarg(x_12, x_19, x_2, x_3, x_4, x_5, x_6, x_15); +x_20 = l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2___rarg(x_12, x_19, x_2, x_3, x_4, x_5, x_6, x_15); return x_20; } else @@ -25773,7 +25773,7 @@ lean_dec(x_16); lean_dec(x_14); lean_dec(x_9); x_22 = l___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_withEqs___rarg___closed__1; -x_23 = l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__3___rarg(x_12, x_22, x_2, x_3, x_4, x_5, x_6, x_15); +x_23 = l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2___rarg(x_12, x_22, x_2, x_3, x_4, x_5, x_6, x_15); return x_23; } else @@ -25786,7 +25786,7 @@ x_26 = l___private_Lean_Meta_Match_Match_0__Lean_Meta_Match_withEqs___rarg___clo x_27 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_Match_withCleanLCtxFor___spec__25(x_9, x_14, x_24, x_25, x_26); lean_dec(x_14); lean_dec(x_9); -x_28 = l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__3___rarg(x_12, x_27, x_2, x_3, x_4, x_5, x_6, x_15); +x_28 = l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2___rarg(x_12, x_27, x_2, x_3, x_4, x_5, x_6, x_15); return x_28; } } diff --git a/stage0/stdlib/Lean/Meta/Match/MatchEqs.c b/stage0/stdlib/Lean/Meta/Match/MatchEqs.c index 93754a4985bf..ad4d512bf203 100644 --- a/stage0/stdlib/Lean/Meta/Match/MatchEqs.c +++ b/stage0/stdlib/Lean/Meta/Match/MatchEqs.c @@ -28,6 +28,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match LEAN_EXPORT uint8_t l_Lean_Meta_Match_forallAltTelescope_isNamedPatternProof___lambda__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertCastEqRec_go___closed__8; LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_SimpH_isDone___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__15; static lean_object* l_panic___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertCastEqRec_go___spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertCastEqRec___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_EXPORT lean_object* l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_proveSubgoalLoop___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -66,6 +67,7 @@ LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_Match_M LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertCastEqRec___spec__4___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_Match_proveCondEqThm___lambda__3___closed__1; static lean_object* l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_isCastEqRec___closed__3; +static lean_object* l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__7; static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Meta_Match_proveCondEqThm_go___spec__2___closed__2; static lean_object* l_Lean_Meta_Match_forallAltTelescope_go___rarg___closed__11; LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertTemplate___spec__7___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*, lean_object*); @@ -114,6 +116,7 @@ lean_object* lean_array_push(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Match_forallAltTelescope_go___rarg___lambda__5___closed__2; lean_object* l_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertCastEqRec_go___spec__5___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_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit_visitLambda___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertTemplate___spec__9___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*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertCastEqRec___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_throwMaxRecDepthAt___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertTemplate___spec__19___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -128,6 +131,7 @@ lean_object* l_ST_Prim_Ref_get___boxed(lean_object*, lean_object*, lean_object*, static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertCastEqRec___spec__6___lambda__2___closed__2; lean_object* l_Lean_Expr_mdata___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_substSomeVar___spec__5(lean_object*, lean_object*, 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_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__4; static lean_object* l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertCastEqRec_go___closed__2; static lean_object* l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_simpH_x3f___lambda__3___closed__4; LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertCastEqRec___spec__8___lambda__2___boxed(lean_object**); @@ -152,7 +156,6 @@ static lean_object* l_Lean_Meta_Match_proveCondEqThm_go___lambda__2___closed__9; static lean_object* l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_substSomeVar___lambda__1___closed__1; uint8_t l_Lean_Expr_isNatLit(lean_object*); lean_object* l_Lean_MVarId_refl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__12; LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkEquationsFor___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_EXPORT lean_object* l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_withSplitterAlts(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*); @@ -165,7 +168,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Matc LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertCastEqRec___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_casesOnStuckLHS___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_simpH_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__2; lean_object* l_Lean_stringToMessageData(lean_object*); lean_object* l_Lean_Expr_mvar___override(lean_object*); LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertTemplate___spec__3___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertTemplate___spec__4___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertTemplate___spec__5(lean_object*, lean_object*); @@ -180,7 +182,6 @@ static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_Match_MatchE LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit_visitLet___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertTemplate___spec__11___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*, 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*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkEquationsFor___spec__5___boxed(lean_object**); -static lean_object* l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__11; lean_object* l_Lean_Exception_toMessageData(lean_object*); LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertCastEqRec___spec__8___lambda__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_Lean_Meta_Match_forallAltTelescope_go___rarg___closed__7; @@ -197,11 +198,11 @@ lean_object* l_Lean_Expr_appArg_x21(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertTemplate___spec__12___rarg(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_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_substSomeVar___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_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_proveSubgoalLoop___closed__2; -static lean_object* l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__8; extern lean_object* l_instInhabitedNat; LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertCastEqRec___spec__8___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_EXPORT lean_object* l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertTemplate___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_mapTR_loop___at_Lean_mkConstWithLevelParams___spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__13; LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertTemplate___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*); static lean_object* l_Lean_Meta_Match_forallAltTelescope_go___rarg___lambda__5___closed__1; static lean_object* l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkEquationsFor___lambda__4___closed__1; @@ -212,10 +213,9 @@ LEAN_EXPORT lean_object* l_Lean_MVarId_withContext___at___private_Lean_Meta_Matc extern lean_object* l_instInhabitedPUnit; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_injectionAny___spec__3(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_panic___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_SimpH_substRHS___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__11; LEAN_EXPORT lean_object* l_Lean_Meta_forallBoundedTelescope___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertCastEqRec___spec__7(lean_object*); -lean_object* l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__3___rarg(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_MatchEqs_0__Lean_Meta_Match_mkSplitterProof___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_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__16; lean_object* l_Lean_Meta_trySubst(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Match_proveCondEqThm_go___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Match_proveCondEqThm___lambda__1___boxed(lean_object*, lean_object*); @@ -224,6 +224,7 @@ lean_object* l_Lean_MVarId_contradiction(lean_object*, lean_object*, lean_object static lean_object* l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_SimpH_substRHS___closed__5; uint8_t l_Lean_Expr_isLambda(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_proveSubgoalLoop___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488_(lean_object*); static lean_object* l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_withSplitterAlts_go___rarg___closed__1; static lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertCastEqRec___spec__8___lambda__9___closed__1; size_t lean_ptr_addr(lean_object*); @@ -273,7 +274,6 @@ lean_object* l_Lean_EnvExtensionInterfaceUnsafe_getState___rarg(lean_object*, le static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkEquationsFor___spec__5___lambda__3___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_SimpH_processNextEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_injectionAny___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_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__5; uint8_t lean_expr_eqv(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_SimpH_isDone(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkPrivateName(lean_object*, lean_object*); @@ -306,12 +306,12 @@ lean_object* l_Lean_Meta_SavedState_restore(lean_object*, lean_object*, lean_obj lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDeclImp___rarg(lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_proveSubgoalLoop___lambda__1___closed__4; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkEquationsFor___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__10; uint8_t l_List_elem___at_Lean_MVarId_getNondepPropHyps___spec__3(lean_object*, lean_object*); static lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_isCastEqRec___spec__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertCastEqRec___spec__8___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_Meta_transform_visit_visitPost___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertTemplate___spec__8(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_Std_Range_forIn_loop___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkEquationsFor___spec__5___lambda__3___boxed(lean_object**); +static lean_object* l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__12; static lean_object* l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_proveSubgoalLoop___lambda__1___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_substSomeVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_substCore(lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -325,7 +325,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match 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*); static lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertCastEqRec___spec__8___lambda__2___closed__4; lean_object* l_Lean_FVarId_getDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__13; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkEquationsFor___spec__5___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* l_Lean_Meta_whnfForall(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertTemplate___spec__14___boxed(lean_object*, lean_object*); @@ -343,17 +342,14 @@ LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Match_M lean_object* l_Lean_Meta_splitIfTarget_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); lean_object* l_Lean_Expr_replaceFVar(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__3; lean_object* l_Lean_RecursorVal_getMajorIdx(lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_find___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_isCastEqRec___spec__1___boxed(lean_object*, lean_object*); lean_object* lean_array_to_list(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkEquationsFor___spec__4(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withMVarContextImp___rarg(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_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkEquationsFor___spec__5___closed__2; lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__6; lean_object* l_Lean_Name_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_SimpH_processNextEq___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Match_forallAltTelescope_go(lean_object*); @@ -361,6 +357,7 @@ lean_object* l_instInhabited___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit_visitLet___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertTemplate___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_EXPORT lean_object* l_Lean_Meta_casesOnStuckLHS_findFVar_x3f___lambda__1(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_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkEquationsFor___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_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__3; static lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertCastEqRec___spec__8___lambda__11___closed__1; static lean_object* l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_simpH_x3f___closed__3; lean_object* l_ST_Prim_Ref_modifyGetUnsafe___rarg___boxed(lean_object*, lean_object*, lean_object*); @@ -379,7 +376,9 @@ static lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Match_MatchE static lean_object* l_Lean_Meta_Match_forallAltTelescope_go___rarg___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertCastEqRec_go___lambda__2___boxed(lean_object**); static lean_object* l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertTemplate___lambda__3___closed__5; +static lean_object* l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__10; LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertCastEqRec___spec__8___lambda__10___boxed(lean_object**); +static lean_object* l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__9; static lean_object* l_List_mapTR_loop___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertTemplate___spec__3___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertTemplate___spec__4___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertTemplate___spec__5___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_withNewAlts___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_Meta_Match_MatchEqs_0__Lean_Meta_Match_simpH_x3f___lambda__3___closed__3; @@ -389,11 +388,12 @@ lean_object* l_Lean_RBNode_insert___at_Lean_FVarIdMap_insert___spec__1___rarg(le LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_mkMap___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_injectionAny___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__14; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_transform___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertTemplate___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* l_Lean_Meta_injection(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_isConstructorApp_x3f(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_SimpH_processNextEq___lambda__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___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_SimpH_processNextEq___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_EXPORT lean_object* l_Lean_Meta_transform_visit___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertTemplate___spec__7(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___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkEquationsFor___lambda__5(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_MatchEqs_0__Lean_Meta_Match_injectionAnyCandidate_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -402,7 +402,6 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Match_M static lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertCastEqRec___spec__8___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_transform___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertTemplate___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* l_Lean_Meta_simpIfTarget(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__1; static lean_object* l_Lean_Meta_Match_proveCondEqThm_go___lambda__2___closed__2; lean_object* l_Lean_Meta_forallTelescope___at___private_Lean_Meta_InferType_0__Lean_Meta_inferForallType___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertTemplate___lambda__3___closed__3; @@ -413,8 +412,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___at___private_Lean_Meta_Ma LEAN_EXPORT lean_object* l_Lean_Meta_casesOnStuckLHS___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Match_getEquationsForImpl___spec__1(lean_object*, lean_object*); lean_object* l_Lean_Meta_saveState___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__7; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_SimpH_processNextEq___lambda__4(lean_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_Match_MatchEqs_0__Lean_Meta_Match_SimpH_processNextEq___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_Meta_forallTelescopeReducing___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertCastEqRec___spec__3___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_Lean_Meta_transform_visit___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertTemplate___spec__7___lambda__2(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_LocalDecl_fvarId(lean_object*); @@ -476,7 +474,6 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Meta_Match lean_object* l_Lean_Meta_forallMetaBoundedTelescope(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertCastEqRec___closed__1; static lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertCastEqRec___spec__8___lambda__2___closed__2; -static lean_object* l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__4; static lean_object* l_List_mapTR_loop___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertTemplate___spec__3___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertTemplate___spec__4___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertTemplate___spec__5___closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_substSomeVar___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_EXPORT lean_object* l_Lean_Meta_Match_proveCondEqThm(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -488,6 +485,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescopeReducing___at___private_Lean lean_object* l_Subarray_get___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertCastEqRec_go___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_mapTR_loop___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertTemplate___spec__3___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertTemplate___spec__4___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertTemplate___spec__5___closed__4; +lean_object* l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2___rarg(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_MatchEqs_0__Lean_Meta_Match_substSomeVar___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_FindImpl_findUnsafe_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertCastEqRec(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -503,6 +501,7 @@ static lean_object* l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSp LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_injectionAny___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertTemplate___lambda__3___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Match_unfoldNamedPattern___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__8; LEAN_EXPORT lean_object* l_Lean_Meta_Match_forallAltTelescope_isNamedPatternProof___lambda__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Match_proveCondEqThm_go___lambda__2___closed__7; static lean_object* l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertTemplate___lambda__2___closed__1; @@ -524,6 +523,7 @@ static lean_object* l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_simp static lean_object* l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertTemplate___lambda__3___closed__4; lean_object* l_Lean_MVarId_isAssigned___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_injectionAny___spec__4___lambda__1___closed__1; +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertCastEqRec___spec__8___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*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertCastEqRec___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_List_mapTR_loop___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertTemplate___spec__3___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertTemplate___spec__4___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertTemplate___spec__5___closed__6; @@ -531,6 +531,7 @@ LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Meta_Match_MatchE lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_proveSubgoal___spec__1(size_t, size_t, lean_object*); lean_object* l_Lean_Meta_mkEqRefl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__2; lean_object* l_Lean_indentExpr(lean_object*); LEAN_EXPORT lean_object* lean_get_match_equations_for(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Array_contains___at___private_Lean_Meta_RecursorInfo_0__Lean_Meta_getMajorPosDepElim___spec__4(lean_object*, lean_object*); @@ -588,6 +589,7 @@ static lean_object* l_Lean_Meta_Match_proveCondEqThm___closed__4; LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_substSomeVar___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Match_proveCondEqThm_go___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit_visitLet___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertTemplate___spec__11___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_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__5; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof___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_Array_mapMUnsafe_map___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertTemplate___spec__15(lean_object*, 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_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertTemplate___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*); @@ -597,7 +599,6 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Match_M static lean_object* l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkEquationsFor___lambda__2___closed__1; static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertCastEqRec___spec__6___lambda__2___closed__5; static lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertCastEqRec___spec__8___lambda__10___closed__2; -static lean_object* l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__14; lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); lean_object* l_Array_ofSubarray___rarg(lean_object*); uint8_t l_Lean_LocalDecl_binderInfo(lean_object*); @@ -649,12 +650,9 @@ lean_object* l_List_redLength___rarg(lean_object*); 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*); static lean_object* l_Lean_Meta_Match_proveCondEqThm_go___closed__2; lean_object* l_Lean_Meta_setInlineAttribute(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Match_forallAltTelescope_go___rarg___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*); lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__15; static lean_object* l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_simpH_x3f___closed__2; -static lean_object* l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__9; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertCastEqRec___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*); extern lean_object* l_Lean_Meta_Match_instInhabitedMatchEqnsExtState; LEAN_EXPORT lean_object* l_Lean_Meta_Match_proveCondEqThm_go___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -711,6 +709,7 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Match_MatchE LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertCastEqRec___spec__8___lambda__8___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_Expr_isConst(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkEquationsFor___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*); +static lean_object* l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__16; static lean_object* l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof___lambda__2___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertTemplate___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_Expr_withAppAux___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertTemplate___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*, lean_object*, lean_object*, lean_object*); @@ -744,6 +743,7 @@ LEAN_EXPORT lean_object* l_ReaderT_bind___at___private_Lean_Meta_Match_MatchEqs_ static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_injectionAny___spec__4___closed__3; LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_mkSplitterProof_convertCastEqRec___spec__8___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_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Match_getEquationsForImpl___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__6; 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___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_injectionAnyCandidate_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Nat_repr(lean_object*); @@ -7529,74 +7529,89 @@ return x_205; } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_SimpH_processNextEq___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_Meta_Match_MatchEqs_0__Lean_Meta_Match_SimpH_processNextEq___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) { _start: { -uint8_t x_11; -x_11 = l_Lean_Expr_isFVar(x_1); -if (x_11 == 0) +uint8_t x_12; +x_12 = l_Lean_Expr_isFVar(x_1); +if (x_12 == 0) { -lean_object* x_12; lean_object* x_13; -lean_dec(x_3); +lean_object* x_13; lean_object* x_14; +lean_dec(x_4); lean_dec(x_1); -x_12 = lean_box(0); -x_13 = lean_apply_7(x_2, x_12, x_5, x_6, x_7, x_8, x_9, x_10); -return x_13; +x_13 = lean_box(0); +x_14 = lean_apply_7(x_2, x_13, x_6, x_7, x_8, x_9, x_10, x_11); +return x_14; } else { -lean_object* x_14; lean_object* x_15; +lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_15 = lean_ctor_get(x_3, 1); +x_16 = l_Lean_Expr_fvarId_x21(x_1); +x_17 = l_List_elem___at_Lean_MVarId_getNondepPropHyps___spec__3(x_16, x_15); +if (x_17 == 0) +{ +lean_object* x_18; lean_object* x_19; +lean_dec(x_16); +lean_dec(x_4); +x_18 = lean_box(0); +x_19 = lean_apply_7(x_2, x_18, x_6, x_7, x_8, x_9, x_10, x_11); +return x_19; +} +else +{ +lean_object* x_20; lean_dec(x_2); -x_14 = l_Lean_Expr_fvarId_x21(x_1); -x_15 = l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_SimpH_substRHS(x_3, x_14, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_15) == 0) +x_20 = l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_SimpH_substRHS(x_4, x_16, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_20) == 0) { -uint8_t x_16; -x_16 = !lean_is_exclusive(x_15); -if (x_16 == 0) +uint8_t x_21; +x_21 = !lean_is_exclusive(x_20); +if (x_21 == 0) { -lean_object* x_17; uint8_t x_18; lean_object* x_19; -x_17 = lean_ctor_get(x_15, 0); -lean_dec(x_17); -x_18 = 1; -x_19 = lean_box(x_18); -lean_ctor_set(x_15, 0, x_19); -return x_15; +lean_object* x_22; uint8_t x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_20, 0); +lean_dec(x_22); +x_23 = 1; +x_24 = lean_box(x_23); +lean_ctor_set(x_20, 0, x_24); +return x_20; } else { -lean_object* x_20; uint8_t x_21; lean_object* x_22; lean_object* x_23; -x_20 = lean_ctor_get(x_15, 1); -lean_inc(x_20); -lean_dec(x_15); -x_21 = 1; -x_22 = lean_box(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_20); -return x_23; +lean_object* x_25; uint8_t x_26; lean_object* x_27; lean_object* x_28; +x_25 = lean_ctor_get(x_20, 1); +lean_inc(x_25); +lean_dec(x_20); +x_26 = 1; +x_27 = lean_box(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_25); +return x_28; } } else { -uint8_t x_24; -x_24 = !lean_is_exclusive(x_15); -if (x_24 == 0) +uint8_t x_29; +x_29 = !lean_is_exclusive(x_20); +if (x_29 == 0) { -return x_15; +return x_20; } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_25 = lean_ctor_get(x_15, 0); -x_26 = lean_ctor_get(x_15, 1); -lean_inc(x_26); -lean_inc(x_25); -lean_dec(x_15); -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_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_20, 0); +x_31 = lean_ctor_get(x_20, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_20); +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; +} } } } @@ -7618,11 +7633,11 @@ lean_dec(x_3); x_10 = l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_SimpH_processNextEq___lambda__5___closed__1; x_11 = lean_ctor_get(x_1, 2); lean_inc(x_11); -lean_dec(x_1); if (lean_obj_tag(x_11) == 0) { lean_object* x_12; lean_object* x_13; lean_dec(x_2); +lean_dec(x_1); x_12 = lean_box(0); x_13 = lean_apply_7(x_10, x_12, x_4, x_5, x_6, x_7, x_8, x_9); return x_13; @@ -7681,6 +7696,7 @@ lean_inc(x_28); if (lean_obj_tag(x_28) == 0) { lean_object* x_29; lean_object* x_30; lean_object* x_31; +lean_dec(x_1); x_29 = lean_ctor_get(x_27, 1); lean_inc(x_29); lean_dec(x_27); @@ -7731,7 +7747,8 @@ x_41 = lean_ctor_get(x_38, 1); lean_inc(x_41); lean_dec(x_38); x_42 = lean_box(0); -x_43 = l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_SimpH_processNextEq___lambda__4(x_37, x_34, x_14, x_42, x_4, x_5, x_6, x_7, x_8, x_41); +x_43 = l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_SimpH_processNextEq___lambda__4(x_37, x_34, x_1, x_14, x_42, x_4, x_5, x_6, x_7, x_8, x_41); +lean_dec(x_1); return x_43; } else @@ -7745,6 +7762,7 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); +lean_dec(x_1); x_44 = !lean_is_exclusive(x_38); if (x_44 == 0) { @@ -7782,6 +7800,7 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); +lean_dec(x_1); x_52 = !lean_is_exclusive(x_38); if (x_52 == 0) { @@ -7814,6 +7833,7 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); +lean_dec(x_1); x_56 = !lean_is_exclusive(x_27); if (x_56 == 0) { @@ -7844,6 +7864,7 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); +lean_dec(x_1); x_60 = !lean_is_exclusive(x_24); if (x_60 == 0) { @@ -7912,6 +7933,7 @@ lean_inc(x_75); if (lean_obj_tag(x_75) == 0) { lean_object* x_76; lean_object* x_77; lean_object* x_78; +lean_dec(x_1); x_76 = lean_ctor_get(x_74, 1); lean_inc(x_76); lean_dec(x_74); @@ -7962,7 +7984,8 @@ x_88 = lean_ctor_get(x_85, 1); lean_inc(x_88); lean_dec(x_85); x_89 = lean_box(0); -x_90 = l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_SimpH_processNextEq___lambda__4(x_84, x_81, x_14, x_89, x_4, x_5, x_6, x_7, x_8, x_88); +x_90 = l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_SimpH_processNextEq___lambda__4(x_84, x_81, x_1, x_14, x_89, x_4, x_5, x_6, x_7, x_8, x_88); +lean_dec(x_1); return x_90; } else @@ -7976,6 +7999,7 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); +lean_dec(x_1); x_91 = lean_ctor_get(x_85, 1); lean_inc(x_91); if (lean_is_exclusive(x_85)) { @@ -8009,6 +8033,7 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); +lean_dec(x_1); x_96 = lean_ctor_get(x_85, 0); lean_inc(x_96); x_97 = lean_ctor_get(x_85, 1); @@ -8043,6 +8068,7 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); +lean_dec(x_1); x_100 = lean_ctor_get(x_74, 0); lean_inc(x_100); x_101 = lean_ctor_get(x_74, 1); @@ -8075,6 +8101,7 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); +lean_dec(x_1); x_104 = lean_ctor_get(x_71, 0); lean_inc(x_104); x_105 = lean_ctor_get(x_71, 1); @@ -8180,13 +8207,14 @@ lean_dec(x_4); return x_11; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_SimpH_processNextEq___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_EXPORT lean_object* l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_SimpH_processNextEq___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) { _start: { -lean_object* x_11; -x_11 = l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_SimpH_processNextEq___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_4); -return x_11; +lean_object* x_12; +x_12 = l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_SimpH_processNextEq___lambda__4(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_5); +lean_dec(x_3); +return x_12; } } LEAN_EXPORT lean_object* l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_SimpH_processNextEq___lambda__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) { @@ -12361,7 +12389,7 @@ lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; uint x_26 = lean_ctor_get(x_25, 1); lean_inc(x_26); lean_dec(x_25); -x_27 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_1, x_5, x_6, x_7, x_8, x_26); +x_27 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_1, x_5, x_6, x_7, x_8, x_26); x_28 = lean_ctor_get(x_27, 0); lean_inc(x_28); x_29 = lean_ctor_get(x_27, 1); @@ -12455,7 +12483,7 @@ lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; uint x_55 = lean_ctor_get(x_54, 1); lean_inc(x_55); lean_dec(x_54); -x_56 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_1, x_5, x_6, x_7, x_8, x_55); +x_56 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_1, x_5, x_6, x_7, x_8, x_55); x_57 = lean_ctor_get(x_56, 0); lean_inc(x_57); x_58 = lean_ctor_get(x_56, 1); @@ -12613,7 +12641,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Match_proveCondEqThm___lambda__4(lean_objec _start: { lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_8 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_1, x_3, x_4, x_5, x_6, x_7); +x_8 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_1, x_3, x_4, x_5, x_6, x_7); x_9 = lean_ctor_get(x_8, 0); lean_inc(x_9); x_10 = lean_ctor_get(x_8, 1); @@ -12712,7 +12740,7 @@ lean_closure_set(x_8, 0, x_2); lean_closure_set(x_8, 1, x_1); x_9 = l_Lean_Meta_Match_proveCondEqThm___closed__7; x_10 = l_Lean_Meta_casesOnStuckLHS___closed__2; -x_11 = l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__3___rarg(x_9, x_10, x_8, x_3, x_4, x_5, x_6, x_7); +x_11 = l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2___rarg(x_9, x_10, x_8, x_3, x_4, x_5, x_6, x_7); return x_11; } } @@ -23933,7 +23961,7 @@ static lean_object* _init_l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Matc _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("( __do_lift._@.Lean.Meta.Match.MatchEqs._hyg.7542.0 )\n ", 58); +x_1 = lean_mk_string_from_bytes("( __do_lift._@.Lean.Meta.Match.MatchEqs._hyg.7550.0 )\n ", 58); return x_1; } } @@ -29082,7 +29110,7 @@ lean_object* x_16; lean_object* x_17; x_16 = lean_ctor_get(x_15, 1); lean_inc(x_16); lean_dec(x_15); -x_17 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_4, x_6, x_7, x_8, x_9, x_16); +x_17 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_4, x_6, x_7, x_8, x_9, x_16); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -32653,7 +32681,7 @@ lean_closure_set(x_9, 0, x_7); lean_closure_set(x_9, 1, x_1); lean_closure_set(x_9, 2, x_8); x_10 = l_Lean_Meta_Match_proveCondEqThm___closed__7; -x_11 = l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__3___rarg(x_10, x_8, x_9, x_2, x_3, x_4, x_5, x_6); +x_11 = l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2___rarg(x_10, x_8, x_9, x_2, x_3, x_4, x_5, x_6); return x_11; } } @@ -33238,7 +33266,7 @@ lean_dec(x_3); return x_5; } } -static lean_object* _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__1() { +static lean_object* _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__1() { _start: { lean_object* x_1; @@ -33246,37 +33274,37 @@ x_1 = lean_mk_string_from_bytes("Lean", 4); return x_1; } } -static lean_object* _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__2() { +static lean_object* _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__1; +x_2 = l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__3() { +static lean_object* _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__2; +x_1 = l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__2; x_2 = l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_simpH_x3f___lambda__3___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__4() { +static lean_object* _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__3; +x_1 = l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__3; x_2 = l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_simpH_x3f___lambda__3___closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__5() { +static lean_object* _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__5() { _start: { lean_object* x_1; @@ -33284,17 +33312,17 @@ x_1 = lean_mk_string_from_bytes("initFn", 6); return x_1; } } -static lean_object* _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__6() { +static lean_object* _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__4; -x_2 = l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__5; +x_1 = l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__4; +x_2 = l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__7() { +static lean_object* _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__7() { _start: { lean_object* x_1; @@ -33302,47 +33330,47 @@ x_1 = lean_mk_string_from_bytes("_@", 2); return x_1; } } -static lean_object* _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__8() { +static lean_object* _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__6; -x_2 = l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__7; +x_1 = l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__6; +x_2 = l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__7; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__9() { +static lean_object* _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__8; -x_2 = l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__1; +x_1 = l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__8; +x_2 = l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__10() { +static lean_object* _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__9; +x_1 = l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__9; x_2 = l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_simpH_x3f___lambda__3___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__11() { +static lean_object* _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__10; +x_1 = l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__10; x_2 = l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_simpH_x3f___lambda__3___closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__12() { +static lean_object* _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__12() { _start: { lean_object* x_1; @@ -33350,17 +33378,17 @@ x_1 = lean_mk_string_from_bytes("MatchEqs", 8); return x_1; } } -static lean_object* _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__13() { +static lean_object* _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__11; -x_2 = l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__12; +x_1 = l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__11; +x_2 = l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__12; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__14() { +static lean_object* _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__14() { _start: { lean_object* x_1; @@ -33368,33 +33396,33 @@ x_1 = lean_mk_string_from_bytes("_hyg", 4); return x_1; } } -static lean_object* _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__15() { +static lean_object* _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__13; -x_2 = l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__14; +x_1 = l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__13; +x_2 = l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__14; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__16() { +static lean_object* _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__15; -x_2 = lean_unsigned_to_nat(15480u); +x_1 = l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__15; +x_2 = lean_unsigned_to_nat(15488u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; x_2 = l___private_Lean_Meta_Match_MatchEqs_0__Lean_Meta_Match_simpH_x3f___lambda__3___closed__4; x_3 = 0; -x_4 = l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__16; +x_4 = l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__16; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } @@ -33836,39 +33864,39 @@ l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Match_getEquationsForImpl___spec l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Match_getEquationsForImpl___spec__2___closed__2 = _init_l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Match_getEquationsForImpl___spec__2___closed__2(); l_Lean_Meta_Match_getEquationsForImpl___closed__1 = _init_l_Lean_Meta_Match_getEquationsForImpl___closed__1(); lean_mark_persistent(l_Lean_Meta_Match_getEquationsForImpl___closed__1); -l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__1 = _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__1(); -lean_mark_persistent(l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__1); -l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__2 = _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__2(); -lean_mark_persistent(l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__2); -l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__3 = _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__3(); -lean_mark_persistent(l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__3); -l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__4 = _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__4(); -lean_mark_persistent(l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__4); -l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__5 = _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__5(); -lean_mark_persistent(l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__5); -l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__6 = _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__6(); -lean_mark_persistent(l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__6); -l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__7 = _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__7(); -lean_mark_persistent(l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__7); -l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__8 = _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__8(); -lean_mark_persistent(l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__8); -l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__9 = _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__9(); -lean_mark_persistent(l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__9); -l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__10 = _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__10(); -lean_mark_persistent(l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__10); -l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__11 = _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__11(); -lean_mark_persistent(l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__11); -l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__12 = _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__12(); -lean_mark_persistent(l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__12); -l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__13 = _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__13(); -lean_mark_persistent(l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__13); -l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__14 = _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__14(); -lean_mark_persistent(l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__14); -l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__15 = _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__15(); -lean_mark_persistent(l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__15); -l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__16 = _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__16(); -lean_mark_persistent(l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480____closed__16); -res = l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15480_(lean_io_mk_world()); +l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__1 = _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__1(); +lean_mark_persistent(l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__1); +l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__2 = _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__2(); +lean_mark_persistent(l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__2); +l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__3 = _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__3(); +lean_mark_persistent(l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__3); +l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__4 = _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__4(); +lean_mark_persistent(l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__4); +l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__5 = _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__5(); +lean_mark_persistent(l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__5); +l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__6 = _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__6(); +lean_mark_persistent(l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__6); +l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__7 = _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__7(); +lean_mark_persistent(l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__7); +l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__8 = _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__8(); +lean_mark_persistent(l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__8); +l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__9 = _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__9(); +lean_mark_persistent(l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__9); +l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__10 = _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__10(); +lean_mark_persistent(l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__10); +l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__11 = _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__11(); +lean_mark_persistent(l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__11); +l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__12 = _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__12(); +lean_mark_persistent(l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__12); +l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__13 = _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__13(); +lean_mark_persistent(l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__13); +l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__14 = _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__14(); +lean_mark_persistent(l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__14); +l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__15 = _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__15(); +lean_mark_persistent(l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__15); +l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__16 = _init_l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__16(); +lean_mark_persistent(l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488____closed__16); +res = l_Lean_Meta_Match_initFn____x40_Lean_Meta_Match_MatchEqs___hyg_15488_(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/Meta/Offset.c b/stage0/stdlib/Lean/Meta/Offset.c index d1bfd6171dff..63b50eadc82e 100644 --- a/stage0/stdlib/Lean/Meta/Offset.c +++ b/stage0/stdlib/Lean/Meta/Offset.c @@ -23,7 +23,6 @@ lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, l LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Offset_0__Lean_Meta_withInstantiatedMVars___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_isNatProjInst___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Meta_evalNat___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Offset_0__Lean_Meta_isOffset___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Offset_0__Lean_Meta_withInstantiatedMVars(lean_object*); static lean_object* l_Lean_Meta_evalNat_visit___closed__5; static lean_object* l_Lean_Meta_isDefEqOffset___closed__2; @@ -51,9 +50,9 @@ lean_object* l_Lean_Meta_isExprDefEqAux___boxed(lean_object*, lean_object*, lean LEAN_EXPORT lean_object* l___private_Lean_Meta_Offset_0__Lean_Meta_isNatZero(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_isNatProjInst___closed__4; lean_object* lean_st_ref_get(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Offset_0__Lean_Meta_getOffsetAux(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_isNatProjInst___closed__8; static lean_object* l_Lean_Meta_isNatProjInst___closed__19; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Offset_0__Lean_Meta_isOffset_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); static lean_object* l_Lean_Meta_evalNat_visit___closed__4; @@ -66,8 +65,8 @@ lean_object* l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Meta_WHNF_0__Lean_ static lean_object* l_Lean_Meta_isNatProjInst___closed__7; lean_object* l_Lean_Meta_getConfig(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_isNatProjInst___closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Meta_Offset_0__Lean_Meta_isOffset(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_isNatProjInst___closed__15; +LEAN_EXPORT lean_object* l___private_Lean_Meta_Offset_0__Lean_Meta_isOffset_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_isNatProjInst___closed__13; LEAN_EXPORT lean_object* l_Lean_Meta_evalNat_visit(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); @@ -91,7 +90,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_isDefEqOffset(lean_object*, lean_object*, l LEAN_EXPORT lean_object* l___private_Lean_Meta_Offset_0__Lean_Meta_withInstantiatedMVars___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_evalNat_visit___closed__2; lean_object* lean_nat_add(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Meta_Offset_0__Lean_Meta_getOffsetAux___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_isNatProjInst___closed__11; LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Offset_0__Lean_Meta_withInstantiatedMVars___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: @@ -2061,693 +2059,450 @@ lean_dec(x_2); return x_7; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Offset_0__Lean_Meta_getOffsetAux(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_EXPORT lean_object* l___private_Lean_Meta_Offset_0__Lean_Meta_getOffset(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: { -if (lean_obj_tag(x_1) == 5) +lean_object* x_7; +lean_inc(x_1); +x_7 = l___private_Lean_Meta_Offset_0__Lean_Meta_isOffset_x3f(x_1, x_2, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_7) == 0) { -lean_object* x_8; lean_object* x_9; -x_8 = lean_ctor_get(x_1, 1); +lean_object* x_8; +x_8 = lean_ctor_get(x_7, 0); lean_inc(x_8); -x_9 = l_Lean_Expr_getAppFn(x_1); -switch (lean_obj_tag(x_9)) { -case 2: -{ -lean_object* x_10; uint8_t x_11; -lean_dec(x_9); -lean_dec(x_8); -x_10 = l_Lean_instantiateMVars___at___private_Lean_Meta_Offset_0__Lean_Meta_withInstantiatedMVars___spec__1(x_1, x_3, x_4, x_5, x_6, x_7); -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; uint8_t x_16; -x_12 = lean_ctor_get(x_10, 0); -x_13 = lean_ctor_get(x_10, 1); -x_14 = lean_ctor_get(x_12, 0); -lean_inc(x_14); -lean_dec(x_12); -x_15 = l_Lean_Expr_getAppFn(x_14); -x_16 = l_Lean_Expr_isMVar(x_15); -lean_dec(x_15); -if (x_16 == 0) -{ -lean_free_object(x_10); -x_1 = x_14; -x_7 = x_13; -goto _start; -} -else +if (lean_obj_tag(x_8) == 0) { -lean_object* x_18; -lean_dec(x_14); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_18 = lean_box(0); -lean_ctor_set(x_10, 0, x_18); -return x_10; -} -} -else +uint8_t x_9; +x_9 = !lean_is_exclusive(x_7); +if (x_9 == 0) { -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; -x_19 = lean_ctor_get(x_10, 0); -x_20 = lean_ctor_get(x_10, 1); -lean_inc(x_20); -lean_inc(x_19); +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = lean_ctor_get(x_7, 0); lean_dec(x_10); -x_21 = lean_ctor_get(x_19, 0); -lean_inc(x_21); -lean_dec(x_19); -x_22 = l_Lean_Expr_getAppFn(x_21); -x_23 = l_Lean_Expr_isMVar(x_22); -lean_dec(x_22); -if (x_23 == 0) -{ -x_1 = x_21; -x_7 = x_20; -goto _start; +x_11 = lean_unsigned_to_nat(0u); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_1); +lean_ctor_set(x_12, 1, x_11); +lean_ctor_set(x_7, 0, x_12); +return x_7; } else { -lean_object* x_25; lean_object* x_26; -lean_dec(x_21); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_25 = lean_box(0); -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; -} -} -} -case 4: -{ -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_62; lean_object* x_175; uint8_t x_176; -x_27 = lean_ctor_get(x_9, 0); -lean_inc(x_27); -lean_dec(x_9); -x_28 = lean_unsigned_to_nat(0u); -x_29 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_28); -x_175 = l_Lean_Meta_evalNat_visit___closed__6; -x_176 = lean_name_eq(x_27, x_175); -if (x_176 == 0) -{ -lean_object* x_177; -lean_dec(x_8); -x_177 = lean_box(0); -x_62 = x_177; -goto block_174; +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_13 = lean_ctor_get(x_7, 1); +lean_inc(x_13); +lean_dec(x_7); +x_14 = lean_unsigned_to_nat(0u); +x_15 = lean_alloc_ctor(0, 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_13); +return x_16; } -else -{ -lean_object* x_178; uint8_t x_179; -x_178 = lean_unsigned_to_nat(1u); -x_179 = lean_nat_dec_eq(x_29, x_178); -if (x_179 == 0) -{ -lean_object* x_180; -lean_dec(x_8); -x_180 = lean_box(0); -x_62 = x_180; -goto block_174; } else { -uint8_t x_181; lean_object* x_182; -lean_dec(x_29); -lean_dec(x_27); +uint8_t x_17; lean_dec(x_1); -x_181 = 0; -x_182 = l___private_Lean_Meta_Offset_0__Lean_Meta_getOffsetAux(x_8, x_181, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_182) == 0) -{ -lean_object* x_183; -x_183 = lean_ctor_get(x_182, 0); -lean_inc(x_183); -if (lean_obj_tag(x_183) == 0) -{ -uint8_t x_184; -x_184 = !lean_is_exclusive(x_182); -if (x_184 == 0) +x_17 = !lean_is_exclusive(x_7); +if (x_17 == 0) { -lean_object* x_185; lean_object* x_186; -x_185 = lean_ctor_get(x_182, 0); -lean_dec(x_185); -x_186 = lean_box(0); -lean_ctor_set(x_182, 0, x_186); -return x_182; +lean_object* x_18; lean_object* x_19; +x_18 = lean_ctor_get(x_7, 0); +lean_dec(x_18); +x_19 = lean_ctor_get(x_8, 0); +lean_inc(x_19); +lean_dec(x_8); +lean_ctor_set(x_7, 0, x_19); +return x_7; } else { -lean_object* x_187; lean_object* x_188; lean_object* x_189; -x_187 = lean_ctor_get(x_182, 1); -lean_inc(x_187); -lean_dec(x_182); -x_188 = lean_box(0); -x_189 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_189, 0, x_188); -lean_ctor_set(x_189, 1, x_187); -return x_189; +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_7, 1); +lean_inc(x_20); +lean_dec(x_7); +x_21 = lean_ctor_get(x_8, 0); +lean_inc(x_21); +lean_dec(x_8); +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 -{ -uint8_t x_190; -x_190 = !lean_is_exclusive(x_183); -if (x_190 == 0) -{ -uint8_t x_191; -x_191 = !lean_is_exclusive(x_182); -if (x_191 == 0) -{ -lean_object* x_192; lean_object* x_193; uint8_t x_194; -x_192 = lean_ctor_get(x_183, 0); -x_193 = lean_ctor_get(x_182, 0); -lean_dec(x_193); -x_194 = !lean_is_exclusive(x_192); -if (x_194 == 0) -{ -lean_object* x_195; lean_object* x_196; -x_195 = lean_ctor_get(x_192, 1); -x_196 = lean_nat_add(x_195, x_178); -lean_dec(x_195); -lean_ctor_set(x_192, 1, x_196); -return x_182; } else { -lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; -x_197 = lean_ctor_get(x_192, 0); -x_198 = lean_ctor_get(x_192, 1); -lean_inc(x_198); -lean_inc(x_197); -lean_dec(x_192); -x_199 = lean_nat_add(x_198, x_178); -lean_dec(x_198); -x_200 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_200, 0, x_197); -lean_ctor_set(x_200, 1, x_199); -lean_ctor_set(x_183, 0, x_200); -return x_182; -} -} -else +uint8_t x_23; +lean_dec(x_1); +x_23 = !lean_is_exclusive(x_7); +if (x_23 == 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; -x_201 = lean_ctor_get(x_183, 0); -x_202 = lean_ctor_get(x_182, 1); -lean_inc(x_202); -lean_dec(x_182); -x_203 = lean_ctor_get(x_201, 0); -lean_inc(x_203); -x_204 = lean_ctor_get(x_201, 1); -lean_inc(x_204); -if (lean_is_exclusive(x_201)) { - lean_ctor_release(x_201, 0); - lean_ctor_release(x_201, 1); - x_205 = x_201; -} else { - lean_dec_ref(x_201); - x_205 = lean_box(0); -} -x_206 = lean_nat_add(x_204, x_178); -lean_dec(x_204); -if (lean_is_scalar(x_205)) { - x_207 = lean_alloc_ctor(0, 2, 0); -} else { - x_207 = x_205; -} -lean_ctor_set(x_207, 0, x_203); -lean_ctor_set(x_207, 1, x_206); -lean_ctor_set(x_183, 0, x_207); -x_208 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_208, 0, x_183); -lean_ctor_set(x_208, 1, x_202); -return x_208; -} +return x_7; } 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; -x_209 = lean_ctor_get(x_183, 0); -lean_inc(x_209); -lean_dec(x_183); -x_210 = lean_ctor_get(x_182, 1); -lean_inc(x_210); -if (lean_is_exclusive(x_182)) { - lean_ctor_release(x_182, 0); - lean_ctor_release(x_182, 1); - x_211 = x_182; -} else { - lean_dec_ref(x_182); - x_211 = lean_box(0); -} -x_212 = lean_ctor_get(x_209, 0); -lean_inc(x_212); -x_213 = lean_ctor_get(x_209, 1); -lean_inc(x_213); -if (lean_is_exclusive(x_209)) { - lean_ctor_release(x_209, 0); - lean_ctor_release(x_209, 1); - x_214 = x_209; -} else { - lean_dec_ref(x_209); - x_214 = lean_box(0); -} -x_215 = lean_nat_add(x_213, x_178); -lean_dec(x_213); -if (lean_is_scalar(x_214)) { - x_216 = lean_alloc_ctor(0, 2, 0); -} else { - x_216 = x_214; -} -lean_ctor_set(x_216, 0, x_212); -lean_ctor_set(x_216, 1, x_215); -x_217 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_217, 0, x_216); -if (lean_is_scalar(x_211)) { - x_218 = lean_alloc_ctor(0, 2, 0); -} else { - x_218 = x_211; -} -lean_ctor_set(x_218, 0, x_217); -lean_ctor_set(x_218, 1, x_210); -return x_218; +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_7, 0); +x_25 = lean_ctor_get(x_7, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_7); +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_219; -x_219 = !lean_is_exclusive(x_182); -if (x_219 == 0) -{ -return x_182; } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_Offset_0__Lean_Meta_isOffset_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) { +_start: { -lean_object* x_220; lean_object* x_221; lean_object* x_222; -x_220 = lean_ctor_get(x_182, 0); -x_221 = lean_ctor_get(x_182, 1); -lean_inc(x_221); -lean_inc(x_220); -lean_dec(x_182); -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; -} -} -} -} -block_61: +if (lean_obj_tag(x_1) == 5) { -lean_object* x_31; uint8_t x_32; -lean_dec(x_30); -x_31 = l_Lean_Meta_isNatProjInst___closed__6; -x_32 = lean_name_eq(x_27, x_31); -lean_dec(x_27); -if (x_32 == 0) +lean_object* x_7; lean_object* x_8; +x_7 = lean_ctor_get(x_1, 1); +lean_inc(x_7); +x_8 = l_Lean_Expr_getAppFn(x_1); +switch (lean_obj_tag(x_8)) { +case 2: { -lean_dec(x_29); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -if (x_2 == 0) +lean_object* x_9; uint8_t x_10; +lean_dec(x_8); +lean_dec(x_7); +x_9 = l_Lean_instantiateMVars___at___private_Lean_Meta_Offset_0__Lean_Meta_withInstantiatedMVars___spec__1(x_1, x_2, x_3, x_4, x_5, x_6); +x_10 = !lean_is_exclusive(x_9); +if (x_10 == 0) { -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_33, 0, x_1); -lean_ctor_set(x_33, 1, x_28); -x_34 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_34, 0, 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_7); -return x_35; -} -else +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_9, 0); +x_12 = lean_ctor_get(x_9, 1); +x_13 = lean_ctor_get(x_11, 0); +lean_inc(x_13); +lean_dec(x_11); +x_14 = l_Lean_Expr_getAppFn(x_13); +x_15 = l_Lean_Expr_isMVar(x_14); +lean_dec(x_14); +if (x_15 == 0) { -lean_object* x_36; lean_object* x_37; -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_7); -return x_37; -} +lean_free_object(x_9); +x_1 = x_13; +x_6 = x_12; +goto _start; } else { -lean_object* x_38; uint8_t x_39; -x_38 = lean_unsigned_to_nat(6u); -x_39 = lean_nat_dec_eq(x_29, x_38); -lean_dec(x_29); -if (x_39 == 0) -{ -lean_dec(x_6); +lean_object* x_17; +lean_dec(x_13); lean_dec(x_5); lean_dec(x_4); -if (x_2 == 0) -{ -lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_1); -lean_ctor_set(x_40, 1, x_28); -x_41 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_41, 0, 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_7); -return x_42; -} -else -{ -lean_object* x_43; lean_object* x_44; -lean_dec(x_1); -x_43 = lean_box(0); -x_44 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_44, 0, x_43); -lean_ctor_set(x_44, 1, x_7); -return x_44; +lean_dec(x_3); +x_17 = lean_box(0); +lean_ctor_set(x_9, 0, x_17); +return x_9; } } else { -lean_object* x_45; -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_45 = l_Lean_Meta_unfoldProjInst_x3f(x_1, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_45) == 0) -{ -lean_object* x_46; -x_46 = lean_ctor_get(x_45, 0); -lean_inc(x_46); -if (lean_obj_tag(x_46) == 0) -{ -uint8_t x_47; -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_47 = !lean_is_exclusive(x_45); -if (x_47 == 0) -{ -lean_object* x_48; lean_object* x_49; -x_48 = lean_ctor_get(x_45, 0); -lean_dec(x_48); -x_49 = lean_box(0); -lean_ctor_set(x_45, 0, x_49); -return x_45; -} -else -{ -lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_50 = lean_ctor_get(x_45, 1); -lean_inc(x_50); -lean_dec(x_45); -x_51 = lean_box(0); -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 +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_18 = lean_ctor_get(x_9, 0); +x_19 = lean_ctor_get(x_9, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_9); +x_20 = lean_ctor_get(x_18, 0); +lean_inc(x_20); +lean_dec(x_18); +x_21 = l_Lean_Expr_getAppFn(x_20); +x_22 = l_Lean_Expr_isMVar(x_21); +lean_dec(x_21); +if (x_22 == 0) { -lean_object* x_53; lean_object* x_54; uint8_t x_55; -x_53 = lean_ctor_get(x_45, 1); -lean_inc(x_53); -lean_dec(x_45); -x_54 = lean_ctor_get(x_46, 0); -lean_inc(x_54); -lean_dec(x_46); -x_55 = 0; -x_1 = x_54; -x_2 = x_55; -x_7 = x_53; +x_1 = x_20; +x_6 = x_19; goto _start; } -} else { -uint8_t x_57; -lean_dec(x_6); +lean_object* x_24; lean_object* x_25; +lean_dec(x_20); lean_dec(x_5); lean_dec(x_4); -x_57 = !lean_is_exclusive(x_45); -if (x_57 == 0) -{ -return x_45; -} -else -{ -lean_object* x_58; lean_object* x_59; lean_object* x_60; -x_58 = lean_ctor_get(x_45, 0); -x_59 = lean_ctor_get(x_45, 1); -lean_inc(x_59); -lean_inc(x_58); -lean_dec(x_45); -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_dec(x_3); +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_19); +return x_25; } } } -block_174: -{ -lean_object* x_63; uint8_t x_64; -lean_dec(x_62); -x_63 = l_Lean_Meta_evalNat_visit___closed__3; -x_64 = lean_name_eq(x_27, x_63); -if (x_64 == 0) +case 4: { -lean_object* x_65; uint8_t x_66; -x_65 = l_Lean_Meta_isNatProjInst___closed__15; -x_66 = lean_name_eq(x_27, x_65); -if (x_66 == 0) +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_96; lean_object* x_167; uint8_t x_168; +x_26 = lean_ctor_get(x_8, 0); +lean_inc(x_26); +lean_dec(x_8); +x_27 = lean_unsigned_to_nat(0u); +x_28 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_27); +x_167 = l_Lean_Meta_evalNat_visit___closed__6; +x_168 = lean_name_eq(x_26, x_167); +if (x_168 == 0) { -lean_object* x_67; -x_67 = lean_box(0); -x_30 = x_67; -goto block_61; +lean_object* x_169; +lean_dec(x_7); +x_169 = lean_box(0); +x_96 = x_169; +goto block_166; } else { -lean_object* x_68; uint8_t x_69; -x_68 = lean_unsigned_to_nat(4u); -x_69 = lean_nat_dec_eq(x_29, x_68); -if (x_69 == 0) +lean_object* x_170; uint8_t x_171; +x_170 = lean_unsigned_to_nat(1u); +x_171 = lean_nat_dec_eq(x_28, x_170); +if (x_171 == 0) { -lean_object* x_70; -x_70 = lean_box(0); -x_30 = x_70; -goto block_61; +lean_object* x_172; +lean_dec(x_7); +x_172 = lean_box(0); +x_96 = x_172; +goto block_166; } else { -lean_object* x_71; -lean_dec(x_29); -lean_dec(x_27); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_71 = l_Lean_Meta_unfoldProjInst_x3f(x_1, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_71) == 0) +lean_object* x_173; +lean_dec(x_28); +lean_dec(x_26); +lean_dec(x_1); +x_173 = l___private_Lean_Meta_Offset_0__Lean_Meta_getOffset(x_7, x_2, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_173) == 0) { -lean_object* x_72; -x_72 = lean_ctor_get(x_71, 0); -lean_inc(x_72); -if (lean_obj_tag(x_72) == 0) +uint8_t x_174; +x_174 = !lean_is_exclusive(x_173); +if (x_174 == 0) { -uint8_t x_73; -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_73 = !lean_is_exclusive(x_71); -if (x_73 == 0) +lean_object* x_175; uint8_t x_176; +x_175 = lean_ctor_get(x_173, 0); +x_176 = !lean_is_exclusive(x_175); +if (x_176 == 0) { -lean_object* x_74; lean_object* x_75; -x_74 = lean_ctor_get(x_71, 0); -lean_dec(x_74); -x_75 = lean_box(0); -lean_ctor_set(x_71, 0, x_75); -return x_71; +lean_object* x_177; lean_object* x_178; lean_object* x_179; +x_177 = lean_ctor_get(x_175, 1); +x_178 = lean_nat_add(x_177, x_170); +lean_dec(x_177); +lean_ctor_set(x_175, 1, x_178); +x_179 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_179, 0, x_175); +lean_ctor_set(x_173, 0, x_179); +return x_173; } else { -lean_object* x_76; lean_object* x_77; lean_object* x_78; -x_76 = lean_ctor_get(x_71, 1); -lean_inc(x_76); -lean_dec(x_71); -x_77 = lean_box(0); -x_78 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_78, 0, x_77); -lean_ctor_set(x_78, 1, x_76); -return x_78; +lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; +x_180 = lean_ctor_get(x_175, 0); +x_181 = lean_ctor_get(x_175, 1); +lean_inc(x_181); +lean_inc(x_180); +lean_dec(x_175); +x_182 = lean_nat_add(x_181, x_170); +lean_dec(x_181); +x_183 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_183, 0, x_180); +lean_ctor_set(x_183, 1, x_182); +x_184 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_184, 0, x_183); +lean_ctor_set(x_173, 0, x_184); +return x_173; } } else { -lean_object* x_79; lean_object* x_80; uint8_t x_81; -x_79 = lean_ctor_get(x_71, 1); -lean_inc(x_79); -lean_dec(x_71); -x_80 = lean_ctor_get(x_72, 0); -lean_inc(x_80); -lean_dec(x_72); -x_81 = 0; -x_1 = x_80; -x_2 = x_81; -x_7 = x_79; -goto _start; +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_ctor_get(x_173, 0); +x_186 = lean_ctor_get(x_173, 1); +lean_inc(x_186); +lean_inc(x_185); +lean_dec(x_173); +x_187 = lean_ctor_get(x_185, 0); +lean_inc(x_187); +x_188 = lean_ctor_get(x_185, 1); +lean_inc(x_188); +if (lean_is_exclusive(x_185)) { + lean_ctor_release(x_185, 0); + lean_ctor_release(x_185, 1); + x_189 = x_185; +} else { + lean_dec_ref(x_185); + x_189 = lean_box(0); +} +x_190 = lean_nat_add(x_188, x_170); +lean_dec(x_188); +if (lean_is_scalar(x_189)) { + x_191 = lean_alloc_ctor(0, 2, 0); +} else { + x_191 = x_189; +} +lean_ctor_set(x_191, 0, x_187); +lean_ctor_set(x_191, 1, x_190); +x_192 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_192, 0, 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_186); +return x_193; } } else { -uint8_t x_83; -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_83 = !lean_is_exclusive(x_71); -if (x_83 == 0) +uint8_t x_194; +x_194 = !lean_is_exclusive(x_173); +if (x_194 == 0) { -return x_71; +return x_173; } else { -lean_object* x_84; lean_object* x_85; lean_object* x_86; -x_84 = lean_ctor_get(x_71, 0); -x_85 = lean_ctor_get(x_71, 1); -lean_inc(x_85); -lean_inc(x_84); -lean_dec(x_71); -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_object* x_195; lean_object* x_196; lean_object* x_197; +x_195 = lean_ctor_get(x_173, 0); +x_196 = lean_ctor_get(x_173, 1); +lean_inc(x_196); +lean_inc(x_195); +lean_dec(x_173); +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 +block_95: { -lean_object* x_87; uint8_t x_88; -x_87 = lean_unsigned_to_nat(2u); -x_88 = lean_nat_dec_eq(x_29, x_87); -if (x_88 == 0) +lean_object* x_30; uint8_t x_31; +lean_dec(x_29); +x_30 = l_Lean_Meta_isNatProjInst___closed__15; +x_31 = lean_name_eq(x_26, x_30); +if (x_31 == 0) { -lean_object* x_89; uint8_t x_90; -x_89 = l_Lean_Meta_isNatProjInst___closed__15; -x_90 = lean_name_eq(x_27, x_89); -if (x_90 == 0) +lean_object* x_32; uint8_t x_33; +x_32 = l_Lean_Meta_isNatProjInst___closed__6; +x_33 = lean_name_eq(x_26, x_32); +lean_dec(x_26); +if (x_33 == 0) { -lean_object* x_91; -x_91 = lean_box(0); -x_30 = x_91; -goto block_61; +lean_object* x_34; lean_object* x_35; +lean_dec(x_28); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_34 = lean_box(0); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_6); +return x_35; } else { -lean_object* x_92; uint8_t x_93; -x_92 = lean_unsigned_to_nat(4u); -x_93 = lean_nat_dec_eq(x_29, x_92); -if (x_93 == 0) +lean_object* x_36; uint8_t x_37; +x_36 = lean_unsigned_to_nat(6u); +x_37 = lean_nat_dec_eq(x_28, x_36); +lean_dec(x_28); +if (x_37 == 0) { -lean_object* x_94; -x_94 = lean_box(0); -x_30 = x_94; -goto block_61; +lean_object* x_38; lean_object* x_39; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_38 = lean_box(0); +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_6); +return x_39; } else { -lean_object* x_95; -lean_dec(x_29); -lean_dec(x_27); -lean_inc(x_6); +lean_object* x_40; lean_inc(x_5); lean_inc(x_4); -x_95 = l_Lean_Meta_unfoldProjInst_x3f(x_1, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_95) == 0) +lean_inc(x_3); +x_40 = l_Lean_Meta_unfoldProjInst_x3f(x_1, x_2, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_40) == 0) { -lean_object* x_96; -x_96 = lean_ctor_get(x_95, 0); -lean_inc(x_96); -if (lean_obj_tag(x_96) == 0) +lean_object* x_41; +x_41 = lean_ctor_get(x_40, 0); +lean_inc(x_41); +if (lean_obj_tag(x_41) == 0) { -uint8_t x_97; -lean_dec(x_6); +uint8_t x_42; lean_dec(x_5); lean_dec(x_4); -x_97 = !lean_is_exclusive(x_95); -if (x_97 == 0) +lean_dec(x_3); +x_42 = !lean_is_exclusive(x_40); +if (x_42 == 0) { -lean_object* x_98; lean_object* x_99; -x_98 = lean_ctor_get(x_95, 0); -lean_dec(x_98); -x_99 = lean_box(0); -lean_ctor_set(x_95, 0, x_99); -return x_95; +lean_object* x_43; lean_object* x_44; +x_43 = lean_ctor_get(x_40, 0); +lean_dec(x_43); +x_44 = lean_box(0); +lean_ctor_set(x_40, 0, x_44); +return x_40; } else { -lean_object* x_100; lean_object* x_101; lean_object* x_102; -x_100 = lean_ctor_get(x_95, 1); -lean_inc(x_100); -lean_dec(x_95); -x_101 = lean_box(0); -x_102 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_102, 0, x_101); -lean_ctor_set(x_102, 1, x_100); -return x_102; +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_40, 1); +lean_inc(x_45); +lean_dec(x_40); +x_46 = lean_box(0); +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_45); +return x_47; } } else { -lean_object* x_103; lean_object* x_104; uint8_t x_105; -x_103 = lean_ctor_get(x_95, 1); -lean_inc(x_103); -lean_dec(x_95); -x_104 = lean_ctor_get(x_96, 0); -lean_inc(x_104); -lean_dec(x_96); -x_105 = 0; -x_1 = x_104; -x_2 = x_105; -x_7 = x_103; +lean_object* x_48; lean_object* x_49; +x_48 = lean_ctor_get(x_40, 1); +lean_inc(x_48); +lean_dec(x_40); +x_49 = lean_ctor_get(x_41, 0); +lean_inc(x_49); +lean_dec(x_41); +x_1 = x_49; +x_6 = x_48; goto _start; } } else { -uint8_t x_107; -lean_dec(x_6); +uint8_t x_51; lean_dec(x_5); lean_dec(x_4); -x_107 = !lean_is_exclusive(x_95); -if (x_107 == 0) +lean_dec(x_3); +x_51 = !lean_is_exclusive(x_40); +if (x_51 == 0) { -return x_95; +return x_40; } else { -lean_object* x_108; lean_object* x_109; lean_object* x_110; -x_108 = lean_ctor_get(x_95, 0); -x_109 = lean_ctor_get(x_95, 1); -lean_inc(x_109); -lean_inc(x_108); -lean_dec(x_95); -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; +lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_52 = lean_ctor_get(x_40, 0); +x_53 = lean_ctor_get(x_40, 1); +lean_inc(x_53); +lean_inc(x_52); +lean_dec(x_40); +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; } } } @@ -2755,679 +2510,575 @@ 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_27); -x_111 = lean_unsigned_to_nat(1u); -x_112 = lean_nat_sub(x_29, x_111); -x_113 = lean_nat_sub(x_112, x_111); -lean_dec(x_112); -lean_inc(x_1); -x_114 = l_Lean_Expr_getRevArg_x21(x_1, x_113); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_115 = l_Lean_Meta_evalNat(x_114, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_115) == 0) +lean_object* x_55; uint8_t x_56; +x_55 = lean_unsigned_to_nat(4u); +x_56 = lean_nat_dec_eq(x_28, x_55); +if (x_56 == 0) { -lean_object* x_116; -x_116 = lean_ctor_get(x_115, 0); -lean_inc(x_116); -if (lean_obj_tag(x_116) == 0) +lean_object* x_57; uint8_t x_58; +x_57 = l_Lean_Meta_isNatProjInst___closed__6; +x_58 = lean_name_eq(x_26, x_57); +lean_dec(x_26); +if (x_58 == 0) { -uint8_t x_117; -lean_dec(x_29); -lean_dec(x_6); +lean_object* x_59; lean_object* x_60; +lean_dec(x_28); lean_dec(x_5); lean_dec(x_4); +lean_dec(x_3); lean_dec(x_1); -x_117 = !lean_is_exclusive(x_115); -if (x_117 == 0) -{ -lean_object* x_118; lean_object* x_119; -x_118 = lean_ctor_get(x_115, 0); -lean_dec(x_118); -x_119 = lean_box(0); -lean_ctor_set(x_115, 0, x_119); -return x_115; -} -else -{ -lean_object* x_120; lean_object* x_121; lean_object* x_122; -x_120 = lean_ctor_get(x_115, 1); -lean_inc(x_120); -lean_dec(x_115); -x_121 = lean_box(0); -x_122 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_122, 0, x_121); -lean_ctor_set(x_122, 1, x_120); -return x_122; -} -} -else -{ -lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; uint8_t x_128; lean_object* x_129; -x_123 = lean_ctor_get(x_115, 1); -lean_inc(x_123); -lean_dec(x_115); -x_124 = lean_ctor_get(x_116, 0); -lean_inc(x_124); -lean_dec(x_116); -x_125 = lean_nat_sub(x_29, x_28); -lean_dec(x_29); -x_126 = lean_nat_sub(x_125, x_111); -lean_dec(x_125); -x_127 = l_Lean_Expr_getRevArg_x21(x_1, x_126); -x_128 = 0; -x_129 = l___private_Lean_Meta_Offset_0__Lean_Meta_getOffsetAux(x_127, x_128, x_3, x_4, x_5, x_6, x_123); -if (lean_obj_tag(x_129) == 0) -{ -lean_object* x_130; -x_130 = lean_ctor_get(x_129, 0); -lean_inc(x_130); -if (lean_obj_tag(x_130) == 0) -{ -uint8_t x_131; -lean_dec(x_124); -x_131 = !lean_is_exclusive(x_129); -if (x_131 == 0) -{ -lean_object* x_132; lean_object* x_133; -x_132 = lean_ctor_get(x_129, 0); -lean_dec(x_132); -x_133 = lean_box(0); -lean_ctor_set(x_129, 0, x_133); -return x_129; -} -else -{ -lean_object* x_134; lean_object* x_135; lean_object* x_136; -x_134 = lean_ctor_get(x_129, 1); -lean_inc(x_134); -lean_dec(x_129); -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_134); -return x_136; -} -} -else -{ -uint8_t x_137; -x_137 = !lean_is_exclusive(x_130); -if (x_137 == 0) -{ -uint8_t x_138; -x_138 = !lean_is_exclusive(x_129); -if (x_138 == 0) -{ -lean_object* x_139; lean_object* x_140; uint8_t x_141; -x_139 = lean_ctor_get(x_130, 0); -x_140 = lean_ctor_get(x_129, 0); -lean_dec(x_140); -x_141 = !lean_is_exclusive(x_139); -if (x_141 == 0) -{ -lean_object* x_142; lean_object* x_143; -x_142 = lean_ctor_get(x_139, 1); -x_143 = lean_nat_add(x_142, x_124); -lean_dec(x_124); -lean_dec(x_142); -lean_ctor_set(x_139, 1, x_143); -return x_129; -} -else -{ -lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; -x_144 = lean_ctor_get(x_139, 0); -x_145 = lean_ctor_get(x_139, 1); -lean_inc(x_145); -lean_inc(x_144); -lean_dec(x_139); -x_146 = lean_nat_add(x_145, x_124); -lean_dec(x_124); -lean_dec(x_145); -x_147 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_147, 0, x_144); -lean_ctor_set(x_147, 1, x_146); -lean_ctor_set(x_130, 0, x_147); -return x_129; -} -} -else -{ -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_148 = lean_ctor_get(x_130, 0); -x_149 = lean_ctor_get(x_129, 1); -lean_inc(x_149); -lean_dec(x_129); -x_150 = lean_ctor_get(x_148, 0); -lean_inc(x_150); -x_151 = lean_ctor_get(x_148, 1); -lean_inc(x_151); -if (lean_is_exclusive(x_148)) { - lean_ctor_release(x_148, 0); - lean_ctor_release(x_148, 1); - x_152 = x_148; -} else { - lean_dec_ref(x_148); - x_152 = lean_box(0); -} -x_153 = lean_nat_add(x_151, x_124); -lean_dec(x_124); -lean_dec(x_151); -if (lean_is_scalar(x_152)) { - x_154 = lean_alloc_ctor(0, 2, 0); -} else { - x_154 = x_152; -} -lean_ctor_set(x_154, 0, x_150); -lean_ctor_set(x_154, 1, x_153); -lean_ctor_set(x_130, 0, x_154); -x_155 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_155, 0, x_130); -lean_ctor_set(x_155, 1, x_149); -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; -x_156 = lean_ctor_get(x_130, 0); -lean_inc(x_156); -lean_dec(x_130); -x_157 = lean_ctor_get(x_129, 1); -lean_inc(x_157); -if (lean_is_exclusive(x_129)) { - lean_ctor_release(x_129, 0); - lean_ctor_release(x_129, 1); - x_158 = x_129; -} else { - lean_dec_ref(x_129); - x_158 = lean_box(0); -} -x_159 = lean_ctor_get(x_156, 0); -lean_inc(x_159); -x_160 = lean_ctor_get(x_156, 1); -lean_inc(x_160); -if (lean_is_exclusive(x_156)) { - lean_ctor_release(x_156, 0); - lean_ctor_release(x_156, 1); - x_161 = x_156; -} else { - lean_dec_ref(x_156); - x_161 = lean_box(0); -} -x_162 = lean_nat_add(x_160, x_124); -lean_dec(x_124); -lean_dec(x_160); -if (lean_is_scalar(x_161)) { - x_163 = lean_alloc_ctor(0, 2, 0); -} else { - x_163 = x_161; -} -lean_ctor_set(x_163, 0, x_159); -lean_ctor_set(x_163, 1, x_162); -x_164 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_164, 0, x_163); -if (lean_is_scalar(x_158)) { - x_165 = lean_alloc_ctor(0, 2, 0); -} else { - x_165 = x_158; -} -lean_ctor_set(x_165, 0, x_164); -lean_ctor_set(x_165, 1, x_157); -return x_165; -} -} -} -else -{ -uint8_t x_166; -lean_dec(x_124); -x_166 = !lean_is_exclusive(x_129); -if (x_166 == 0) -{ -return x_129; +x_59 = lean_box(0); +x_60 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_60, 0, x_59); +lean_ctor_set(x_60, 1, x_6); +return x_60; } else { -lean_object* x_167; lean_object* x_168; lean_object* x_169; -x_167 = lean_ctor_get(x_129, 0); -x_168 = lean_ctor_get(x_129, 1); -lean_inc(x_168); -lean_inc(x_167); -lean_dec(x_129); -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; -} -} -} -} -else +lean_object* x_61; uint8_t x_62; +x_61 = lean_unsigned_to_nat(6u); +x_62 = lean_nat_dec_eq(x_28, x_61); +lean_dec(x_28); +if (x_62 == 0) { -uint8_t x_170; -lean_dec(x_29); -lean_dec(x_6); +lean_object* x_63; lean_object* x_64; lean_dec(x_5); lean_dec(x_4); +lean_dec(x_3); lean_dec(x_1); -x_170 = !lean_is_exclusive(x_115); -if (x_170 == 0) -{ -return x_115; +x_63 = lean_box(0); +x_64 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_64, 0, x_63); +lean_ctor_set(x_64, 1, x_6); +return x_64; } else { -lean_object* x_171; lean_object* x_172; lean_object* x_173; -x_171 = lean_ctor_get(x_115, 0); -x_172 = lean_ctor_get(x_115, 1); -lean_inc(x_172); -lean_inc(x_171); -lean_dec(x_115); -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; -} -} -} -} -} -} -default: +lean_object* x_65; +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_65 = l_Lean_Meta_unfoldProjInst_x3f(x_1, x_2, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_65) == 0) { -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_6); +lean_object* x_66; +x_66 = lean_ctor_get(x_65, 0); +lean_inc(x_66); +if (lean_obj_tag(x_66) == 0) +{ +uint8_t x_67; lean_dec(x_5); lean_dec(x_4); -if (x_2 == 0) +lean_dec(x_3); +x_67 = !lean_is_exclusive(x_65); +if (x_67 == 0) { -lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; -x_223 = lean_unsigned_to_nat(0u); -x_224 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_224, 0, x_1); -lean_ctor_set(x_224, 1, x_223); -x_225 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_225, 0, x_224); -x_226 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_226, 0, x_225); -lean_ctor_set(x_226, 1, x_7); -return x_226; +lean_object* x_68; lean_object* x_69; +x_68 = lean_ctor_get(x_65, 0); +lean_dec(x_68); +x_69 = lean_box(0); +lean_ctor_set(x_65, 0, x_69); +return x_65; } else { -lean_object* x_227; lean_object* x_228; -lean_dec(x_1); -x_227 = lean_box(0); -x_228 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_228, 0, x_227); -lean_ctor_set(x_228, 1, x_7); -return x_228; +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_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; } } +else +{ +lean_object* x_73; lean_object* x_74; +x_73 = lean_ctor_get(x_65, 1); +lean_inc(x_73); +lean_dec(x_65); +x_74 = lean_ctor_get(x_66, 0); +lean_inc(x_74); +lean_dec(x_66); +x_1 = x_74; +x_6 = x_73; +goto _start; } } else { -lean_dec(x_6); +uint8_t x_76; lean_dec(x_5); lean_dec(x_4); -if (x_2 == 0) -{ -lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; -x_229 = lean_unsigned_to_nat(0u); -x_230 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_230, 0, x_1); -lean_ctor_set(x_230, 1, x_229); -x_231 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_231, 0, x_230); -x_232 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_232, 0, x_231); -lean_ctor_set(x_232, 1, x_7); -return x_232; -} -else +lean_dec(x_3); +x_76 = !lean_is_exclusive(x_65); +if (x_76 == 0) { -lean_object* x_233; lean_object* x_234; -lean_dec(x_1); -x_233 = lean_box(0); -x_234 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_234, 0, x_233); -lean_ctor_set(x_234, 1, x_7); -return x_234; -} -} +return x_65; } -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_Offset_0__Lean_Meta_getOffsetAux___boxed(lean_object* x_1, 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: +else { -uint8_t x_8; lean_object* x_9; -x_8 = lean_unbox(x_2); -lean_dec(x_2); -x_9 = l___private_Lean_Meta_Offset_0__Lean_Meta_getOffsetAux(x_1, x_8, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_3); -return x_9; -} +lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_77 = lean_ctor_get(x_65, 0); +x_78 = lean_ctor_get(x_65, 1); +lean_inc(x_78); +lean_inc(x_77); +lean_dec(x_65); +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; } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Offset_0__Lean_Meta_getOffset(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 = 1; -x_8 = l___private_Lean_Meta_Offset_0__Lean_Meta_getOffsetAux(x_1, x_7, x_2, x_3, x_4, x_5, x_6); -return x_8; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Offset_0__Lean_Meta_getOffset___boxed(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___private_Lean_Meta_Offset_0__Lean_Meta_getOffset(x_1, x_2, x_3, x_4, x_5, x_6); -lean_dec(x_2); -return x_7; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Offset_0__Lean_Meta_isOffset(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: -{ -if (lean_obj_tag(x_1) == 5) +else { -lean_object* x_7; -x_7 = l_Lean_Expr_getAppFn(x_1); -switch (lean_obj_tag(x_7)) { -case 2: +lean_object* x_80; +lean_dec(x_28); +lean_dec(x_26); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_80 = l_Lean_Meta_unfoldProjInst_x3f(x_1, x_2, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_80) == 0) { -lean_object* x_8; uint8_t x_9; -lean_dec(x_7); -x_8 = l_Lean_instantiateMVars___at___private_Lean_Meta_Offset_0__Lean_Meta_withInstantiatedMVars___spec__1(x_1, x_2, x_3, x_4, x_5, x_6); -x_9 = !lean_is_exclusive(x_8); -if (x_9 == 0) +lean_object* x_81; +x_81 = lean_ctor_get(x_80, 0); +lean_inc(x_81); +if (lean_obj_tag(x_81) == 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, 0); -x_11 = lean_ctor_get(x_8, 1); -x_12 = lean_ctor_get(x_10, 0); -lean_inc(x_12); -lean_dec(x_10); -x_13 = l_Lean_Expr_getAppFn(x_12); -x_14 = l_Lean_Expr_isMVar(x_13); -lean_dec(x_13); -if (x_14 == 0) +uint8_t x_82; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_82 = !lean_is_exclusive(x_80); +if (x_82 == 0) { -lean_free_object(x_8); -x_1 = x_12; -x_6 = x_11; -goto _start; +lean_object* x_83; lean_object* x_84; +x_83 = lean_ctor_get(x_80, 0); +lean_dec(x_83); +x_84 = lean_box(0); +lean_ctor_set(x_80, 0, x_84); +return x_80; } else { -lean_object* x_16; -lean_dec(x_12); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -x_16 = lean_box(0); -lean_ctor_set(x_8, 0, x_16); -return x_8; +lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_85 = lean_ctor_get(x_80, 1); +lean_inc(x_85); +lean_dec(x_80); +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_85); +return x_87; } } else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_17 = lean_ctor_get(x_8, 0); -x_18 = lean_ctor_get(x_8, 1); -lean_inc(x_18); -lean_inc(x_17); -lean_dec(x_8); -x_19 = lean_ctor_get(x_17, 0); -lean_inc(x_19); -lean_dec(x_17); -x_20 = l_Lean_Expr_getAppFn(x_19); -x_21 = l_Lean_Expr_isMVar(x_20); -lean_dec(x_20); -if (x_21 == 0) -{ -x_1 = x_19; -x_6 = x_18; +lean_object* x_88; lean_object* x_89; +x_88 = lean_ctor_get(x_80, 1); +lean_inc(x_88); +lean_dec(x_80); +x_89 = lean_ctor_get(x_81, 0); +lean_inc(x_89); +lean_dec(x_81); +x_1 = x_89; +x_6 = x_88; goto _start; } +} else { -lean_object* x_23; lean_object* x_24; -lean_dec(x_19); +uint8_t x_91; lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -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_18); -return x_24; +x_91 = !lean_is_exclusive(x_80); +if (x_91 == 0) +{ +return x_80; } +else +{ +lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_92 = lean_ctor_get(x_80, 0); +x_93 = lean_ctor_get(x_80, 1); +lean_inc(x_93); +lean_inc(x_92); +lean_dec(x_80); +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; } } -case 4: +} +} +} +block_166: { -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_40; lean_object* x_64; uint8_t x_65; -x_25 = lean_ctor_get(x_7, 0); -lean_inc(x_25); -lean_dec(x_7); -x_26 = lean_unsigned_to_nat(0u); -x_27 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_26); -x_64 = l_Lean_Meta_evalNat_visit___closed__6; -x_65 = lean_name_eq(x_25, x_64); -if (x_65 == 0) +lean_object* x_97; uint8_t x_98; +lean_dec(x_96); +x_97 = l_Lean_Meta_evalNat_visit___closed__3; +x_98 = lean_name_eq(x_26, x_97); +if (x_98 == 0) { -lean_object* x_66; -x_66 = lean_box(0); -x_40 = x_66; -goto block_63; +lean_object* x_99; +x_99 = lean_box(0); +x_29 = x_99; +goto block_95; } else { -lean_object* x_67; uint8_t x_68; -x_67 = lean_unsigned_to_nat(1u); -x_68 = lean_nat_dec_eq(x_27, x_67); -if (x_68 == 0) +lean_object* x_100; uint8_t x_101; +x_100 = lean_unsigned_to_nat(2u); +x_101 = lean_nat_dec_eq(x_28, x_100); +if (x_101 == 0) { -lean_object* x_69; -x_69 = lean_box(0); -x_40 = x_69; -goto block_63; +lean_object* x_102; +x_102 = lean_box(0); +x_29 = x_102; +goto block_95; } else { -uint8_t x_70; lean_object* x_71; -lean_dec(x_27); -lean_dec(x_25); -x_70 = 1; -x_71 = l___private_Lean_Meta_Offset_0__Lean_Meta_getOffsetAux(x_1, x_70, x_2, x_3, x_4, x_5, x_6); -return x_71; -} -} -block_39: +lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; +lean_dec(x_26); +x_103 = lean_unsigned_to_nat(1u); +x_104 = lean_nat_sub(x_28, x_103); +x_105 = lean_nat_sub(x_104, x_103); +lean_dec(x_104); +lean_inc(x_1); +x_106 = l_Lean_Expr_getRevArg_x21(x_1, x_105); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_107 = l_Lean_Meta_evalNat(x_106, x_2, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_107) == 0) { -lean_object* x_29; uint8_t x_30; -lean_dec(x_28); -x_29 = l_Lean_Meta_isNatProjInst___closed__6; -x_30 = lean_name_eq(x_25, x_29); -lean_dec(x_25); -if (x_30 == 0) +lean_object* x_108; +x_108 = lean_ctor_get(x_107, 0); +lean_inc(x_108); +if (lean_obj_tag(x_108) == 0) { -lean_object* x_31; lean_object* x_32; -lean_dec(x_27); +uint8_t x_109; +lean_dec(x_28); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -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_6); -return x_32; -} -else -{ -lean_object* x_33; uint8_t x_34; -x_33 = lean_unsigned_to_nat(6u); -x_34 = lean_nat_dec_eq(x_27, x_33); -lean_dec(x_27); -if (x_34 == 0) +x_109 = !lean_is_exclusive(x_107); +if (x_109 == 0) { -lean_object* x_35; lean_object* x_36; -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -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_6); -return x_36; +lean_object* x_110; lean_object* x_111; +x_110 = lean_ctor_get(x_107, 0); +lean_dec(x_110); +x_111 = lean_box(0); +lean_ctor_set(x_107, 0, x_111); +return x_107; } else { -uint8_t x_37; lean_object* x_38; -x_37 = 1; -x_38 = l___private_Lean_Meta_Offset_0__Lean_Meta_getOffsetAux(x_1, x_37, x_2, x_3, x_4, x_5, x_6); -return x_38; +lean_object* x_112; lean_object* x_113; lean_object* x_114; +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_113); +lean_ctor_set(x_114, 1, x_112); +return x_114; } } -} -block_63: +else { -lean_object* x_41; uint8_t x_42; -lean_dec(x_40); -x_41 = l_Lean_Meta_evalNat_visit___closed__3; -x_42 = lean_name_eq(x_25, x_41); -if (x_42 == 0) +lean_object* x_115; uint8_t x_116; +x_115 = lean_ctor_get(x_107, 1); +lean_inc(x_115); +lean_dec(x_107); +x_116 = !lean_is_exclusive(x_108); +if (x_116 == 0) { -lean_object* x_43; uint8_t x_44; -x_43 = l_Lean_Meta_isNatProjInst___closed__15; -x_44 = lean_name_eq(x_25, x_43); -if (x_44 == 0) +lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; +x_117 = lean_ctor_get(x_108, 0); +x_118 = lean_nat_sub(x_28, x_27); +lean_dec(x_28); +x_119 = lean_nat_sub(x_118, x_103); +lean_dec(x_118); +x_120 = l_Lean_Expr_getRevArg_x21(x_1, x_119); +x_121 = l___private_Lean_Meta_Offset_0__Lean_Meta_getOffset(x_120, x_2, x_3, x_4, x_5, x_115); +if (lean_obj_tag(x_121) == 0) +{ +uint8_t x_122; +x_122 = !lean_is_exclusive(x_121); +if (x_122 == 0) +{ +lean_object* x_123; uint8_t x_124; +x_123 = lean_ctor_get(x_121, 0); +x_124 = !lean_is_exclusive(x_123); +if (x_124 == 0) { -lean_object* x_45; -x_45 = lean_box(0); -x_28 = x_45; -goto block_39; +lean_object* x_125; lean_object* x_126; +x_125 = lean_ctor_get(x_123, 1); +x_126 = lean_nat_add(x_125, x_117); +lean_dec(x_117); +lean_dec(x_125); +lean_ctor_set(x_123, 1, x_126); +lean_ctor_set(x_108, 0, x_123); +lean_ctor_set(x_121, 0, x_108); +return x_121; } else { -lean_object* x_46; uint8_t x_47; -x_46 = lean_unsigned_to_nat(4u); -x_47 = lean_nat_dec_eq(x_27, x_46); -if (x_47 == 0) -{ -lean_object* x_48; -x_48 = lean_box(0); -x_28 = x_48; -goto block_39; +lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; +x_127 = lean_ctor_get(x_123, 0); +x_128 = lean_ctor_get(x_123, 1); +lean_inc(x_128); +lean_inc(x_127); +lean_dec(x_123); +x_129 = lean_nat_add(x_128, x_117); +lean_dec(x_117); +lean_dec(x_128); +x_130 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_130, 0, x_127); +lean_ctor_set(x_130, 1, x_129); +lean_ctor_set(x_108, 0, x_130); +lean_ctor_set(x_121, 0, x_108); +return x_121; +} } else { -uint8_t x_49; lean_object* x_50; -lean_dec(x_27); -lean_dec(x_25); -x_49 = 1; -x_50 = l___private_Lean_Meta_Offset_0__Lean_Meta_getOffsetAux(x_1, x_49, x_2, x_3, x_4, x_5, x_6); -return x_50; +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_131 = lean_ctor_get(x_121, 0); +x_132 = lean_ctor_get(x_121, 1); +lean_inc(x_132); +lean_inc(x_131); +lean_dec(x_121); +x_133 = lean_ctor_get(x_131, 0); +lean_inc(x_133); +x_134 = lean_ctor_get(x_131, 1); +lean_inc(x_134); +if (lean_is_exclusive(x_131)) { + lean_ctor_release(x_131, 0); + lean_ctor_release(x_131, 1); + x_135 = x_131; +} else { + lean_dec_ref(x_131); + x_135 = lean_box(0); +} +x_136 = lean_nat_add(x_134, x_117); +lean_dec(x_117); +lean_dec(x_134); +if (lean_is_scalar(x_135)) { + x_137 = lean_alloc_ctor(0, 2, 0); +} else { + x_137 = x_135; } +lean_ctor_set(x_137, 0, x_133); +lean_ctor_set(x_137, 1, x_136); +lean_ctor_set(x_108, 0, x_137); +x_138 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_138, 0, x_108); +lean_ctor_set(x_138, 1, x_132); +return x_138; } } else { -lean_object* x_51; uint8_t x_52; -x_51 = lean_unsigned_to_nat(2u); -x_52 = lean_nat_dec_eq(x_27, x_51); -if (x_52 == 0) +uint8_t x_139; +lean_free_object(x_108); +lean_dec(x_117); +x_139 = !lean_is_exclusive(x_121); +if (x_139 == 0) { -lean_object* x_53; uint8_t x_54; -x_53 = l_Lean_Meta_isNatProjInst___closed__15; -x_54 = lean_name_eq(x_25, x_53); -if (x_54 == 0) +return x_121; +} +else { -lean_object* x_55; -x_55 = lean_box(0); -x_28 = x_55; -goto block_39; +lean_object* x_140; lean_object* x_141; lean_object* x_142; +x_140 = lean_ctor_get(x_121, 0); +x_141 = lean_ctor_get(x_121, 1); +lean_inc(x_141); +lean_inc(x_140); +lean_dec(x_121); +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 { -lean_object* x_56; uint8_t x_57; -x_56 = lean_unsigned_to_nat(4u); -x_57 = lean_nat_dec_eq(x_27, x_56); -if (x_57 == 0) +lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; +x_143 = lean_ctor_get(x_108, 0); +lean_inc(x_143); +lean_dec(x_108); +x_144 = lean_nat_sub(x_28, x_27); +lean_dec(x_28); +x_145 = lean_nat_sub(x_144, x_103); +lean_dec(x_144); +x_146 = l_Lean_Expr_getRevArg_x21(x_1, x_145); +x_147 = l___private_Lean_Meta_Offset_0__Lean_Meta_getOffset(x_146, x_2, x_3, x_4, x_5, x_115); +if (lean_obj_tag(x_147) == 0) { -lean_object* x_58; -x_58 = lean_box(0); -x_28 = x_58; -goto block_39; +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_148 = lean_ctor_get(x_147, 0); +lean_inc(x_148); +x_149 = lean_ctor_get(x_147, 1); +lean_inc(x_149); +if (lean_is_exclusive(x_147)) { + lean_ctor_release(x_147, 0); + lean_ctor_release(x_147, 1); + x_150 = x_147; +} else { + lean_dec_ref(x_147); + x_150 = lean_box(0); +} +x_151 = lean_ctor_get(x_148, 0); +lean_inc(x_151); +x_152 = lean_ctor_get(x_148, 1); +lean_inc(x_152); +if (lean_is_exclusive(x_148)) { + lean_ctor_release(x_148, 0); + lean_ctor_release(x_148, 1); + x_153 = x_148; +} else { + lean_dec_ref(x_148); + x_153 = lean_box(0); +} +x_154 = lean_nat_add(x_152, x_143); +lean_dec(x_143); +lean_dec(x_152); +if (lean_is_scalar(x_153)) { + x_155 = lean_alloc_ctor(0, 2, 0); +} else { + x_155 = x_153; +} +lean_ctor_set(x_155, 0, x_151); +lean_ctor_set(x_155, 1, x_154); +x_156 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_156, 0, x_155); +if (lean_is_scalar(x_150)) { + x_157 = lean_alloc_ctor(0, 2, 0); +} else { + x_157 = x_150; +} +lean_ctor_set(x_157, 0, x_156); +lean_ctor_set(x_157, 1, x_149); +return x_157; } else { -uint8_t x_59; lean_object* x_60; -lean_dec(x_27); -lean_dec(x_25); -x_59 = 1; -x_60 = l___private_Lean_Meta_Offset_0__Lean_Meta_getOffsetAux(x_1, x_59, x_2, x_3, x_4, x_5, x_6); -return x_60; +lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; +lean_dec(x_143); +x_158 = lean_ctor_get(x_147, 0); +lean_inc(x_158); +x_159 = lean_ctor_get(x_147, 1); +lean_inc(x_159); +if (lean_is_exclusive(x_147)) { + lean_ctor_release(x_147, 0); + lean_ctor_release(x_147, 1); + x_160 = x_147; +} else { + lean_dec_ref(x_147); + x_160 = lean_box(0); +} +if (lean_is_scalar(x_160)) { + x_161 = lean_alloc_ctor(1, 2, 0); +} else { + x_161 = x_160; +} +lean_ctor_set(x_161, 0, x_158); +lean_ctor_set(x_161, 1, x_159); +return x_161; +} } } } else { -uint8_t x_61; lean_object* x_62; -lean_dec(x_27); -lean_dec(x_25); -x_61 = 1; -x_62 = l___private_Lean_Meta_Offset_0__Lean_Meta_getOffsetAux(x_1, x_61, x_2, x_3, x_4, x_5, x_6); -return x_62; +uint8_t x_162; +lean_dec(x_28); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_162 = !lean_is_exclusive(x_107); +if (x_162 == 0) +{ +return x_107; +} +else +{ +lean_object* x_163; lean_object* x_164; lean_object* x_165; +x_163 = lean_ctor_get(x_107, 0); +x_164 = lean_ctor_get(x_107, 1); +lean_inc(x_164); +lean_inc(x_163); +lean_dec(x_107); +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; +} +} } } } } default: { -lean_object* x_72; lean_object* x_73; +lean_object* x_198; lean_object* x_199; +lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_72 = lean_box(0); -x_73 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_73, 0, x_72); -lean_ctor_set(x_73, 1, x_6); -return x_73; +x_198 = lean_box(0); +x_199 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_199, 0, x_198); +lean_ctor_set(x_199, 1, x_6); +return x_199; } } } else { -lean_object* x_74; lean_object* x_75; +lean_object* x_200; lean_object* x_201; lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -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_6); -return x_75; +x_200 = lean_box(0); +x_201 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_201, 0, x_200); +lean_ctor_set(x_201, 1, x_6); +return x_201; +} } } +LEAN_EXPORT lean_object* l___private_Lean_Meta_Offset_0__Lean_Meta_getOffset___boxed(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___private_Lean_Meta_Offset_0__Lean_Meta_getOffset(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_2); +return x_7; +} } -LEAN_EXPORT lean_object* l___private_Lean_Meta_Offset_0__Lean_Meta_isOffset___boxed(lean_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_Meta_Offset_0__Lean_Meta_isOffset_x3f___boxed(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___private_Lean_Meta_Offset_0__Lean_Meta_isOffset(x_1, x_2, x_3, x_4, x_5, x_6); +x_7 = l___private_Lean_Meta_Offset_0__Lean_Meta_isOffset_x3f(x_1, x_2, x_3, x_4, x_5, x_6); lean_dec(x_2); return x_7; } @@ -3722,7 +3373,7 @@ lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_1); -x_20 = l___private_Lean_Meta_Offset_0__Lean_Meta_isOffset(x_1, x_3, x_4, x_5, x_6, x_19); +x_20 = l___private_Lean_Meta_Offset_0__Lean_Meta_isOffset_x3f(x_1, x_3, x_4, x_5, x_6, x_19); if (lean_obj_tag(x_20) == 0) { lean_object* x_21; @@ -3791,7 +3442,7 @@ lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_2); -x_35 = l___private_Lean_Meta_Offset_0__Lean_Meta_isOffset(x_2, x_3, x_4, x_5, x_6, x_33); +x_35 = l___private_Lean_Meta_Offset_0__Lean_Meta_isOffset_x3f(x_2, x_3, x_4, x_5, x_6, x_33); if (lean_obj_tag(x_35) == 0) { lean_object* x_36; @@ -4448,7 +4099,7 @@ lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_2); -x_181 = l___private_Lean_Meta_Offset_0__Lean_Meta_isOffset(x_2, x_3, x_4, x_5, x_6, x_178); +x_181 = l___private_Lean_Meta_Offset_0__Lean_Meta_isOffset_x3f(x_2, x_3, x_4, x_5, x_6, x_178); if (lean_obj_tag(x_181) == 0) { lean_object* x_182; diff --git a/stage0/stdlib/Lean/Meta/PPGoal.c b/stage0/stdlib/Lean/Meta/PPGoal.c index 41ac1a7d50cc..815875a8841b 100644 --- a/stage0/stdlib/Lean/Meta/PPGoal.c +++ b/stage0/stdlib/Lean/Meta/PPGoal.c @@ -13,7 +13,6 @@ #ifdef __cplusplus extern "C" { #endif -uint8_t l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getGoalPrefix(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_Lean_Meta_getGoalPrefix___boxed(lean_object*); @@ -55,7 +54,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_PPGoal___hyg_50_(le LEAN_EXPORT lean_object* l_Lean_Meta_ppGoal_pushPending(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Std_Format_isNil(lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_PPGoal___hyg_6____closed__7; -lean_object* l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_isLHSGoal_x3f(lean_object*); static lean_object* l_Lean_Meta_ppGoal_ppVars___closed__2; size_t lean_usize_of_nat(lean_object*); @@ -78,7 +76,6 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_ppGoal___spec_ static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_PPGoal___hyg_6____closed__6; static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_PPGoal___hyg_130____closed__5; static lean_object* l_Lean_Meta_ppGoal_pushPending___closed__3; -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_ppGoal___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_EXPORT lean_object* l_Lean_PersistentArray_foldlM___at_Lean_Meta_ppGoal___spec__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*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_PPGoal___hyg_130____closed__3; @@ -87,6 +84,7 @@ static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_PPGoal___hyg_130____clos static lean_object* l_Lean_Meta_ppGoal_pushPending___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_pp_inaccessibleNames; static lean_object* l_Lean_Meta_ppGoal___closed__1; +uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_PPGoal___hyg_90____closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_PPGoal___hyg_90_(lean_object*); static lean_object* l_Lean_Meta_ppGoal_ppVars___closed__1; @@ -97,7 +95,6 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_ppGoal___spec_ static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_PPGoal___hyg_90____closed__4; lean_object* lean_usize_to_nat(size_t); LEAN_EXPORT lean_object* l_Std_Format_joinSep___at_Lean_Meta_ppGoal_pushPending___spec__1(lean_object*, lean_object*); -lean_object* l_Lean_Option_register___at_Std_Format_initFn____x40_Lean_Data_Format___hyg_95____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_ppGoal___closed__8; uint8_t l_Lean_LocalDecl_isImplementationDetail(lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_PPGoal___hyg_90____closed__1; @@ -105,10 +102,12 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_Persisten static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_PPGoal___hyg_90____closed__5; static lean_object* l_Lean_Meta_ppGoal_pushPending___closed__1; static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_PPGoal___hyg_50____closed__1; +lean_object* l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_simp_macro_scopes(lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_PPGoal___hyg_130____closed__2; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_PPGoal___hyg_50____closed__3; +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_ppGoal_ppVars___closed__4; LEAN_EXPORT lean_object* l_Lean_PersistentArray_foldlM___at_Lean_Meta_ppGoal___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* l_Lean_Name_mkStr2(lean_object*, lean_object*); @@ -116,6 +115,7 @@ static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_PPGoal___hyg_6____closed LEAN_EXPORT lean_object* l_Lean_Meta_ppGoal_ppVars___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_initFn____x40_Lean_Meta_PPGoal___hyg_90____closed__2; lean_object* lean_nat_sub(lean_object*, lean_object*); +lean_object* l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_ppGoal___lambda__1(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_erase_macro_scopes(lean_object*); lean_object* l_List_reverse___rarg(lean_object*); @@ -231,7 +231,7 @@ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_PPGoal___hyg_6____closed__3; x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_PPGoal___hyg_6____closed__5; x_4 = l_Lean_Meta_initFn____x40_Lean_Meta_PPGoal___hyg_6____closed__8; -x_5 = l_Lean_Option_register___at_Std_Format_initFn____x40_Lean_Data_Format___hyg_95____spec__1(x_2, x_3, x_4, x_1); +x_5 = l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__1(x_2, x_3, x_4, x_1); return x_5; } } @@ -295,7 +295,7 @@ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_PPGoal___hyg_50____closed__2; x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_PPGoal___hyg_50____closed__4; x_4 = l_Lean_Meta_initFn____x40_Lean_Meta_PPGoal___hyg_50____closed__5; -x_5 = l_Lean_Option_register___at_Std_Format_initFn____x40_Lean_Data_Format___hyg_95____spec__1(x_2, x_3, x_4, x_1); +x_5 = l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__1(x_2, x_3, x_4, x_1); return x_5; } } @@ -359,7 +359,7 @@ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_PPGoal___hyg_90____closed__2; x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_PPGoal___hyg_90____closed__4; x_4 = l_Lean_Meta_initFn____x40_Lean_Meta_PPGoal___hyg_90____closed__5; -x_5 = l_Lean_Option_register___at_Std_Format_initFn____x40_Lean_Data_Format___hyg_95____spec__1(x_2, x_3, x_4, x_1); +x_5 = l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__1(x_2, x_3, x_4, x_1); return x_5; } } @@ -423,7 +423,7 @@ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_PPGoal___hyg_130____closed__2; x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_PPGoal___hyg_130____closed__4; x_4 = l_Lean_Meta_initFn____x40_Lean_Meta_PPGoal___hyg_130____closed__5; -x_5 = l_Lean_Option_register___at_Std_Format_initFn____x40_Lean_Data_Format___hyg_95____spec__1(x_2, x_3, x_4, x_1); +x_5 = l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__1(x_2, x_3, x_4, x_1); return x_5; } } @@ -867,7 +867,7 @@ x_13 = lean_ctor_get(x_6, 3); lean_inc(x_13); lean_dec(x_6); x_14 = lean_simp_macro_scopes(x_12); -x_15 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_13, x_7, x_8, x_9, x_10, x_11); +x_15 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_13, x_7, x_8, x_9, x_10, x_11); x_16 = !lean_is_exclusive(x_15); if (x_16 == 0) { @@ -1147,7 +1147,7 @@ x_82 = lean_ctor_get(x_80, 1); lean_inc(x_82); lean_dec(x_80); x_83 = l___private_Lean_Meta_PPGoal_0__Lean_Meta_addLine(x_81); -x_84 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_77, x_7, x_8, x_9, x_10, x_82); +x_84 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_77, x_7, x_8, x_9, x_10, x_82); x_85 = lean_ctor_get(x_84, 0); lean_inc(x_85); x_86 = lean_ctor_get(x_84, 1); @@ -1186,7 +1186,7 @@ return x_97; else { lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; -x_98 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_78, x_7, x_8, x_9, x_10, x_89); +x_98 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_78, x_7, x_8, x_9, x_10, x_89); x_99 = lean_ctor_get(x_98, 0); lean_inc(x_99); x_100 = lean_ctor_get(x_98, 1); @@ -7021,7 +7021,7 @@ lean_dec(x_21); x_24 = l___private_Lean_Meta_PPGoal_0__Lean_Meta_addLine(x_22); x_25 = lean_ctor_get(x_7, 2); lean_inc(x_25); -x_26 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_25, x_8, x_9, x_10, x_11, x_23); +x_26 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_25, x_8, x_9, x_10, x_11, x_23); x_27 = lean_ctor_get(x_26, 0); lean_inc(x_27); x_28 = lean_ctor_get(x_26, 1); @@ -7336,11 +7336,11 @@ lean_dec(x_12); x_15 = lean_ctor_get(x_4, 2); lean_inc(x_15); x_16 = l_Lean_Meta_ppGoal___closed__4; -x_17 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_15, x_16); +x_17 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_15, x_16); x_18 = l_Lean_Meta_ppGoal___closed__5; -x_19 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_15, x_18); +x_19 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_15, x_18); x_20 = l_Lean_Meta_ppGoal___closed__6; -x_21 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_15, x_20); +x_21 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_15, x_20); x_22 = lean_ctor_get(x_14, 1); lean_inc(x_22); x_23 = lean_box(0); @@ -7368,7 +7368,7 @@ lean_closure_set(x_33, 3, x_32); lean_closure_set(x_33, 4, x_26); lean_closure_set(x_33, 5, x_29); lean_closure_set(x_33, 6, x_14); -x_34 = l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__3___rarg(x_26, x_27, x_33, x_2, x_3, x_4, x_5, x_10); +x_34 = l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2___rarg(x_26, x_27, x_33, x_2, x_3, x_4, x_5, x_10); return x_34; } } @@ -7406,11 +7406,11 @@ lean_dec(x_38); x_42 = lean_ctor_get(x_4, 2); lean_inc(x_42); x_43 = l_Lean_Meta_ppGoal___closed__4; -x_44 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_42, x_43); +x_44 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_42, x_43); x_45 = l_Lean_Meta_ppGoal___closed__5; -x_46 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_42, x_45); +x_46 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_42, x_45); x_47 = l_Lean_Meta_ppGoal___closed__6; -x_48 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_42, x_47); +x_48 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_42, x_47); x_49 = lean_ctor_get(x_41, 1); lean_inc(x_49); x_50 = lean_box(0); @@ -7438,7 +7438,7 @@ lean_closure_set(x_60, 3, x_59); lean_closure_set(x_60, 4, x_53); lean_closure_set(x_60, 5, x_56); lean_closure_set(x_60, 6, x_41); -x_61 = l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__3___rarg(x_53, x_54, x_60, x_2, x_3, x_4, x_5, x_36); +x_61 = l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2___rarg(x_53, x_54, x_60, x_2, x_3, x_4, x_5, x_36); return x_61; } } diff --git a/stage0/stdlib/Lean/Meta/SizeOf.c b/stage0/stdlib/Lean/Meta/SizeOf.c index 5ecada794fc6..2fe4afcd8936 100644 --- a/stage0/stdlib/Lean/Meta/SizeOf.c +++ b/stage0/stdlib/Lean/Meta/SizeOf.c @@ -14,22 +14,25 @@ extern "C" { #endif LEAN_EXPORT lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfMotives___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266_(lean_object*); lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); -uint8_t l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Meta_mkSizeOfSpecLemmaInstance___closed__2; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Meta_SizeOf_0__Lean_Meta_SizeOfSpecNested_mkSizeOfAuxLemmaProof___spec__5(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* l___private_Init_Util_0__outOfBounds___rarg(lean_object*); LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Meta_mkSizeOfInstances___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*); +static lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_ignoreFieldType___closed__2; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Meta_SizeOf_0__Lean_Meta_isRecField_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_SizeOfSpecNested_mkMinorProof___lambda__1___closed__1; static lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_SizeOfSpecNested_mkMinorProof___lambda__1___closed__6; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_SizeOf_0__Lean_Meta_SizeOfSpecNested_mkSizeOfAuxLemmaProof___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfSpecTheorem___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_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8426____closed__4; static lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfMinors___rarg___closed__6; LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Meta_SizeOf_0__Lean_Meta_SizeOfSpecNested_mkMinorProof___spec__1(lean_object*, 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*); lean_object* l_Lean_Meta_mkAppOptM(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mkSizeOfFn___lambda__8___boxed(lean_object**); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__4; lean_object* lean_mk_empty_array_with_capacity(lean_object*); lean_object* l_Lean_mkAppN(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_SizeOfSpecNested_mkMinorProof___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -40,10 +43,10 @@ static lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_isRecField_x3f___c static lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfMinors_loop___rarg___lambda__2___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_SizeOfSpecNested_mkSizeOfAuxLemma___lambda__4(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_getAppNumArgsAux(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__11; static lean_object* l_List_forIn_loop___at_Lean_Meta_mkSizeOfFns___spec__3___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_mkSizeOfInstances(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_whnfI(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__3; lean_object* l_Lean_Meta_unfoldDefinition(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_SizeOf_0__Lean_Meta_isInductiveHypothesis_x3f___spec__2(lean_object*, lean_object*, size_t, size_t); static lean_object* l_panic___at___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfMinors___spec__1___rarg___closed__1; @@ -66,16 +69,17 @@ static lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_SizeOfSpecNested_m LEAN_EXPORT lean_object* l_Lean_Meta_SizeOfSpecNested_main_step(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_Meta_mkSizeOfFns___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* l_Lean_Meta_isExprDefEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__7; LEAN_EXPORT lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_ignoreField(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__8; LEAN_EXPORT lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfRecFieldFormIH(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isAppOfArity(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfMinors_loop___spec__1(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___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfMinors___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isApp(lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8465____closed__1; LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfMinors___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Meta_mkSizeOfFns___spec__3(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_SizeOf___hyg_9251____closed__1; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8426____closed__6; LEAN_EXPORT lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkLocalInstances_loop(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_mkSizeOfSpecLemmaInstance___spec__1___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_sort___override(lean_object*); @@ -84,7 +88,6 @@ lean_object* lean_array_push(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mkSizeOfFn___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_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Meta_SizeOf_0__Lean_Meta_isRecField_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*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8411____closed__4; LEAN_EXPORT lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_isRecField_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfSpecTheorem___lambda__2___closed__1; static lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkLocalInstances_loop___rarg___closed__2; @@ -108,9 +111,7 @@ static lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfMotives_lo LEAN_EXPORT lean_object* l_Lean_Meta_mkSizeOfFn___lambda__6___boxed(lean_object**); lean_object* lean_environment_find(lean_object*, lean_object*); lean_object* l_Lean_RecursorVal_getFirstMinorIdx(lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8411____closed__6; LEAN_EXPORT lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfSpecTheorem___lambda__8___boxed(lean_object**); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__12; static lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_SizeOfSpecNested_mkMinorProofStep___closed__2; LEAN_EXPORT lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfSpecTheorem___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* 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*); @@ -119,6 +120,7 @@ LEAN_EXPORT lean_object* l_List_mapM_loop___at___private_Lean_Meta_SizeOf_0__Lea static lean_object* l_Lean_Meta_mkSizeOfInstances___closed__2; lean_object* l_Lean_Meta_forallBoundedTelescope___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_mkSizeOfFn___lambda__4___closed__1; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__10; static lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_SizeOfSpecNested_mkMinorProof___closed__4; lean_object* l_Lean_stringToMessageData(lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Meta_SizeOf_0__Lean_Meta_SizeOfSpecNested_mkSizeOfAuxLemmaProof___spec__6___lambda__1___closed__1; @@ -126,6 +128,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_SizeOfSpecNes uint8_t l___private_Lean_Util_Trace_0__Lean_checkTraceOption(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfSpecTheorems___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___private_Lean_Meta_SizeOf_0__Lean_Meta_SizeOfSpecNested_mkSizeOfAuxLemma___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*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8426____closed__5; static lean_object* l_Lean_getConstInfoRec___at_Lean_Meta_mkSizeOfFn___spec__1___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfMotives_loop___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_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_mkSizeOfFn___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -139,7 +142,6 @@ static lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfMotives_lo LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfMinors_loop___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_SizeOf_0__Lean_Meta_mkLocalInstances___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getConstInfoRec___at___private_Lean_Meta_SizeOf_0__Lean_Meta_SizeOfSpecNested_mkSizeOfAuxLemmaProof___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8411____closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_mkSizeOfFn___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*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_appArg_x21(lean_object*); static lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkLocalInstances_loop___rarg___closed__1; @@ -154,8 +156,11 @@ static lean_object* l_List_forIn_loop___at_Lean_Meta_mkSizeOfInstances___spec__2 static lean_object* l_Lean_Meta_mkSizeOfFn___lambda__3___closed__2; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Meta_SizeOf_0__Lean_Meta_SizeOfSpecNested_mkSizeOfAuxLemmaProof___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*); static lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfSpecTheorem___lambda__7___closed__4; +LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8426_(lean_object*); static lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_SizeOfSpecNested_mkMinorProof___lambda__1___closed__5; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8465____closed__3; static lean_object* l_Lean_Meta_mkSizeOfFn___lambda__2___closed__1; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__6; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Meta_SizeOf_0__Lean_Meta_SizeOfSpecNested_mkSizeOfAuxLemmaProof_mkSizeOf___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_EXPORT lean_object* l_Lean_Meta_SizeOfSpecNested_throwFailed___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Meta_SizeOf_0__Lean_Meta_SizeOfSpecNested_mkMinorProof___spec__1___closed__1; @@ -167,17 +172,16 @@ lean_object* l_Lean_Expr_natAdd_x3f(lean_object*); static lean_object* l_Lean_getConstInfoInduct___at_Lean_Meta_mkSizeOfFns___spec__1___closed__2; static lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_SizeOfSpecNested_mkMinorProof___lambda__1___closed__3; size_t lean_usize_of_nat(lean_object*); +static lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_ignoreFieldType___closed__1; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8465____closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfMinors_loop___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_EXPORT lean_object* l_Lean_Meta_mkSizeOfFn___lambda__7___boxed(lean_object**); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Meta_SizeOf_0__Lean_Meta_isRecField_x3f___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_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251_(lean_object*); static lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_SizeOfSpecNested_mkSizeOfAuxLemma___lambda__4___closed__1; static lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_SizeOfSpecNested_mkMinorProof___lambda__1___closed__7; LEAN_EXPORT lean_object* l_Lean_Meta_SizeOfSpecNested_main_loop___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_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__3; lean_object* l_Lean_throwError___at_Lean_Expr_abstractRangeM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_mkSizeOfFn___closed__1; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__9; lean_object* lean_st_ref_take(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfMinors_loop___rarg___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_getConstInfoRec___at_Lean_Meta_mkSizeOfFn___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -185,10 +189,10 @@ uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l_Lean_Meta_SizeOfSpecNested_throwFailed___rarg___closed__1; uint8_t lean_expr_eqv(lean_object*, lean_object*); lean_object* l_Lean_RecursorVal_getFirstIndexIdx(lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8450____closed__4; static lean_object* l_Lean_Meta_SizeOfSpecNested_main_loop___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_mkSizeOfSpecLemmaInstance___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfSpecTheorem___lambda__4___closed__2; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__7; static lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfSpecTheorem___lambda__5___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_SizeOfSpecNested_mkMinorProofStep(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_SizeOf_0__Lean_Meta_mkSizeOfMotives_loop___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -203,12 +207,10 @@ static lean_object* l_Lean_Meta_mkSizeOfFns___closed__1; static lean_object* l_Lean_Meta_mkSizeOfFn___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Lean_getConstInfoInduct___at_Lean_Meta_mkSizeOfFns___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Meta_SizeOf_0__Lean_Meta_SizeOfSpecNested_mkSizeOfAuxLemmaProof_mkSizeOf___spec__1(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*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__10; lean_object* l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescopeReducing___at___private_Lean_Meta_SizeOf_0__Lean_Meta_SizeOfSpecNested_mkSizeOfAuxLemma___spec__3(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_SizeOf_0__Lean_Meta_SizeOfSpecNested_mkSizeOfAuxLemma___spec__2(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfMinors(lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__2; lean_object* l_Lean_Meta_isInductivePredicate(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkLocalInstances_loop___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___private_Lean_Meta_SizeOf_0__Lean_Meta_SizeOfSpecNested_mkSizeOfAuxLemma___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*); @@ -243,11 +245,9 @@ LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Meta_Siz LEAN_EXPORT lean_object* l_Lean_Meta_mkSizeOfSpecLemmaInstance___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfSpecTheorems___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_levelZero; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8411____closed__2; extern lean_object* l_Lean_instInhabitedExpr; lean_object* l_Lean_RBNode_insert___at_Lean_NameMap_insert___spec__1___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfSpecTheorems___spec__2(lean_object*, 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_initFn____x40_Lean_Meta_SizeOf___hyg_8450____closed__5; LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Meta_mkSizeOfInstances___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_mkSizeOfFn___closed__2; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_mkSizeOfFn___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -265,6 +265,7 @@ LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Meta_mkSizeOfFns___spe LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_SizeOf_0__Lean_Meta_SizeOfSpecNested_mkSizeOfAuxLemmaProof___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkLocalInstances_loop___rarg___lambda__1___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_SizeOfSpecNested_mkSizeOfAuxLemmaProof___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*); +uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfMinors___rarg___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_SizeOfSpecNested_mkSizeOfAuxLemmaProof_mkSizeOf___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfSpecTheorem___lambda__7___boxed(lean_object**); @@ -282,7 +283,7 @@ lean_object* l_Lean_Expr_appFn_x21(lean_object*); static lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_SizeOfSpecNested_mkSizeOfAuxLemma___lambda__2___closed__1; static lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_SizeOfSpecNested_mkSizeOfAuxLemma___lambda__3___closed__1; static lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfSpecTheorem___lambda__3___closed__2; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8411____closed__1; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8465____closed__5; static lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_SizeOfSpecNested_mkSizeOfAuxLemma___lambda__3___closed__2; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_SizeOfSpecNested_throwUnexpected___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_forIn_loop___at_Lean_Meta_mkSizeOfInstances___spec__2___lambda__2___closed__7; @@ -292,11 +293,11 @@ lean_object* l_Lean_Meta_instInhabitedMetaM___boxed(lean_object*, lean_object*, static lean_object* l_Lean_getConstInfoInduct___at_Lean_Meta_mkSizeOfFns___spec__1___closed__1; lean_object* l_Array_append___rarg(lean_object*, lean_object*); uint8_t l_Lean_Environment_contains(lean_object*, lean_object*); -lean_object* l_Lean_Option_register___at_Std_Format_initFn____x40_Lean_Data_Format___hyg_95____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_inheritedTraceOptions; LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_SizeOf_0__Lean_Meta_SizeOfSpecNested_mkSizeOfAuxLemma___spec__2___rarg___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_mkSizeOfFns___closed__4; lean_object* l_Lean_MessageData_ofExpr(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8465_(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfMinors_loop___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_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_SizeOfSpecNested_throwUnexpected___rarg___closed__1; static lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfSpecTheorem___lambda__4___closed__1; @@ -312,8 +313,12 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_SizeOfSpecNested_throw static lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfMotives_loop___rarg___lambda__1___closed__2; lean_object* lean_name_append_index_after(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Meta_SizeOf_0__Lean_Meta_isRecField_x3f___spec__1___lambda__1(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_SizeOf___hyg_8465____closed__4; static lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfSpecTheorem___lambda__3___closed__1; lean_object* l_Lean_Meta_mkCongr(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8426____closed__3; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_ignoreFieldType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_getConstInfoRec___at_Lean_Meta_mkSizeOfFn___spec__1___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_SizeOfSpecNested_mkSizeOfAuxLemma___lambda__3___boxed(lean_object**); LEAN_EXPORT lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkLocalInstances(lean_object*); @@ -340,8 +345,6 @@ uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Meta_SizeOf_0__Lean_Meta_SizeOfSpecNested_mkMinorProofStep___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*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_SizeOfSpecNested_mkSizeOfAuxLemmaProof___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_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8450_(lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8450____closed__2; static lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfRecFieldFormIH___closed__3; LEAN_EXPORT lean_object* l_Array_contains___at___private_Lean_Meta_SizeOf_0__Lean_Meta_isInductiveHypothesis_x3f___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkLocalInstances_loop___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -352,18 +355,21 @@ lean_object* l_Lean_indentExpr(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescopeReducing___at_Lean_Meta_mkSizeOfInstances___spec__1(lean_object*); static lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfSpecTheorem___lambda__5___closed__1; static lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_SizeOfSpecNested_mkSizeOfAuxLemma___lambda__2___closed__2; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8426____closed__2; static lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_SizeOfSpecNested_recToSizeOf___closed__2; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__11; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__12; uint8_t l_Lean_Expr_isConstOf(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*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__2; LEAN_EXPORT lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfSpecTheorem___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_EXPORT uint8_t l_Array_contains___at___private_Lean_Meta_SizeOf_0__Lean_Meta_isInductiveHypothesis_x3f___spec__1(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_Meta_SizeOf_0__Lean_Meta_SizeOfSpecNested_mkMinorProof(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_SizeOf_0__Lean_Meta_ignoreField___closed__1; lean_object* lean_panic_fn(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SizeOfSpecNested_throwFailed(lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__5; lean_object* lean_nat_sub(lean_object*, lean_object*); +lean_object* l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfMotives_loop___rarg___closed__7; lean_object* l_Lean_Expr_getAppFn(lean_object*); static lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkLocalInstances___rarg___closed__1; @@ -375,7 +381,6 @@ static lean_object* l_Lean_Meta_mkSizeOfFns___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfSpecTheorem___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*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfRecFieldFormIH___closed__2; LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfMinors___spec__1(lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8450____closed__1; static lean_object* l_Lean_Meta_mkSizeOfSpecLemmaInstance___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_SizeOfSpecNested_throwUnexpected___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_ofSubarray___rarg(lean_object*); @@ -386,21 +391,21 @@ lean_object* l_Array_back___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfSpecTheorem(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_reverse___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallBoundedTelescope___at___private_Lean_Meta_SizeOf_0__Lean_Meta_SizeOfSpecNested_mkSizeOfAuxLemmaProof___spec__7___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_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__8; LEAN_EXPORT lean_object* l_Lean_Meta_SizeOfSpecNested_main_loop(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mkSizeOfFn(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfSpecTheorem___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___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfMotives_loop___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_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__9; static lean_object* l_Lean_Meta_mkSizeOfSpecLemmaName___closed__1; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_SizeOfSpecNested_throwFailed___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfMotives_loop___rarg___closed__8; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8426____closed__1; size_t lean_usize_add(size_t, size_t); extern lean_object* l_Lean_instInhabitedName; LEAN_EXPORT lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_isRecField_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_SizeOf_0__Lean_Meta_SizeOfSpecNested_mkSizeOfAuxLemma___spec__2___rarg___lambda__1(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); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_mkSizeOfFn___lambda__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_Meta_SizeOfSpecNested_throwUnexpected___rarg___boxed(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*); @@ -413,7 +418,6 @@ lean_object* l_List_redLength___rarg(lean_object*); 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*); lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(lean_object*, lean_object*, lean_object*); static lean_object* l_List_forIn_loop___at_Lean_Meta_mkSizeOfInstances___spec__2___lambda__2___closed__3; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8450____closed__3; LEAN_EXPORT lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_SizeOfSpecNested_mkMinorProofStep___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getConstInfoRec___at_Lean_Meta_mkSizeOfFn___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -425,14 +429,12 @@ lean_object* lean_array_get_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfSpecTheorem___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfSpecTheorem___lambda__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_mkSizeOfSpecLemmaName(lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__5; static lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfMinors___rarg___closed__4; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8411____closed__5; lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Meta_mkSizeOfFn___lambda__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_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8411_(lean_object*); -static lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_ignoreField___closed__2; 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_List_forIn_loop___at_Lean_Meta_mkSizeOfInstances___spec__2___lambda__2___closed__5; lean_object* lean_nat_add(lean_object*, lean_object*); @@ -451,7 +453,6 @@ static lean_object* l_Lean_Meta_mkSizeOfFn___lambda__4___closed__2; lean_object* l_Lean_MessageData_ofName(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkLocalInstances_loop___rarg___lambda__1(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_Meta_mkSizeOfInstances___spec__2___lambda__2___closed__1; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescopeReducing___at_Lean_Meta_mkSizeOfInstances___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_SizeOfSpecNested_main_loop___closed__2; static lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfSpecTheorem___lambda__7___closed__1; @@ -2083,7 +2084,7 @@ x_2 = lean_alloc_closure((void*)(l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSiz return x_2; } } -static lean_object* _init_l___private_Lean_Meta_SizeOf_0__Lean_Meta_ignoreField___closed__1() { +static lean_object* _init_l___private_Lean_Meta_SizeOf_0__Lean_Meta_ignoreFieldType___closed__1() { _start: { lean_object* x_1; @@ -2091,17 +2092,17 @@ x_1 = lean_mk_string_from_bytes("Unit", 4); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_SizeOf_0__Lean_Meta_ignoreField___closed__2() { +static lean_object* _init_l___private_Lean_Meta_SizeOf_0__Lean_Meta_ignoreFieldType___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_Meta_SizeOf_0__Lean_Meta_ignoreField___closed__1; +x_2 = l___private_Lean_Meta_SizeOf_0__Lean_Meta_ignoreFieldType___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_ignoreField(lean_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_Meta_SizeOf_0__Lean_Meta_ignoreFieldType(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; @@ -2109,88 +2110,75 @@ lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_7 = lean_infer_type(x_1, x_2, x_3, x_4, x_5, x_6); +x_7 = lean_whnf(x_1, x_2, x_3, x_4, x_5, x_6); if (lean_obj_tag(x_7) == 0) { -lean_object* x_8; lean_object* x_9; lean_object* x_10; -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); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -lean_inc(x_2); -x_10 = lean_whnf(x_8, x_2, x_3, x_4, x_5, x_9); -if (lean_obj_tag(x_10) == 0) +uint8_t x_8; +x_8 = !lean_is_exclusive(x_7); +if (x_8 == 0) { -uint8_t x_11; -x_11 = !lean_is_exclusive(x_10); +lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_9 = lean_ctor_get(x_7, 0); +x_10 = lean_ctor_get(x_7, 1); +x_11 = l_Lean_Expr_isForall(x_9); if (x_11 == 0) { -lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_12 = lean_ctor_get(x_10, 0); -x_13 = lean_ctor_get(x_10, 1); -x_14 = l_Lean_Expr_isForall(x_12); -if (x_14 == 0) -{ -uint8_t x_15; lean_object* x_16; -lean_dec(x_12); +uint8_t x_12; lean_object* x_13; +lean_dec(x_9); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_15 = 0; -x_16 = lean_box(x_15); -lean_ctor_set(x_10, 0, x_16); -return x_10; +x_12 = 0; +x_13 = lean_box(x_12); +lean_ctor_set(x_7, 0, x_13); +return x_7; } else { -uint8_t x_17; -x_17 = l_Lean_Expr_isArrow(x_12); -if (x_17 == 0) +uint8_t x_14; +x_14 = l_Lean_Expr_isArrow(x_9); +if (x_14 == 0) { -uint8_t x_18; lean_object* x_19; -lean_dec(x_12); +uint8_t x_15; lean_object* x_16; +lean_dec(x_9); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_18 = 1; -x_19 = lean_box(x_18); -lean_ctor_set(x_10, 0, x_19); -return x_10; +x_15 = 1; +x_16 = lean_box(x_15); +lean_ctor_set(x_7, 0, x_16); +return x_7; } else { -lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_20 = l_Lean_Expr_bindingDomain_x21(x_12); -x_21 = l___private_Lean_Meta_SizeOf_0__Lean_Meta_ignoreField___closed__2; -x_22 = l_Lean_Expr_isConstOf(x_20, x_21); -lean_dec(x_20); -if (x_22 == 0) +lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_17 = l_Lean_Expr_bindingDomain_x21(x_9); +x_18 = l___private_Lean_Meta_SizeOf_0__Lean_Meta_ignoreFieldType___closed__2; +x_19 = l_Lean_Expr_isConstOf(x_17, x_18); +lean_dec(x_17); +if (x_19 == 0) { -uint8_t x_23; lean_object* x_24; -lean_dec(x_12); +uint8_t x_20; lean_object* x_21; +lean_dec(x_9); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_23 = 1; -x_24 = lean_box(x_23); -lean_ctor_set(x_10, 0, x_24); -return x_10; +x_20 = 1; +x_21 = lean_box(x_20); +lean_ctor_set(x_7, 0, x_21); +return x_7; } else { -lean_object* x_25; -lean_free_object(x_10); -x_25 = l_Lean_Expr_bindingBody_x21(x_12); -lean_dec(x_12); -x_1 = x_25; -x_6 = x_13; +lean_object* x_22; +lean_free_object(x_7); +x_22 = l_Lean_Expr_bindingBody_x21(x_9); +lean_dec(x_9); +x_1 = x_22; +x_6 = x_10; goto _start; } } @@ -2198,76 +2186,76 @@ goto _start; } else { -lean_object* x_27; lean_object* x_28; uint8_t x_29; -x_27 = lean_ctor_get(x_10, 0); -x_28 = lean_ctor_get(x_10, 1); -lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_10); -x_29 = l_Lean_Expr_isForall(x_27); -if (x_29 == 0) +lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_24 = lean_ctor_get(x_7, 0); +x_25 = lean_ctor_get(x_7, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_7); +x_26 = l_Lean_Expr_isForall(x_24); +if (x_26 == 0) { -uint8_t x_30; lean_object* x_31; lean_object* x_32; -lean_dec(x_27); +uint8_t x_27; lean_object* x_28; lean_object* x_29; +lean_dec(x_24); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -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_28); -return x_32; +x_27 = 0; +x_28 = lean_box(x_27); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_25); +return x_29; } else { -uint8_t x_33; -x_33 = l_Lean_Expr_isArrow(x_27); -if (x_33 == 0) +uint8_t x_30; +x_30 = l_Lean_Expr_isArrow(x_24); +if (x_30 == 0) { -uint8_t x_34; lean_object* x_35; lean_object* x_36; -lean_dec(x_27); +uint8_t x_31; lean_object* x_32; lean_object* x_33; +lean_dec(x_24); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_34 = 1; -x_35 = lean_box(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_28); -return x_36; +x_31 = 1; +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_25); +return x_33; } else { -lean_object* x_37; lean_object* x_38; uint8_t x_39; -x_37 = l_Lean_Expr_bindingDomain_x21(x_27); -x_38 = l___private_Lean_Meta_SizeOf_0__Lean_Meta_ignoreField___closed__2; -x_39 = l_Lean_Expr_isConstOf(x_37, x_38); -lean_dec(x_37); -if (x_39 == 0) +lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_34 = l_Lean_Expr_bindingDomain_x21(x_24); +x_35 = l___private_Lean_Meta_SizeOf_0__Lean_Meta_ignoreFieldType___closed__2; +x_36 = l_Lean_Expr_isConstOf(x_34, x_35); +lean_dec(x_34); +if (x_36 == 0) { -uint8_t x_40; lean_object* x_41; lean_object* x_42; -lean_dec(x_27); +uint8_t x_37; lean_object* x_38; lean_object* x_39; +lean_dec(x_24); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_40 = 1; -x_41 = lean_box(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_28); -return x_42; +x_37 = 1; +x_38 = lean_box(x_37); +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_25); +return x_39; } else { -lean_object* x_43; -x_43 = l_Lean_Expr_bindingBody_x21(x_27); -lean_dec(x_27); -x_1 = x_43; -x_6 = x_28; +lean_object* x_40; +x_40 = l_Lean_Expr_bindingBody_x21(x_24); +lean_dec(x_24); +x_1 = x_40; +x_6 = x_25; goto _start; } } @@ -2276,55 +2264,76 @@ goto _start; } else { -uint8_t x_45; +uint8_t x_42; lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_45 = !lean_is_exclusive(x_10); -if (x_45 == 0) +x_42 = !lean_is_exclusive(x_7); +if (x_42 == 0) { -return x_10; +return x_7; } else { -lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_46 = lean_ctor_get(x_10, 0); -x_47 = lean_ctor_get(x_10, 1); -lean_inc(x_47); -lean_inc(x_46); -lean_dec(x_10); -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_object* x_43; lean_object* x_44; lean_object* x_45; +x_43 = lean_ctor_get(x_7, 0); +x_44 = lean_ctor_get(x_7, 1); +lean_inc(x_44); +lean_inc(x_43); +lean_dec(x_7); +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_EXPORT lean_object* l___private_Lean_Meta_SizeOf_0__Lean_Meta_ignoreField(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_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +x_7 = lean_infer_type(x_1, x_2, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_7) == 0) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; +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___private_Lean_Meta_SizeOf_0__Lean_Meta_ignoreFieldType(x_8, x_2, x_3, x_4, x_5, x_9); +return x_10; +} else { -uint8_t x_49; +uint8_t x_11; lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_49 = !lean_is_exclusive(x_7); -if (x_49 == 0) +x_11 = !lean_is_exclusive(x_7); +if (x_11 == 0) { return x_7; } else { -lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_50 = lean_ctor_get(x_7, 0); -x_51 = lean_ctor_get(x_7, 1); -lean_inc(x_51); -lean_inc(x_50); +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_7, 0); +x_13 = lean_ctor_get(x_7, 1); +lean_inc(x_13); +lean_inc(x_12); lean_dec(x_7); -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; +x_14 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_14, 0, x_12); +lean_ctor_set(x_14, 1, x_13); +return x_14; } } } @@ -2341,7 +2350,7 @@ static lean_object* _init_l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfRecF _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Meta_SizeOf_0__Lean_Meta_ignoreField___closed__1; +x_1 = l___private_Lean_Meta_SizeOf_0__Lean_Meta_ignoreFieldType___closed__1; x_2 = l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfRecFieldFormIH___closed__1; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; @@ -4723,7 +4732,7 @@ static lean_object* _init_l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfMino 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_Meta_SizeOf_0__Lean_Meta_mkSizeOfMinors___rarg___closed__4; x_2 = l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfMinors___rarg___closed__5; -x_3 = lean_unsigned_to_nat(95u); +x_3 = lean_unsigned_to_nat(99u); x_4 = lean_unsigned_to_nat(2u); x_5 = l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfMinors___rarg___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -14901,7 +14910,7 @@ lean_dec(x_1); return x_9; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8411____closed__1() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8426____closed__1() { _start: { lean_object* x_1; @@ -14909,17 +14918,17 @@ x_1 = lean_mk_string_from_bytes("genSizeOf", 9); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8411____closed__2() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8426____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8411____closed__1; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8426____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8411____closed__3() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8426____closed__3() { _start: { lean_object* x_1; @@ -14927,13 +14936,13 @@ x_1 = lean_mk_string_from_bytes("generate `SizeOf` instance for inductive types return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8411____closed__4() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8426____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 = 1; x_2 = l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfMotives_loop___rarg___closed__7; -x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8411____closed__3; +x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8426____closed__3; x_4 = lean_box(x_1); x_5 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_5, 0, x_4); @@ -14942,7 +14951,7 @@ lean_ctor_set(x_5, 2, x_3); return x_5; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8411____closed__5() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8426____closed__5() { _start: { lean_object* x_1; @@ -14950,29 +14959,29 @@ x_1 = lean_mk_string_from_bytes("Lean", 4); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8411____closed__6() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8426____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8411____closed__5; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8426____closed__5; x_2 = l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfMotives_loop___rarg___closed__2; -x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8411____closed__1; +x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8426____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8411_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8426_(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_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8411____closed__2; -x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8411____closed__4; -x_4 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8411____closed__6; -x_5 = l_Lean_Option_register___at_Std_Format_initFn____x40_Lean_Data_Format___hyg_95____spec__1(x_2, x_3, x_4, x_1); +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8426____closed__2; +x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8426____closed__4; +x_4 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8426____closed__6; +x_5 = l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__1(x_2, x_3, x_4, x_1); return x_5; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8450____closed__1() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8465____closed__1() { _start: { lean_object* x_1; @@ -14980,17 +14989,17 @@ x_1 = lean_mk_string_from_bytes("genSizeOfSpec", 13); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8450____closed__2() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8465____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8450____closed__1; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8465____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8450____closed__3() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8465____closed__3() { _start: { lean_object* x_1; @@ -14998,13 +15007,13 @@ x_1 = lean_mk_string_from_bytes("generate `SizeOf` specificiation theorems for a return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8450____closed__4() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8465____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 = 1; x_2 = l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfMotives_loop___rarg___closed__7; -x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8450____closed__3; +x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8465____closed__3; x_4 = lean_box(x_1); x_5 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_5, 0, x_4); @@ -15013,25 +15022,25 @@ lean_ctor_set(x_5, 2, x_3); return x_5; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8450____closed__5() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8465____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8411____closed__5; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8426____closed__5; x_2 = l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfMotives_loop___rarg___closed__2; -x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8450____closed__1; +x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8465____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8450_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8465_(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_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8450____closed__2; -x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8450____closed__4; -x_4 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8450____closed__5; -x_5 = l_Lean_Option_register___at_Std_Format_initFn____x40_Lean_Data_Format___hyg_95____spec__1(x_2, x_3, x_4, x_1); +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8465____closed__2; +x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8465____closed__4; +x_4 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8465____closed__5; +x_5 = l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__1(x_2, x_3, x_4, x_1); return x_5; } } @@ -15962,7 +15971,7 @@ else { lean_object* x_19; uint8_t x_20; x_19 = l_Lean_Meta_mkSizeOfInstances___closed__1; -x_20 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_11, x_19); +x_20 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_11, x_19); if (x_20 == 0) { lean_object* x_21; @@ -16042,7 +16051,7 @@ x_38 = lean_ctor_get(x_36, 1); x_39 = lean_ctor_get(x_36, 0); lean_dec(x_39); x_40 = l_Lean_Meta_mkSizeOfInstances___closed__2; -x_41 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_11, x_40); +x_41 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_11, x_40); lean_dec(x_11); if (x_41 == 0) { @@ -16078,7 +16087,7 @@ x_47 = lean_ctor_get(x_36, 1); lean_inc(x_47); lean_dec(x_36); x_48 = l_Lean_Meta_mkSizeOfInstances___closed__2; -x_49 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_11, x_48); +x_49 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_11, x_48); lean_dec(x_11); if (x_49 == 0) { @@ -16278,7 +16287,7 @@ else { lean_object* x_81; uint8_t x_82; x_81 = l_Lean_Meta_mkSizeOfInstances___closed__1; -x_82 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_11, x_81); +x_82 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_11, x_81); if (x_82 == 0) { lean_object* x_83; lean_object* x_84; @@ -16362,7 +16371,7 @@ if (lean_is_exclusive(x_99)) { x_101 = lean_box(0); } x_102 = l_Lean_Meta_mkSizeOfInstances___closed__2; -x_103 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_11, x_102); +x_103 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_11, x_102); lean_dec(x_11); if (x_103 == 0) { @@ -16616,27 +16625,27 @@ lean_dec(x_7); return x_13; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__1() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8411____closed__5; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8426____closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__2() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__1; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__1; x_2 = l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfMotives_loop___rarg___closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__3() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__3() { _start: { lean_object* x_1; @@ -16644,17 +16653,17 @@ x_1 = lean_mk_string_from_bytes("initFn", 6); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__4() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__2; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__3; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__2; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__3; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__5() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__5() { _start: { lean_object* x_1; @@ -16662,47 +16671,47 @@ x_1 = lean_mk_string_from_bytes("_@", 2); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__6() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__4; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__5; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__4; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__7() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__6; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8411____closed__5; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__6; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8426____closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__8() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__7; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__7; x_2 = l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfMotives_loop___rarg___closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__9() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__8; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__8; x_2 = l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkLocalInstances_loop___rarg___lambda__1___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__10() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__10() { _start: { lean_object* x_1; @@ -16710,33 +16719,33 @@ x_1 = lean_mk_string_from_bytes("_hyg", 4); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__11() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__9; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__10; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__9; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__10; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__12() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__11; -x_2 = lean_unsigned_to_nat(9251u); +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__11; +x_2 = lean_unsigned_to_nat(9266u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; x_2 = l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfMotives_loop___rarg___closed__4; x_3 = 0; -x_4 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__12; +x_4 = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__12; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } @@ -16794,10 +16803,10 @@ l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfMotives_loop___rarg___closed__ lean_mark_persistent(l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfMotives_loop___rarg___closed__7); l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfMotives_loop___rarg___closed__8 = _init_l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfMotives_loop___rarg___closed__8(); lean_mark_persistent(l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfMotives_loop___rarg___closed__8); -l___private_Lean_Meta_SizeOf_0__Lean_Meta_ignoreField___closed__1 = _init_l___private_Lean_Meta_SizeOf_0__Lean_Meta_ignoreField___closed__1(); -lean_mark_persistent(l___private_Lean_Meta_SizeOf_0__Lean_Meta_ignoreField___closed__1); -l___private_Lean_Meta_SizeOf_0__Lean_Meta_ignoreField___closed__2 = _init_l___private_Lean_Meta_SizeOf_0__Lean_Meta_ignoreField___closed__2(); -lean_mark_persistent(l___private_Lean_Meta_SizeOf_0__Lean_Meta_ignoreField___closed__2); +l___private_Lean_Meta_SizeOf_0__Lean_Meta_ignoreFieldType___closed__1 = _init_l___private_Lean_Meta_SizeOf_0__Lean_Meta_ignoreFieldType___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_SizeOf_0__Lean_Meta_ignoreFieldType___closed__1); +l___private_Lean_Meta_SizeOf_0__Lean_Meta_ignoreFieldType___closed__2 = _init_l___private_Lean_Meta_SizeOf_0__Lean_Meta_ignoreFieldType___closed__2(); +lean_mark_persistent(l___private_Lean_Meta_SizeOf_0__Lean_Meta_ignoreFieldType___closed__2); l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfRecFieldFormIH___closed__1 = _init_l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfRecFieldFormIH___closed__1(); lean_mark_persistent(l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfRecFieldFormIH___closed__1); l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfRecFieldFormIH___closed__2 = _init_l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfRecFieldFormIH___closed__2(); @@ -16980,34 +16989,34 @@ l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfSpecTheorem___closed__1 = _ini lean_mark_persistent(l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfSpecTheorem___closed__1); l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfSpecTheorem___closed__2 = _init_l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfSpecTheorem___closed__2(); lean_mark_persistent(l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfSpecTheorem___closed__2); -l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8411____closed__1 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8411____closed__1(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8411____closed__1); -l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8411____closed__2 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8411____closed__2(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8411____closed__2); -l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8411____closed__3 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8411____closed__3(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8411____closed__3); -l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8411____closed__4 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8411____closed__4(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8411____closed__4); -l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8411____closed__5 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8411____closed__5(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8411____closed__5); -l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8411____closed__6 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8411____closed__6(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8411____closed__6); -if (builtin) {res = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8411_(lean_io_mk_world()); +l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8426____closed__1 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8426____closed__1(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8426____closed__1); +l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8426____closed__2 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8426____closed__2(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8426____closed__2); +l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8426____closed__3 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8426____closed__3(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8426____closed__3); +l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8426____closed__4 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8426____closed__4(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8426____closed__4); +l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8426____closed__5 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8426____closed__5(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8426____closed__5); +l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8426____closed__6 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8426____closed__6(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8426____closed__6); +if (builtin) {res = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8426_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Meta_genSizeOf = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Meta_genSizeOf); lean_dec_ref(res); -}l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8450____closed__1 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8450____closed__1(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8450____closed__1); -l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8450____closed__2 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8450____closed__2(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8450____closed__2); -l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8450____closed__3 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8450____closed__3(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8450____closed__3); -l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8450____closed__4 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8450____closed__4(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8450____closed__4); -l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8450____closed__5 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8450____closed__5(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8450____closed__5); -if (builtin) {res = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8450_(lean_io_mk_world()); +}l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8465____closed__1 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8465____closed__1(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8465____closed__1); +l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8465____closed__2 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8465____closed__2(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8465____closed__2); +l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8465____closed__3 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8465____closed__3(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8465____closed__3); +l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8465____closed__4 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8465____closed__4(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8465____closed__4); +l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8465____closed__5 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8465____closed__5(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8465____closed__5); +if (builtin) {res = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_8465_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Meta_genSizeOfSpec = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Meta_genSizeOfSpec); @@ -17036,31 +17045,31 @@ l_Lean_Meta_mkSizeOfInstances___closed__1 = _init_l_Lean_Meta_mkSizeOfInstances_ lean_mark_persistent(l_Lean_Meta_mkSizeOfInstances___closed__1); l_Lean_Meta_mkSizeOfInstances___closed__2 = _init_l_Lean_Meta_mkSizeOfInstances___closed__2(); lean_mark_persistent(l_Lean_Meta_mkSizeOfInstances___closed__2); -l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__1 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__1(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__1); -l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__2 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__2(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__2); -l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__3 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__3(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__3); -l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__4 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__4(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__4); -l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__5 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__5(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__5); -l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__6 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__6(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__6); -l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__7 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__7(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__7); -l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__8 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__8(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__8); -l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__9 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__9(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__9); -l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__10 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__10(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__10); -l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__11 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__11(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__11); -l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__12 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__12(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251____closed__12); -res = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9251_(lean_io_mk_world()); +l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__1 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__1(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__1); +l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__2 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__2(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__2); +l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__3 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__3(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__3); +l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__4 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__4(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__4); +l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__5 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__5(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__5); +l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__6 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__6(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__6); +l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__7 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__7(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__7); +l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__8 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__8(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__8); +l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__9 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__9(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__9); +l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__10 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__10(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__10); +l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__11 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__11(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__11); +l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__12 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__12(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266____closed__12); +res = l_Lean_Meta_initFn____x40_Lean_Meta_SizeOf___hyg_9266_(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/Meta/SynthInstance.c b/stage0/stdlib/Lean/Meta/SynthInstance.c index 1517c233745a..c2e60fc94f3e 100644 --- a/stage0/stdlib/Lean/Meta/SynthInstance.c +++ b/stage0/stdlib/Lean/Meta/SynthInstance.c @@ -22,18 +22,16 @@ static lean_object* l_Lean_Meta_SynthInstance_resume___closed__3; lean_object* l_Lean_throwError___at_Lean_Meta_collectForwardDeps___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); static lean_object* l_Lean_Meta_SynthInstance_MkTableKey_instMonadMCtxM___closed__2; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_88____closed__2; -uint8_t l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Meta_SynthInstance_MkTableKey_instMonadMCtxM___closed__3; static lean_object* l_Lean_Meta_SynthInstance_resume___lambda__1___closed__1; LEAN_EXPORT lean_object* l_List_anyM___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Init_Util_0__outOfBounds___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_tryResolve___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_withTraceNode_x27___at_Lean_Meta_SynthInstance_newSubgoal___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_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_tryResolve___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_tryAnswer___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_SynthInstance_tryResolve___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_tryResolve___spec__1(lean_object*); -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_generate___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_synthInstance_x3f_assignOutParams(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Exception_isMaxHeartbeat(lean_object*); static lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_synthInstance_x3f___spec__7___closed__4; @@ -42,6 +40,8 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_SynthInstance_con static lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_SynthInstance_newSubgoal___spec__11___rarg___closed__3; LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_SynthInstance_newSubgoal___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_Meta_SynthInstance_mkGeneratorNode_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_profiler; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__1; lean_object* lean_mk_empty_array_with_capacity(lean_object*); lean_object* l_Lean_mkAppN(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_isEqvAux___at_Lean_Meta_synthInstance_x3f___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -55,13 +55,11 @@ LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_instInhabitedConsumerNode; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_SynthInstance_getInstances___spec__9___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_instInhabitedInstance; LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_resume___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*); -LEAN_EXPORT lean_object* l_Lean_withSeconds___at_Lean_Meta_SynthInstance_addAnswer___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkFreshLevelMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_consume___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_MVarId_isAssigned___at_Lean_Meta_SynthInstance_consume___spec__1(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_withMCtxImp___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withSeconds___at_Lean_Meta_SynthInstance_generate___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___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_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_generate___spec__1___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*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_SynthInstance_getInstances___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -72,26 +70,27 @@ LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Meta_Synt LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_newSubgoal___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_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*); lean_object* l_Lean_PersistentArray_toArray___rarg(lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__8; lean_object* l_Lean_profileitIOUnsafe___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_resume___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_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_getInstances___spec__7___closed__1; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_synthInstance_x3f___lambda__1___boxed(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_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__4(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_EXPORT lean_object* l_Lean_Meta_SynthInstance_getInstances(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_TableEntry_answers___default; static lean_object* l_Lean_Meta_SynthInstance_resume___closed__1; extern lean_object* l_Lean_maxRecDepthErrorMessage; lean_object* l_Lean_MetavarContext_getDecl(lean_object*, lean_object*); uint8_t lean_usize_dec_le(size_t, size_t); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Meta_SynthInstance_newSubgoal___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*); double lean_float_div(double, double); LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_synth(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_SynthInstance_resume___lambda__1___closed__2; LEAN_EXPORT lean_object* l_StateT_bind___at_Lean_Meta_SynthInstance_MkTableKey_instMonadMCtxM___spec__2___rarg(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__3; uint64_t lean_uint64_mix_hash(uint64_t, uint64_t); LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_MkTableKey_State_lmap___default; -LEAN_EXPORT lean_object* l_Lean_withOptProfile___at_Lean_Meta_SynthInstance_addAnswer___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_newSubgoal___spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_synthInstance_x3f___lambda__6___closed__3; lean_object* l_Lean_Meta_isExprDefEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_insertionSort_swapLoop___at_Lean_Meta_SynthInstance_getInstances___spec__3(lean_object*, lean_object*, lean_object*); @@ -106,13 +105,12 @@ static lean_object* l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__10 lean_object* lean_array_fswap(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_synthInstance_x3f___spec__13___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_EXPORT lean_object* l_Lean_Meta_SynthInstance_instInhabitedSynthM___rarg(lean_object*); -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_synthInstance_x3f___spec__16(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_consume___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_isNewAnswer___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_withOptProfile___at_Lean_Meta_SynthInstance_newSubgoal___spec__12___closed__2; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_synthInstance_x3f___spec__7(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_SynthInstance_getSubgoals___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Level_succ___override(lean_object*); +double l_Lean_trace_profiler_threshold_getSecs(lean_object*); static lean_object* l_Lean_Meta_SynthInstance_resume___lambda__5___closed__2; lean_object* l_Lean_Expr_sort___override(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_trySynthInstance(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -121,22 +119,22 @@ lean_object* lean_array_push(lean_object*, lean_object*); lean_object* l_Lean_mkLevelMax_x27(lean_object*, lean_object*); lean_object* l_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_49_(lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__10; static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_49____closed__5; lean_object* l_Lean_instBEqLocalInstance___boxed(lean_object*, lean_object*); size_t lean_usize_mul(size_t, size_t); static lean_object* l_Lean_Meta_SynthInstance_addAnswer___lambda__1___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_SynthInstance_newSubgoal___spec__11___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___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_mk_array(lean_object*, lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); LEAN_EXPORT lean_object* l_StateT_bind___at_Lean_Meta_SynthInstance_MkTableKey_instMonadMCtxM___spec__2(lean_object*, lean_object*); lean_object* l_Lean_Expr_mdata___override(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__16___rarg(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_SynthInstance_MkTableKey_instMonadMCtxM; -LEAN_EXPORT lean_object* l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__18___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_Meta_SynthInstance_newSubgoal___spec__5(lean_object*, lean_object*); -lean_object* l_Lean_Option_get___at_Std_Format_pretty_x27___spec__1(lean_object*, lean_object*); lean_object* l_Lean_replaceRef(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_proj___override(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_bvar___override(lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); @@ -147,17 +145,19 @@ static lean_object* l_Lean_Meta_synthInstance_x3f_assignOutParams___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_consume___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_expr_has_loose_bvar(lean_object*, lean_object*); static lean_object* l_Lean_Meta_SynthInstance_generate___lambda__1___closed__3; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_addAnswer___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_object*); static lean_object* l_Lean_Meta_SynthInstance_tryResolve___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_mkTableKey___at_Lean_Meta_SynthInstance_main___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_addAnswer___spec__3___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_Expr_fvarId_x21(lean_object*); LEAN_EXPORT lean_object* l_ReaderT_read___at_Lean_Meta_SynthInstance_addAnswer___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_consume___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t lean_float_decLt(double, double); static lean_object* l_Lean_Meta_SynthInstance_tryResolve___lambda__4___closed__1; -LEAN_EXPORT lean_object* l_Lean_withSeconds___at_Lean_Meta_SynthInstance_newSubgoal___spec__14(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessLevels(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_synthInstance_x3f___spec__13___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_EXPORT lean_object* l_Lean_withTraceNode_x27___at_Lean_Meta_SynthInstance_newSubgoal___spec__8___lambda__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_Meta_SynthInstance_main___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_withOptProfile___at_Lean_Meta_SynthInstance_newSubgoal___spec__12___closed__1; 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_SynthInstance_resume___lambda__1___closed__4; static lean_object* l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__7; @@ -171,10 +171,11 @@ static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_6___ static lean_object* l_Lean_Meta_SynthInstance_newSubgoal___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_getMaxHeartbeats___boxed(lean_object*); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___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_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_addAnswer___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_SynthInstance_wakeUp___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__15(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* l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_RecDepth___hyg_6____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_isEqvAux___at_Lean_Meta_synthInstance_x3f___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___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_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_SynthInstance_isNewAnswer___spec__1(lean_object*, lean_object*, size_t, size_t); lean_object* l_Lean_stringToMessageData(lean_object*); @@ -190,6 +191,7 @@ static lean_object* l_Lean_Meta_synthInstance_x3f_assignOutParams___closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_SynthInstance_getInstances___spec__9(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_Array_foldrMUnsafe_fold___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___spec__4(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_withTraceNode___at_Lean_Meta_SynthInstance_tryResolve___spec__2___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*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__12; LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_State_tableEntries___default; LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_mkTableKeyFor___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_PersistentHashMap_insertAux_traverse___at_Lean_Meta_synthInstance_x3f___spec__9(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -197,8 +199,10 @@ extern lean_object* l_Lean_Expr_instBEqExpr; lean_object* l_Lean_Exception_toMessageData(lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_replace___at_Lean_Meta_SynthInstance_newSubgoal___spec__6(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_generate___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_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__4___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_generate___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_instBEqProd___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__9; LEAN_EXPORT lean_object* l_Lean_Core_withCurrHeartbeats___at_Lean_Meta_SynthInstance_main___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_SynthInstance_getEntry___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_MkTableKey_normLevel(lean_object*, lean_object*); @@ -209,15 +213,16 @@ LEAN_EXPORT lean_object* l_Lean_Meta_synthInstance_x3f_assignOutParams___lambda_ static lean_object* l_Lean_Meta_SynthInstance_wakeUp___closed__2; static lean_object* l_Lean_Meta_synthInstance_x3f___lambda__6___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_mkTableKey(lean_object*); +LEAN_EXPORT lean_object* l_Lean_withSeconds___at_Lean_Meta_SynthInstance_generate___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_generate___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static size_t l_Lean_PersistentHashMap_findAux___at_Lean_Meta_synthInstance_x3f___spec__2___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_SynthInstance_newSubgoal___spec__11___rarg___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_getBoolOption___at_Lean_Meta_SynthInstance_newSubgoal___spec__13(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_HashMapImp_find_x3f___at_Lean_Meta_AbstractMVars_abstractExprMVars___spec__2(lean_object*, lean_object*); static lean_object* l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__3; uint8_t l_Lean_Expr_hasMVar(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_getTop(lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_addAnswer___spec__3(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_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_generate___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_tryResolve___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_instHashableArray___rarg___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Meta_synthInstance_x3f___lambda__6___closed__5; @@ -226,6 +231,7 @@ lean_object* l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(lean_object* static lean_object* l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__2; size_t lean_ptr_addr(lean_object*); lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_synthInstance_x3f___spec__13___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___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___lambda__2___closed__4; LEAN_EXPORT lean_object* l_Lean_HashMapImp_expand___at_Lean_Meta_SynthInstance_newSubgoal___spec__3(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_synthInstance_x3f___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -234,31 +240,31 @@ static lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_synthInstance size_t lean_usize_of_nat(lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_6____closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_getNextToResume___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_generate___spec__1___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_Lean_Meta_SynthInstance_generate___lambda__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_addAnswer___spec__3___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_SynthInstance_instInhabitedGeneratorNode___closed__6; static lean_object* l_Lean_Meta_synthInstance_x3f___lambda__7___closed__1; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Meta_SynthInstance_getInstances___spec__8___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_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___closed__2; static lean_object* l_Lean_Meta_SynthInstance_main___lambda__1___closed__5; LEAN_EXPORT lean_object* l_Array_isEqvAux___at_Lean_Meta_synthInstance_x3f___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__12; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__4; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__1(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_SynthInstance___hyg_11245____closed__13; lean_object* l_Lean_throwError___at_Lean_Expr_abstractRangeM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_getEntry(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_synthInstance_x3f___spec__7___closed__2; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__5; LEAN_EXPORT lean_object* l_Lean_Meta_synthInstance_x3f___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355_(lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_processPostponed___spec__2___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessOutParam(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245_(lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__10; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); lean_object* l_Lean_getOutParamPositions_x3f(lean_object*, lean_object*); uint8_t lean_expr_eqv(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__18(lean_object*); static lean_object* l_Lean_Meta_SynthInstance_main___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_tryAnswer___lambda__1___boxed(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_SynthInstance___hyg_88____closed__5; static lean_object* l_Lean_Meta_SynthInstance_resume___lambda__5___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_modifyTop(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_HashMap_insert___at_Lean_Meta_AbstractMVars_abstractExprMVars___spec__4(lean_object*, lean_object*, lean_object*); @@ -269,9 +275,8 @@ LEAN_EXPORT lean_object* l_List_forIn_loop___at___private_Lean_Meta_SynthInstanc static lean_object* l_Lean_Meta_SynthInstance_newSubgoal___lambda__2___closed__4; static lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__3___closed__2; static lean_object* l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__8; +lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__6; -LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Meta_synthInstance_x3f___spec__17___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_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__13; lean_object* l_Lean_HashMap_insert___at___private_Lean_Meta_AbstractMVars_0__Lean_Meta_AbstractMVars_abstractLevelMVars___spec__3(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_SynthInstance_newSubgoal___lambda__2___closed__2; LEAN_EXPORT uint8_t l_Lean_Meta_SynthInstance_Waiter_isRoot(lean_object*); @@ -281,12 +286,14 @@ LEAN_EXPORT lean_object* l_Lean_withTraceNode_x27___at_Lean_Meta_SynthInstance_n static lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessLevels___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_MkTableKey_instMonadMCtxM___lambda__1___boxed(lean_object*, lean_object*); lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDeclImp___rarg(lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_addAnswer___spec__5___boxed(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_SynthInstance___hyg_6____closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_mkTableKey___at_Lean_Meta_SynthInstance_mkTableKeyFor___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_generate___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_Lean_Meta_SynthInstance_instInhabitedGeneratorNode___closed__5; LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_resume___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_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_88____closed__3; +lean_object* l_Lean_PersistentArray_append___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Meta_synthInstance_x3f___spec__16___rarg(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_SynthInstance_addAnswer___spec__2(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_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_SynthInstance_newSubgoal___spec__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -294,8 +301,8 @@ LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_tryAnswer___lambda__2___boxed lean_object* l_Lean_Meta_check(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_SynthInstance_addAnswer___lambda__1___closed__1; static lean_object* l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__7; +LEAN_EXPORT lean_object* l_Lean_withSeconds___at_Lean_Meta_SynthInstance_addAnswer___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___lambda__2___closed__5; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__7; static lean_object* l_Lean_Meta_SynthInstance_generate___lambda__4___closed__1; static lean_object* l_Lean_Meta_SynthInstance_getSubgoals___closed__1; static lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___lambda__2___closed__6; @@ -310,17 +317,18 @@ static lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstan LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_mkTableKey___rarg___lambda__2(lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_State_generatorStack___default; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_88____closed__4; static lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_synthInstance_x3f___spec__8___closed__1; lean_object* lean_array_pop(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Meta_synthInstance_x3f___spec__16(lean_object*); LEAN_EXPORT lean_object* lean_synth_pending(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__4; static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_49____closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_addAnswer___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_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_getInstances___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_array_to_list(lean_object*, lean_object*); -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_addAnswer___spec__3___lambda__3(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* l_Lean_Meta_abstractMVars(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static size_t l_Lean_PersistentHashMap_findAux___at_Lean_Meta_synthInstance_x3f___spec__2___closed__2; lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withMVarContextImp___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -337,15 +345,14 @@ LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_mapMetaM(lean_object*, lean_o LEAN_EXPORT uint8_t l_Array_isEqvAux___at_Lean_Meta_synthInstance_x3f___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_instInhabitedAnswer; LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocess(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_newSubgoal___spec__17(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_getResetTraces___at_Lean_Meta_SynthInstance_newSubgoal___spec__11___rarg___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_SynthInstance_newSubgoal___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Expr_instHashableExpr; LEAN_EXPORT lean_object* l_List_forIn_loop___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessLevels___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__11; static lean_object* l_Lean_Meta_SynthInstance_tryResolve___lambda__1___closed__3; LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Meta_SynthInstance_getInstances___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_levelZero; +static double l_Lean_withSeconds___at_Lean_Meta_SynthInstance_newSubgoal___spec__12___closed__1; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_generate___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_EXPORT lean_object* l_Lean_Meta_SynthInstance_MkTableKey_State_nextIdx___default; lean_object* lean_io_mono_nanos_now(lean_object*); @@ -354,15 +361,14 @@ LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_tryResolve___lambda__3___boxe static lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_SynthInstance_newSubgoal___spec__11___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at_Lean_Meta_SynthInstance_mkTableKeyFor___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Level_hasMVar(lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__6; extern lean_object* l_Lean_instInhabitedExpr; -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___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_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__13___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_SynthInstance_getNextToResume___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_synthInstance_x3f___spec__7___closed__1; LEAN_EXPORT uint8_t l_Lean_AssocList_contains___at_Lean_Meta_SynthInstance_newSubgoal___spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___lambda__2___closed__4; lean_object* l_Lean_Expr_constName_x3f(lean_object*); -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_tryResolve___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__6; static lean_object* l_Lean_Meta_SynthInstance_resume___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_synthInstance_x3f___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_tryResolve___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -376,15 +382,13 @@ static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_getIn LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_SynthInstance_getInstances___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_SynthInstance_consume___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_Lean_Meta_SynthInstance_instInhabitedConsumerNode___closed__1; -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_generate___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_getBoolOption___at_Lean_Meta_processPostponed___spec__4(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__5; LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_generate___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_instHashableLocalInstance___boxed(lean_object*); static lean_object* l_Lean_Meta_SynthInstance_getEntry___closed__3; lean_object* l_Lean_Expr_letE___override(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); lean_object* l_instHashableProd___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___closed__3; uint8_t lean_name_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_synthInstance_x3f___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); @@ -397,9 +401,9 @@ LEAN_EXPORT lean_object* l_List_filterAuxM___at_Lean_Meta_SynthInstance_consume_ LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9(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_Meta_getErasedInstances___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Meta_SynthInstance_instInhabitedInstance___closed__3; -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_newSubgoal___spec__17___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_synthInstance_x3f___lambda__6___closed__9; static lean_object* l_Lean_Meta_SynthInstance_instInhabitedInstance___closed__1; +uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_6____closed__7; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_synthInstance_x3f___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessArgs(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -411,18 +415,18 @@ lean_object* l_Lean_Meta_forallTelescope___at___private_Lean_Meta_InferType_0__L LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_Meta_SynthInstance_State_tableEntries___default___spec__1(lean_object*); static lean_object* l_Lean_Meta_SynthInstance_mkTableKeyFor___closed__1; lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_synthInstance_x3f___spec__13___lambda__3(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* l_Lean_simpLevelIMax_x27(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___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_Meta_SynthInstance_State_resumeStack___default; -LEAN_EXPORT lean_object* l_Lean_getBoolOption___at_Lean_Meta_SynthInstance_newSubgoal___spec__13___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_SynthInstance_instInhabitedSynthM___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___closed__4; +LEAN_EXPORT lean_object* l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__16(lean_object*); LEAN_EXPORT uint8_t l_Array_isEqvAux___at_Lean_Meta_synthInstance_x3f___spec__12(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___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_Meta_synthInstance_x3f___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_tryResolve___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_Meta_SynthInstance_checkMaxHeartbeats(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__9; +static lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__3; 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___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessArgs___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_SynthInstance_newSubgoal___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -431,7 +435,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_MkTableKey_State_emap___defau LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_SynthInstance_getInstances___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___lambda__2___closed__3; static lean_object* l_Lean_Meta_SynthInstance_newSubgoal___lambda__2___closed__3; -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_addAnswer___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_SynthInstance_tryResolve___lambda__5___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_getMaxHeartbeats(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_synthInstance_x3f___spec__2(lean_object*, size_t, lean_object*); @@ -439,6 +442,7 @@ LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Meta_SynthInstance_isN LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_synthInstance_maxSize; LEAN_EXPORT lean_object* l_List_anyM___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withSeconds___at_Lean_Meta_synthInstance_x3f___spec__14(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_generate(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_wakeUp___boxed(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_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__3___closed__1; @@ -447,9 +451,11 @@ LEAN_EXPORT lean_object* l_Lean_Meta_synthInstance_x3f___lambda__5___boxed(lean_ LEAN_EXPORT lean_object* l_Lean_Meta_synthInstance_x3f___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_mkTableKeyFor___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__2___closed__5; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__3; LEAN_EXPORT lean_object* l_List_filterAuxM___at_Lean_Meta_SynthInstance_consume___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_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_6____closed__6; static lean_object* l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__6; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_tryResolve___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*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_arrayExpr_toMessageData(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_SynthInstance_getEntry___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_SynthInstance_getSubgoals___spec__2(lean_object*, size_t, size_t, lean_object*); @@ -458,7 +464,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_mkTableKey___rarg(lean_object lean_object* l_Lean_Meta_DiscrTree_getUnify___rarg(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_SynthInstance_instInhabitedAnswer___closed__1; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_tryResolve___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withSeconds___at_Lean_Meta_SynthInstance_tryResolve___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_usize_to_nat(size_t); size_t lean_data_hashmap_mk_idx(lean_object*, uint64_t); @@ -469,19 +474,20 @@ LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_findEntry_x3f(lean_object*, l static lean_object* l_Lean_Meta_SynthInstance_checkMaxHeartbeats___closed__1; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___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*); LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_newSubgoal(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_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_tryResolve___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_mkTableKey___at_Lean_Meta_SynthInstance_mkTableKeyFor___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_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__2___closed__6; -lean_object* l_Lean_Option_register___at_Std_Format_initFn____x40_Lean_Data_Format___hyg_95____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_synthInstance_x3f___spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_SynthInstance_instInhabitedAnswer___closed__2; extern lean_object* l_Lean_inheritedTraceOptions; lean_object* l_Lean_MessageData_ofExpr(lean_object*); +LEAN_EXPORT lean_object* l_Lean_withSeconds___at_Lean_Meta_SynthInstance_tryResolve___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_checkMaxHeartbeats___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_newSubgoal___spec__15(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_addAnswer___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_initFn____x40_Lean_Meta_SynthInstance___hyg_88____closed__1; static lean_object* l_Lean_Meta_trySynthInstance___closed__1; LEAN_EXPORT lean_object* l_ReaderT_read___at_Lean_Meta_SynthInstance_addAnswer___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_SynthInstance_resume___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_tryResolve___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_SynthInstance_instInhabitedGeneratorNode___closed__4; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_synthInstance_x3f___spec__8(lean_object*, size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -492,8 +498,8 @@ static lean_object* l_Lean_Meta_SynthInstance_newSubgoal___lambda__2___closed__1 extern lean_object* l_Lean_crossEmoji; LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_instInhabitedGeneratorNode; static lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___lambda__2___closed__3; -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_synthInstance_x3f___spec__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_generate___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* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_processPostponed___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_HashMapImp_find_x3f___at___private_Lean_Meta_AbstractMVars_0__Lean_Meta_AbstractMVars_abstractLevelMVars___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_MkTableKey_instMonadMCtxM___lambda__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -506,12 +512,9 @@ static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_6___ uint8_t l_Lean_Expr_hasAnyFVar_visit___at_Lean_Expr_containsFVar___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_SynthInstance_getInstances___spec__9___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_isEqvAux___at_Lean_Meta_synthInstance_x3f___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_getNextToResume___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___lambda__2___closed__2; -static double l_Lean_withSeconds___at_Lean_Meta_SynthInstance_newSubgoal___spec__14___closed__1; LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Meta_SynthInstance_wakeUp___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_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_addAnswer___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_synthInstance_x3f_assignOutParams___closed__5; LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_getTop___boxed(lean_object*); @@ -519,16 +522,17 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_SynthInstance_g LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_mkTableKey___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_app___override(lean_object*, lean_object*); static lean_object* l_Lean_Meta_SynthInstance_instInhabitedSynthM___rarg___closed__2; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_tryResolve___spec__2___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_EXPORT lean_object* l_Lean_Meta_SynthInstance_tryResolve___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_Meta_SynthInstance_MkTableKey_State_lmap___default___closed__1; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__2; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); lean_object* l_Lean_MVarId_getDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMap_insert___at_Lean_Meta_SynthInstance_newSubgoal___spec__1(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__11; lean_object* l_Lean_Meta_getFVarLocalDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at_Lean_Meta_SynthInstance_findEntry_x3f___spec__2___boxed(lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Meta_synthInstance_x3f___spec__17___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_synthInstance_x3f___spec__15(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_resume___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_EXPORT lean_object* l_Lean_MVarId_isAssigned___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_mkTableKeyFor___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -537,20 +541,20 @@ static lean_object* l_Lean_Meta_synthInstance_x3f___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_synthInstance_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_float_to_string(double); static lean_object* l_Lean_Meta_SynthInstance_main___lambda__1___closed__4; +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_tryResolve___spec__2___lambda__4(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__12; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_synthInstance_x3f___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_88_(lean_object*); static lean_object* l_Lean_Meta_SynthInstance_instInhabitedInstance___closed__2; LEAN_EXPORT lean_object* l_Lean_MVarId_isAssignable___at_Lean_Meta_SynthInstance_MkTableKey_normExpr___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Meta_synthInstance_x3f___lambda__6___closed__4; -LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Meta_synthInstance_x3f___spec__17(lean_object*); lean_object* l_Lean_Meta_getLocalInstances(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_SynthInstance_resume___lambda__1___closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_main(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_generate___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_withOptProfile___at_Lean_Meta_SynthInstance_generate___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l___private_Lean_Expr_0__Lean_beqBinderInfo____x40_Lean_Expr___hyg_473_(uint8_t, uint8_t); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Meta_SynthInstance_newSubgoal___spec__14(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_Lean_Meta_SynthInstance_addAnswer___closed__1; static lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__2___closed__3; static lean_object* l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__11; @@ -564,17 +568,17 @@ static lean_object* l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed_ static lean_object* l_Lean_Meta_SynthInstance_MkTableKey_instMonadMCtxM___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_modifyTop___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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*); -static lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_consume___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_HashMapImp_find_x3f___at_Lean_Meta_SynthInstance_findEntry_x3f___spec__1(lean_object*, lean_object*); -uint8_t l_Lean_KVMap_getBool(lean_object*, lean_object*, uint8_t); lean_object* l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_ptrEqList___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_SynthInstance_getInstances___spec__9___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_tryResolve___spec__2___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__7; lean_object* lean_panic_fn(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_generate___spec__1___lambda__3(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*); static lean_object* l_Lean_Meta_SynthInstance_addAnswer___lambda__1___closed__2; LEAN_EXPORT uint8_t l_Array_isEqvAux___at_Lean_Meta_synthInstance_x3f___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withOptProfile___at_Lean_Meta_SynthInstance_newSubgoal___spec__12(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_State_result_x3f___default; static lean_object* l_panic___at_Lean_Meta_SynthInstance_getEntry___spec__1___closed__1; static lean_object* l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__1; @@ -587,6 +591,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_step(lean_object*, lean_objec lean_object* lean_nat_mul(lean_object*, lean_object*); static lean_object* l_Lean_Meta_SynthInstance_resume___lambda__5___closed__1; static lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessArgs___closed__1; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_synthInstance_x3f___spec__13___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_synthInstance_x3f___spec__11(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_main___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -594,15 +599,14 @@ static lean_object* l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__1; static lean_object* l_Lean_Meta_SynthInstance_tryResolve___lambda__5___closed__1; lean_object* l_Lean_PersistentHashMap_mkCollisionNode___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_addAnswer___lambda__4(lean_object*, 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_SynthInstance___hyg_11245____closed__8; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_tryResolve___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*); static lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___lambda__2___closed__1; static lean_object* l_Lean_Meta_synthInstance_x3f_assignOutParams___closed__2; lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Meta_SynthInstance_newSubgoal___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*); uint8_t l_Lean_LocalDecl_binderInfo(lean_object*); static lean_object* l_Lean_Meta_synthInstance_x3f___lambda__6___closed__2; LEAN_EXPORT lean_object* l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Meta_SynthInstance_newSubgoal___spec__16(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* l_Lean_mkLevelIMax_x27(lean_object*, lean_object*); lean_object* l_Array_back___rarg(lean_object*, lean_object*); lean_object* lean_expr_instantiate_rev(lean_object*, lean_object*); @@ -633,9 +637,11 @@ LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Meta_SynthInstance_MkTableKe static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_49____closed__2; static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_6____closed__8; LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_SynthInstance_getEntry___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_generate___spec__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_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_contains___at_Lean_Meta_SynthInstance_newSubgoal___spec__2___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Meta_SynthInstance_newSubgoal___lambda__2___closed__5; static lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessOutParam___lambda__1___closed__1; +static lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__1; lean_object* l_Lean_Expr_betaRev(lean_object*, lean_object*, uint8_t, uint8_t); LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_6_(lean_object*); static lean_object* l_Lean_Meta_SynthInstance_instInhabitedGeneratorNode___closed__3; @@ -643,14 +649,15 @@ LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_tryAnswer(lean_object*, lean_ LEAN_EXPORT lean_object* l_Array_filterMapM___at_Lean_Meta_SynthInstance_getInstances___spec__6___boxed(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_object* l_Lean_Expr_fvar___override(lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessOutParam___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_trace_profiler; static lean_object* l_Lean_Meta_synthInstance_x3f___lambda__6___closed__6; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_SynthInstance_getInstances___spec__1___boxed(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_object* l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_getNextToResume(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_hasUnusedArguments___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_synthInstance_etaExperiment; LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_consume___lambda__1___boxed(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); static lean_object* l_Lean_Meta_SynthInstance_getMaxHeartbeats___closed__1; @@ -660,12 +667,12 @@ lean_object* l_Lean_Meta_openAbstractMVarsResult(lean_object*, lean_object*, lea lean_object* l_List_redLength___rarg(lean_object*); lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_Meta_SynthInstance_State_tableEntries___default___spec__1___boxed(lean_object*); -static lean_object* l_Lean_Meta_synthInstance_x3f___lambda__6___closed__10; static lean_object* l_Lean_Meta_SynthInstance_generate___lambda__4___closed__2; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_synthInstance_x3f___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_SynthInstance_MkTableKey_normLevel___closed__1; uint8_t l_Lean_PersistentHashMap_contains___at_Lean_NameSSet_contains___spec__3(lean_object*, lean_object*); lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_generate___spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__4; LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_SynthInstance_newSubgoal___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); @@ -684,7 +691,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Synt uint8_t lean_usize_dec_lt(size_t, size_t); uint8_t l_Array_contains___at___private_Lean_Meta_FunInfo_0__Lean_Meta_collectDeps_visit___spec__2(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_isEqvAux___at_Lean_Meta_synthInstance_x3f___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withOptProfile___at_Lean_Meta_synthInstance_x3f___spec__14(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_synthInstance_x3f___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_synthInstance_x3f___spec__7___closed__5; 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*); @@ -706,11 +713,11 @@ LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_mkTableKey___rarg___lambda__1 LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_mvarId_x21(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_MkTableKey_instMonadMCtxM___lambda__1(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__5; static lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_SynthInstance_newSubgoal___spec__10___closed__1; -LEAN_EXPORT lean_object* l_Lean_withSeconds___at_Lean_Meta_synthInstance_x3f___spec__15(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withSeconds___at_Lean_Meta_SynthInstance_newSubgoal___spec__12(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_resume___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_Meta_SynthInstance_addAnswer___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_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_generate___spec__3___boxed(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_SynthInstance_getResult(lean_object*); lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); @@ -718,12 +725,14 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_SynthInstance_g lean_object* lean_expr_instantiate1(lean_object*, lean_object*); static lean_object* l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_mkTableKeyFor___lambda__1(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_addTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__13(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_Lean_Meta_SynthInstance_tryResolve___lambda__1___closed__4; static lean_object* l_Lean_Meta_SynthInstance_instInhabitedGeneratorNode___closed__7; LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_tryAnswer___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_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___lambda__2___closed__5; LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_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* l_Nat_repr(lean_object*); +LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Meta_synthInstance_x3f___spec__16___rarg___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_withTraceNode___at_Lean_Meta_SynthInstance_addAnswer___spec__3___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*); LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Meta_SynthInstance_getInstances___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_getResult___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -732,13 +741,11 @@ lean_object* l_Option_toLOption___rarg(lean_object*); size_t lean_usize_land(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_tryResolve(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_SynthInstance_MkTableKey_normLevel___closed__2; -LEAN_EXPORT lean_object* l_Lean_withOptProfile___at_Lean_Meta_SynthInstance_tryResolve___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_toArrayAux___rarg(lean_object*, lean_object*); LEAN_EXPORT uint64_t l_Array_foldlMUnsafe_fold___at_Lean_Meta_synthInstance_x3f___spec__6(lean_object*, size_t, size_t, uint64_t); LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_resume(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Data_HashMap_0__Lean_numBucketsForCapacity(lean_object*); static lean_object* l_Lean_Meta_SynthInstance_consume___closed__1; -lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_processPostponed___spec__6(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_Lean_Meta_SynthInstance_addAnswer___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_throwError___at_Lean_Meta_SynthInstance_main___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessArgs___closed__2; @@ -903,70 +910,6 @@ x_5 = l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_RecDepth___hyg_6_ return x_5; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_88____closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("etaExperiment", 13); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_88____closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_6____closed__1; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_88____closed__1; -x_3 = l_Lean_Name_mkStr2(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_88____closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("[DO NOT USE EXCEPT FOR TESTING] enable structure eta for type-classes during type-class search", 94); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_88____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_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_6____closed__4; -x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_88____closed__3; -x_4 = lean_box(x_1); -x_5 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_5, 0, x_4); -lean_ctor_set(x_5, 1, x_2); -lean_ctor_set(x_5, 2, x_3); -return x_5; -} -} -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_88____closed__5() { -_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_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_6____closed__7; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_6____closed__8; -x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_6____closed__1; -x_4 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_88____closed__1; -x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); -return x_5; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_88_(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_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_88____closed__2; -x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_88____closed__4; -x_4 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_88____closed__5; -x_5 = l_Lean_Option_register___at_Std_Format_initFn____x40_Lean_Data_Format___hyg_95____spec__1(x_2, x_3, x_4, x_1); -return x_5; -} -} static lean_object* _init_l_Lean_Meta_SynthInstance_getMaxHeartbeats___closed__1() { _start: { @@ -980,7 +923,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_getMaxHeartbeats(lean_object* { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_2 = l_Lean_Meta_SynthInstance_getMaxHeartbeats___closed__1; -x_3 = l_Lean_Option_get___at_Std_Format_pretty_x27___spec__1(x_1, x_2); +x_3 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_2); x_4 = lean_unsigned_to_nat(1000u); x_5 = lean_nat_mul(x_3, x_4); lean_dec(x_3); @@ -3414,7 +3357,7 @@ static lean_object* _init_l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance 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_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_getInstances___spec__7___closed__1; x_2 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_getInstances___spec__7___closed__2; -x_3 = lean_unsigned_to_nat(216u); +x_3 = lean_unsigned_to_nat(211u); x_4 = lean_unsigned_to_nat(15u); x_5 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_getInstances___spec__7___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -4856,7 +4799,7 @@ lean_inc(x_9); x_10 = lean_ctor_get(x_8, 1); lean_inc(x_10); lean_dec(x_8); -x_11 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_9, x_3, x_4, x_5, x_6, x_10); +x_11 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_9, x_3, x_4, x_5, x_6, x_10); x_12 = lean_ctor_get(x_11, 0); lean_inc(x_12); x_13 = lean_ctor_get(x_11, 1); @@ -5618,20 +5561,7 @@ x_6 = lean_alloc_closure((void*)(l___private_Lean_Util_Trace_0__Lean_getResetTra return x_6; } } -LEAN_EXPORT lean_object* l_Lean_getBoolOption___at_Lean_Meta_SynthInstance_newSubgoal___spec__13(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: -{ -lean_object* x_10; uint8_t x_11; lean_object* x_12; lean_object* x_13; -x_10 = lean_ctor_get(x_7, 2); -x_11 = l_Lean_KVMap_getBool(x_10, x_1, x_2); -x_12 = lean_box(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_9); -return x_13; -} -} -static double _init_l_Lean_withSeconds___at_Lean_Meta_SynthInstance_newSubgoal___spec__14___closed__1() { +static double _init_l_Lean_withSeconds___at_Lean_Meta_SynthInstance_newSubgoal___spec__12___closed__1() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; double x_4; @@ -5642,7 +5572,7 @@ x_4 = l_Float_ofScientific(x_1, x_2, x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_withSeconds___at_Lean_Meta_SynthInstance_newSubgoal___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_EXPORT lean_object* l_Lean_withSeconds___at_Lean_Meta_SynthInstance_newSubgoal___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) { _start: { lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; @@ -5674,7 +5604,7 @@ x_19 = 0; x_20 = lean_unsigned_to_nat(0u); x_21 = l_Float_ofScientific(x_18, x_19, x_20); lean_dec(x_18); -x_22 = l_Lean_withSeconds___at_Lean_Meta_SynthInstance_newSubgoal___spec__14___closed__1; +x_22 = l_Lean_withSeconds___at_Lean_Meta_SynthInstance_newSubgoal___spec__12___closed__1; x_23 = lean_float_div(x_21, x_22); x_24 = lean_box_float(x_23); x_25 = lean_alloc_ctor(0, 2, 0); @@ -5698,7 +5628,7 @@ x_29 = 0; x_30 = lean_unsigned_to_nat(0u); x_31 = l_Float_ofScientific(x_28, x_29, x_30); lean_dec(x_28); -x_32 = l_Lean_withSeconds___at_Lean_Meta_SynthInstance_newSubgoal___spec__14___closed__1; +x_32 = l_Lean_withSeconds___at_Lean_Meta_SynthInstance_newSubgoal___spec__12___closed__1; x_33 = lean_float_div(x_31, x_32); x_34 = lean_box_float(x_33); x_35 = lean_alloc_ctor(0, 2, 0); @@ -5735,201 +5665,7 @@ return x_40; } } } -static lean_object* _init_l_Lean_withOptProfile___at_Lean_Meta_SynthInstance_newSubgoal___spec__12___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("profiler", 8); -return x_1; -} -} -static lean_object* _init_l_Lean_withOptProfile___at_Lean_Meta_SynthInstance_newSubgoal___spec__12___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_withOptProfile___at_Lean_Meta_SynthInstance_newSubgoal___spec__12___closed__1; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Lean_withOptProfile___at_Lean_Meta_SynthInstance_newSubgoal___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) { -_start: -{ -lean_object* x_9; uint8_t x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_9 = l_Lean_withOptProfile___at_Lean_Meta_SynthInstance_newSubgoal___spec__12___closed__2; -x_10 = 0; -x_11 = l_Lean_getBoolOption___at_Lean_Meta_SynthInstance_newSubgoal___spec__13(x_9, x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_unbox(x_12); -lean_dec(x_12); -if (x_13 == 0) -{ -lean_object* x_14; lean_object* x_15; -x_14 = lean_ctor_get(x_11, 1); -lean_inc(x_14); -lean_dec(x_11); -x_15 = lean_apply_7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_14); -if (lean_obj_tag(x_15) == 0) -{ -uint8_t x_16; -x_16 = !lean_is_exclusive(x_15); -if (x_16 == 0) -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_17 = lean_ctor_get(x_15, 0); -x_18 = lean_box(0); -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_17); -lean_ctor_set(x_19, 1, x_18); -lean_ctor_set(x_15, 0, x_19); -return x_15; -} -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_box(0); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_20); -lean_ctor_set(x_23, 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_21); -return x_24; -} -} -else -{ -uint8_t x_25; -x_25 = !lean_is_exclusive(x_15); -if (x_25 == 0) -{ -return x_15; -} -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_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; -} -} -} -else -{ -lean_object* x_29; lean_object* x_30; -x_29 = lean_ctor_get(x_11, 1); -lean_inc(x_29); -lean_dec(x_11); -x_30 = l_Lean_withSeconds___at_Lean_Meta_SynthInstance_newSubgoal___spec__14(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_29); -if (lean_obj_tag(x_30) == 0) -{ -uint8_t x_31; -x_31 = !lean_is_exclusive(x_30); -if (x_31 == 0) -{ -lean_object* x_32; uint8_t x_33; -x_32 = lean_ctor_get(x_30, 0); -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, 1); -x_35 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_35, 0, x_34); -lean_ctor_set(x_32, 1, x_35); -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 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_38, 0, x_37); -x_39 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_39, 0, x_36); -lean_ctor_set(x_39, 1, x_38); -lean_ctor_set(x_30, 0, x_39); -return x_30; -} -} -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; -x_40 = lean_ctor_get(x_30, 0); -x_41 = lean_ctor_get(x_30, 1); -lean_inc(x_41); -lean_inc(x_40); -lean_dec(x_30); -x_42 = lean_ctor_get(x_40, 0); -lean_inc(x_42); -x_43 = lean_ctor_get(x_40, 1); -lean_inc(x_43); -if (lean_is_exclusive(x_40)) { - lean_ctor_release(x_40, 0); - lean_ctor_release(x_40, 1); - x_44 = x_40; -} else { - lean_dec_ref(x_40); - x_44 = lean_box(0); -} -x_45 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_45, 0, x_43); -if (lean_is_scalar(x_44)) { - x_46 = lean_alloc_ctor(0, 2, 0); -} else { - x_46 = x_44; -} -lean_ctor_set(x_46, 0, x_42); -lean_ctor_set(x_46, 1, 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_41); -return x_47; -} -} -else -{ -uint8_t x_48; -x_48 = !lean_is_exclusive(x_30); -if (x_48 == 0) -{ -return x_30; -} -else -{ -lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_49 = lean_ctor_get(x_30, 0); -x_50 = lean_ctor_get(x_30, 1); -lean_inc(x_50); -lean_inc(x_49); -lean_dec(x_30); -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_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Meta_SynthInstance_newSubgoal___spec__16(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) { +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Meta_SynthInstance_newSubgoal___spec__14(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; uint8_t x_16; @@ -6048,7 +5784,7 @@ return x_53; } } } -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__15(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) { +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__13(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: { uint8_t x_13; @@ -6066,7 +5802,7 @@ lean_inc(x_17); x_18 = lean_ctor_get(x_16, 1); lean_inc(x_18); lean_dec(x_16); -x_19 = l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Meta_SynthInstance_newSubgoal___spec__16(x_1, x_2, x_3, x_17, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_18); +x_19 = l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Meta_SynthInstance_newSubgoal___spec__14(x_1, x_2, x_3, x_17, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_18); lean_dec(x_10); return x_19; } @@ -6116,13 +5852,13 @@ lean_inc(x_34); x_35 = lean_ctor_get(x_33, 1); lean_inc(x_35); lean_dec(x_33); -x_36 = l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Meta_SynthInstance_newSubgoal___spec__16(x_1, x_2, x_3, x_34, x_5, x_6, x_7, x_8, x_9, x_32, x_11, x_35); +x_36 = l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Meta_SynthInstance_newSubgoal___spec__14(x_1, x_2, x_3, x_34, x_5, x_6, x_7, x_8, x_9, x_32, x_11, x_35); lean_dec(x_32); return x_36; } } } -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_newSubgoal___spec__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_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_newSubgoal___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) { _start: { if (lean_obj_tag(x_1) == 0) @@ -6218,17 +5954,25 @@ LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSu { lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_inc(x_12); -x_15 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__15(x_1, x_2, x_3, x_6, x_4, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +x_15 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__13(x_1, x_2, x_3, x_6, x_4, x_8, x_9, x_10, x_11, x_12, x_13, x_14); x_16 = lean_ctor_get(x_15, 1); lean_inc(x_16); lean_dec(x_15); -x_17 = l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_newSubgoal___spec__17(x_5, x_8, x_9, x_10, x_11, x_12, x_13, x_16); +x_17 = l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_newSubgoal___spec__15(x_5, x_8, x_9, x_10, x_11, x_12, x_13, x_16); lean_dec(x_12); lean_dec(x_8); return x_17; } } -static lean_object* _init_l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___closed__1() { +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_profiler; +return x_1; +} +} +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__2() { _start: { lean_object* x_1; @@ -6236,16 +5980,16 @@ x_1 = lean_mk_string_from_bytes("[", 1); return x_1; } } -static lean_object* _init_l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___closed__2() { +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__3() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___closed__1; +x_1 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__2; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___closed__3() { +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__4() { _start: { lean_object* x_1; @@ -6253,169 +5997,148 @@ x_1 = lean_mk_string_from_bytes("s] ", 3); return x_1; } } -static lean_object* _init_l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___closed__4() { +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___closed__3; +x_1 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__4; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9(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_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, uint8_t x_7, double 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_12; lean_object* x_13; uint8_t x_14; -lean_inc(x_1); -x_12 = l_Lean_isTracingEnabledFor___at_Lean_Meta_SynthInstance_newSubgoal___spec__10(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -x_13 = lean_ctor_get(x_12, 0); +lean_object* x_17; lean_object* x_18; +lean_dec(x_9); +x_17 = lean_ctor_get(x_14, 5); +lean_inc(x_17); +lean_inc(x_15); +lean_inc(x_14); lean_inc(x_13); -x_14 = lean_unbox(x_13); -lean_dec(x_13); -if (x_14 == 0) +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_2); +x_18 = lean_apply_8(x_1, x_2, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_18) == 0) { -lean_object* x_15; lean_object* x_16; -lean_dec(x_2); -lean_dec(x_1); -x_15 = lean_ctor_get(x_12, 1); -lean_inc(x_15); +lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +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_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__1; +x_22 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_6, x_21); +lean_dec(x_6); +if (x_22 == 0) +{ +if (x_7 == 0) +{ +lean_object* x_23; lean_object* x_24; +x_23 = lean_box(0); +x_24 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__2(x_3, x_4, x_17, x_5, x_2, x_19, x_23, x_10, x_11, x_12, x_13, x_14, x_15, x_20); +lean_dec(x_15); +lean_dec(x_13); lean_dec(x_12); -x_16 = lean_apply_7(x_3, x_5, x_6, x_7, x_8, x_9, x_10, x_15); -return x_16; +lean_dec(x_11); +lean_dec(x_2); +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; -x_17 = lean_ctor_get(x_12, 1); -lean_inc(x_17); -lean_dec(x_12); -x_18 = lean_ctor_get(x_9, 5); -lean_inc(x_18); -x_19 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_SynthInstance_newSubgoal___spec__11___rarg(x_10, x_17); -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_alloc_closure((void*)(l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__1), 8, 1); -lean_closure_set(x_22, 0, x_3); -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_23 = l_Lean_withOptProfile___at_Lean_Meta_SynthInstance_newSubgoal___spec__12(x_22, x_5, x_6, x_7, x_8, x_9, x_10, x_21); -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; -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_24, 0); -lean_inc(x_26); -x_27 = lean_ctor_get(x_24, 1); -lean_inc(x_27); -lean_dec(x_24); -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_26); -x_28 = lean_apply_8(x_2, x_26, x_5, x_6, x_7, x_8, x_9, x_10, x_25); -if (lean_obj_tag(x_28) == 0) -{ -if (lean_obj_tag(x_27) == 0) -{ -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; -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_box(0); -x_32 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__2(x_20, x_1, x_18, x_4, x_26, x_29, x_31, x_5, x_6, x_7, x_8, x_9, x_10, x_30); -lean_dec(x_10); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_26); -return x_32; +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; 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_25 = lean_float_to_string(x_8); +x_26 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_26, 0, x_25); +x_27 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_27, 0, x_26); +x_28 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__3; +x_29 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_27); +x_30 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__5; +x_31 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_30); +x_32 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_19); +x_33 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; +x_34 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +x_35 = lean_box(0); +x_36 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__2(x_3, x_4, x_17, x_5, x_2, x_34, x_35, x_10, x_11, x_12, x_13, x_14, x_15, x_20); +lean_dec(x_15); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_2); +return x_36; +} } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; double x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; 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_28, 0); -lean_inc(x_33); -x_34 = lean_ctor_get(x_28, 1); -lean_inc(x_34); -lean_dec(x_28); -x_35 = lean_ctor_get(x_27, 0); -lean_inc(x_35); -lean_dec(x_27); -x_36 = lean_unbox_float(x_35); -lean_dec(x_35); -x_37 = lean_float_to_string(x_36); +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; 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_37 = lean_float_to_string(x_8); x_38 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_38, 0, x_37); x_39 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_39, 0, x_38); -x_40 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___closed__2; +x_40 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__3; x_41 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_41, 0, x_40); lean_ctor_set(x_41, 1, x_39); -x_42 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___closed__4; +x_42 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__5; x_43 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_43, 0, x_41); lean_ctor_set(x_43, 1, x_42); x_44 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_44, 0, x_43); -lean_ctor_set(x_44, 1, x_33); +lean_ctor_set(x_44, 1, x_19); x_45 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; x_46 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_46, 0, x_44); lean_ctor_set(x_46, 1, x_45); x_47 = lean_box(0); -x_48 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__2(x_20, x_1, x_18, x_4, x_26, x_46, x_47, x_5, x_6, x_7, x_8, x_9, x_10, x_34); -lean_dec(x_10); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_26); +x_48 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__2(x_3, x_4, x_17, x_5, x_2, x_46, x_47, x_10, x_11, x_12, x_13, x_14, x_15, x_20); +lean_dec(x_15); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_2); return x_48; } } else { uint8_t x_49; -lean_dec(x_27); -lean_dec(x_26); -lean_dec(x_20); -lean_dec(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_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_1); -x_49 = !lean_is_exclusive(x_28); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_49 = !lean_is_exclusive(x_18); if (x_49 == 0) { -return x_28; +return x_18; } else { lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_50 = lean_ctor_get(x_28, 0); -x_51 = lean_ctor_get(x_28, 1); +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_28); +lean_dec(x_18); x_52 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_52, 0, x_50); lean_ctor_set(x_52, 1, x_51); @@ -6423,256 +6146,602 @@ return x_52; } } } -else -{ -uint8_t x_53; -lean_dec(x_20); -lean_dec(x_18); -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_53 = !lean_is_exclusive(x_23); -if (x_53 == 0) -{ -return x_23; } -else +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__4___closed__1() { +_start: { -lean_object* x_54; lean_object* x_55; lean_object* x_56; -x_54 = lean_ctor_get(x_23, 0); -x_55 = lean_ctor_get(x_23, 1); -lean_inc(x_55); -lean_inc(x_54); -lean_dec(x_23); -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_object* x_1; +x_1 = l_Lean_trace_profiler; +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_withTraceNode_x27___at_Lean_Meta_SynthInstance_newSubgoal___spec__8___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_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t 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, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { -if (lean_obj_tag(x_1) == 0) -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_9 = lean_ctor_get(x_1, 0); +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +lean_dec(x_7); +x_15 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_SynthInstance_newSubgoal___spec__11___rarg(x_13, x_14); +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_closure((void*)(l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__1), 8, 1); +lean_closure_set(x_18, 0, x_1); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); lean_inc(x_9); -lean_dec(x_1); -x_10 = l_Lean_Exception_toMessageData(x_9); -x_11 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_11, 0, x_10); -lean_ctor_set(x_11, 1, x_8); -return x_11; +lean_inc(x_8); +x_19 = l_Lean_withSeconds___at_Lean_Meta_SynthInstance_newSubgoal___spec__12(x_18, x_8, x_9, x_10, x_11, x_12, x_13, x_17); +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_69; uint8_t x_70; +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_20, 0); +lean_inc(x_22); +x_23 = lean_ctor_get(x_20, 1); +lean_inc(x_23); +lean_dec(x_20); +x_69 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__4___closed__1; +x_70 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_5, x_69); +if (x_70 == 0) +{ +uint8_t x_71; +x_71 = 0; +x_24 = x_71; +goto block_68; } else { -lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_12 = lean_ctor_get(x_1, 0); -lean_inc(x_12); -lean_dec(x_1); -x_13 = lean_ctor_get(x_12, 1); -lean_inc(x_13); -lean_dec(x_12); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_13); -lean_ctor_set(x_14, 1, x_8); -return x_14; -} +double x_72; double x_73; uint8_t x_74; +x_72 = l_Lean_trace_profiler_threshold_getSecs(x_5); +x_73 = lean_unbox_float(x_23); +x_74 = lean_float_decLt(x_72, x_73); +x_24 = x_74; +goto block_68; } -} -static lean_object* _init_l_Lean_withTraceNode_x27___at_Lean_Meta_SynthInstance_newSubgoal___spec__8___closed__1() { -_start: +block_68: { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_withTraceNode_x27___at_Lean_Meta_SynthInstance_newSubgoal___spec__8___lambda__1___boxed), 8, 0); -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_withTraceNode_x27___at_Lean_Meta_SynthInstance_newSubgoal___spec__8(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: +if (x_6 == 0) { -lean_object* x_11; lean_object* x_12; -x_11 = l_Lean_withTraceNode_x27___at_Lean_Meta_SynthInstance_newSubgoal___spec__8___closed__1; -x_12 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9(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) +if (x_24 == 0) { -uint8_t x_13; -x_13 = !lean_is_exclusive(x_12); -if (x_13 == 0) +lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; +lean_dec(x_23); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +x_25 = lean_st_ref_take(x_13, x_21); +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_is_exclusive(x_26); +if (x_28 == 0) { -lean_object* x_14; lean_object* x_15; -x_14 = lean_ctor_get(x_12, 0); -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -lean_dec(x_14); -lean_ctor_set(x_12, 0, x_15); -return x_12; +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, 3); +x_30 = l_Lean_PersistentArray_append___rarg(x_16, x_29); +lean_ctor_set(x_26, 3, x_30); +x_31 = lean_st_ref_set(x_13, x_26, x_27); +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +lean_dec(x_31); +x_33 = l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_newSubgoal___spec__15(x_22, x_8, x_9, x_10, x_11, x_12, x_13, x_32); +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_22); +if (lean_obj_tag(x_33) == 0) +{ +uint8_t x_34; +x_34 = !lean_is_exclusive(x_33); +if (x_34 == 0) +{ +return x_33; } 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); -x_18 = lean_ctor_get(x_16, 0); -lean_inc(x_18); -lean_dec(x_16); -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_18); -lean_ctor_set(x_19, 1, x_17); -return x_19; +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_33, 0); +x_36 = lean_ctor_get(x_33, 1); +lean_inc(x_36); +lean_inc(x_35); +lean_dec(x_33); +x_37 = lean_alloc_ctor(0, 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_20; -x_20 = !lean_is_exclusive(x_12); -if (x_20 == 0) +uint8_t x_38; +x_38 = !lean_is_exclusive(x_33); +if (x_38 == 0) { -return x_12; +return x_33; } else { -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_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_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_33, 0); +x_40 = lean_ctor_get(x_33, 1); +lean_inc(x_40); +lean_inc(x_39); +lean_dec(x_33); +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_EXPORT lean_object* l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__18___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) { -_start: -{ -lean_object* x_10; lean_object* x_11; -x_10 = lean_apply_2(x_2, x_3, x_4); -x_11 = l___private_Lean_Meta_Basic_0__Lean_Meta_withMCtxImp___rarg(x_1, x_10, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_11) == 0) +else { -uint8_t x_12; -x_12 = !lean_is_exclusive(x_11); -if (x_12 == 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; +x_42 = lean_ctor_get(x_26, 0); +x_43 = lean_ctor_get(x_26, 1); +x_44 = lean_ctor_get(x_26, 2); +x_45 = lean_ctor_get(x_26, 3); +x_46 = lean_ctor_get(x_26, 4); +x_47 = lean_ctor_get(x_26, 5); +x_48 = lean_ctor_get(x_26, 6); +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_inc(x_42); +lean_dec(x_26); +x_49 = l_Lean_PersistentArray_append___rarg(x_16, x_45); +x_50 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_50, 0, x_42); +lean_ctor_set(x_50, 1, x_43); +lean_ctor_set(x_50, 2, x_44); +lean_ctor_set(x_50, 3, x_49); +lean_ctor_set(x_50, 4, x_46); +lean_ctor_set(x_50, 5, x_47); +lean_ctor_set(x_50, 6, x_48); +x_51 = lean_st_ref_set(x_13, x_50, x_27); +x_52 = lean_ctor_get(x_51, 1); +lean_inc(x_52); +lean_dec(x_51); +x_53 = l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_newSubgoal___spec__15(x_22, x_8, x_9, x_10, x_11, x_12, x_13, x_52); +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_22); +if (lean_obj_tag(x_53) == 0) { -return x_11; +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_54 = lean_ctor_get(x_53, 0); +lean_inc(x_54); +x_55 = lean_ctor_get(x_53, 1); +lean_inc(x_55); +if (lean_is_exclusive(x_53)) { + lean_ctor_release(x_53, 0); + lean_ctor_release(x_53, 1); + x_56 = x_53; +} else { + lean_dec_ref(x_53); + x_56 = lean_box(0); +} +if (lean_is_scalar(x_56)) { + x_57 = lean_alloc_ctor(0, 2, 0); +} else { + x_57 = x_56; +} +lean_ctor_set(x_57, 0, x_54); +lean_ctor_set(x_57, 1, x_55); +return x_57; } else { -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_inc(x_13); -lean_dec(x_11); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_13); -lean_ctor_set(x_15, 1, x_14); -return x_15; +lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_58 = lean_ctor_get(x_53, 0); +lean_inc(x_58); +x_59 = lean_ctor_get(x_53, 1); +lean_inc(x_59); +if (lean_is_exclusive(x_53)) { + lean_ctor_release(x_53, 0); + lean_ctor_release(x_53, 1); + x_60 = x_53; +} else { + lean_dec_ref(x_53); + x_60 = lean_box(0); +} +if (lean_is_scalar(x_60)) { + x_61 = lean_alloc_ctor(1, 2, 0); +} else { + x_61 = x_60; +} +lean_ctor_set(x_61, 0, x_58); +lean_ctor_set(x_61, 1, x_59); +return x_61; +} } } else { -uint8_t x_16; -x_16 = !lean_is_exclusive(x_11); -if (x_16 == 0) +lean_object* x_62; double x_63; lean_object* x_64; +x_62 = lean_box(0); +x_63 = lean_unbox_float(x_23); +lean_dec(x_23); +x_64 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3(x_2, x_22, x_16, x_3, x_4, x_5, x_24, x_63, x_62, x_8, x_9, x_10, x_11, x_12, x_13, x_21); +return x_64; +} +} +else { -return x_11; +lean_object* x_65; double x_66; lean_object* x_67; +x_65 = lean_box(0); +x_66 = lean_unbox_float(x_23); +lean_dec(x_23); +x_67 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3(x_2, x_22, x_16, x_3, x_4, x_5, x_24, x_66, x_65, x_8, x_9, x_10, x_11, x_12, x_13, x_21); +return x_67; +} +} } else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_17 = lean_ctor_get(x_11, 0); -x_18 = lean_ctor_get(x_11, 1); -lean_inc(x_18); -lean_inc(x_17); +uint8_t x_75; +lean_dec(x_16); +lean_dec(x_13); +lean_dec(x_12); lean_dec(x_11); -x_19 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_19, 0, x_17); -lean_ctor_set(x_19, 1, x_18); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +x_75 = !lean_is_exclusive(x_19); +if (x_75 == 0) +{ return x_19; } +else +{ +lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_76 = lean_ctor_get(x_19, 0); +x_77 = lean_ctor_get(x_19, 1); +lean_inc(x_77); +lean_inc(x_76); +lean_dec(x_19); +x_78 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_78, 0, x_76); +lean_ctor_set(x_78, 1, x_77); +return x_78; +} } } } -LEAN_EXPORT lean_object* l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__18(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9(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) { _start: { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__18___rarg), 9, 0); -return x_2; -} +lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_12 = lean_ctor_get(x_9, 2); +lean_inc(x_12); +lean_inc(x_1); +x_13 = l_Lean_isTracingEnabledFor___at_Lean_Meta_SynthInstance_newSubgoal___spec__10(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_unbox(x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); +lean_dec(x_13); +x_17 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__4___closed__1; +x_18 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_12, x_17); +if (x_18 == 0) +{ +lean_object* x_19; +lean_dec(x_14); +lean_dec(x_12); +lean_dec(x_2); +lean_dec(x_1); +x_19 = lean_apply_7(x_3, x_5, x_6, x_7, x_8, x_9, x_10, x_16); +if (lean_obj_tag(x_19) == 0) +{ +uint8_t x_20; +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) +{ +return x_19; } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_newSubgoal___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: +else { -lean_object* x_10; -x_10 = l_Lean_Meta_SynthInstance_mkGeneratorNode_x3f(x_1, x_2, x_5, x_6, x_7, x_8, x_9); -return x_10; +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(0, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +return x_23; } } -static lean_object* _init_l_Lean_Meta_SynthInstance_newSubgoal___lambda__2___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("no instances for ", 17); -return x_1; +uint8_t x_24; +x_24 = !lean_is_exclusive(x_19); +if (x_24 == 0) +{ +return x_19; +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_19, 0); +x_26 = lean_ctor_get(x_19, 1); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_19); +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; } } -static lean_object* _init_l_Lean_Meta_SynthInstance_newSubgoal___lambda__2___closed__2() { -_start: +} +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_SynthInstance_newSubgoal___lambda__2___closed__1; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +lean_object* x_28; uint8_t x_29; lean_object* x_30; +x_28 = lean_box(0); +x_29 = lean_unbox(x_14); +lean_dec(x_14); +x_30 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__4(x_3, x_2, x_1, x_4, x_12, x_29, x_28, x_5, x_6, x_7, x_8, x_9, x_10, x_16); +return x_30; } } -static lean_object* _init_l_Lean_Meta_SynthInstance_newSubgoal___lambda__2___closed__3() { -_start: +else { -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_object* x_31; lean_object* x_32; uint8_t x_33; lean_object* x_34; +x_31 = lean_ctor_get(x_13, 1); +lean_inc(x_31); +lean_dec(x_13); +x_32 = lean_box(0); +x_33 = lean_unbox(x_14); +lean_dec(x_14); +x_34 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__4(x_3, x_2, x_1, x_4, x_12, x_33, x_32, x_5, x_6, x_7, x_8, x_9, x_10, x_31); +return x_34; } } -static lean_object* _init_l_Lean_Meta_SynthInstance_newSubgoal___lambda__2___closed__4() { +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode_x27___at_Lean_Meta_SynthInstance_newSubgoal___spec__8___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_1; -x_1 = lean_mk_string_from_bytes("new goal ", 9); -return x_1; +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_9 = lean_ctor_get(x_1, 0); +lean_inc(x_9); +lean_dec(x_1); +x_10 = l_Lean_Exception_toMessageData(x_9); +x_11 = lean_alloc_ctor(0, 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_12; lean_object* x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_1, 0); +lean_inc(x_12); +lean_dec(x_1); +x_13 = lean_ctor_get(x_12, 1); +lean_inc(x_13); +lean_dec(x_12); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_8); +return x_14; } } -static lean_object* _init_l_Lean_Meta_SynthInstance_newSubgoal___lambda__2___closed__5() { +} +static lean_object* _init_l_Lean_withTraceNode_x27___at_Lean_Meta_SynthInstance_newSubgoal___spec__8___closed__1() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_SynthInstance_newSubgoal___lambda__2___closed__4; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_withTraceNode_x27___at_Lean_Meta_SynthInstance_newSubgoal___spec__8___lambda__1___boxed), 8, 0); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_newSubgoal___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_Lean_withTraceNode_x27___at_Lean_Meta_SynthInstance_newSubgoal___spec__8(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: { -if (lean_obj_tag(x_3) == 0) +lean_object* x_11; lean_object* x_12; +x_11 = l_Lean_withTraceNode_x27___at_Lean_Meta_SynthInstance_newSubgoal___spec__8___closed__1; +x_12 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9(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) +{ +uint8_t x_13; +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); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +lean_dec(x_14); +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); +x_18 = lean_ctor_get(x_16, 0); +lean_inc(x_18); +lean_dec(x_16); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_17); +return x_19; +} +} +else +{ +uint8_t x_20; +x_20 = !lean_is_exclusive(x_12); +if (x_20 == 0) +{ +return x_12; +} +else +{ +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_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_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__16___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) { +_start: +{ +lean_object* x_10; lean_object* x_11; +x_10 = lean_apply_2(x_2, x_3, x_4); +x_11 = l___private_Lean_Meta_Basic_0__Lean_Meta_withMCtxImp___rarg(x_1, x_10, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_11) == 0) +{ +uint8_t x_12; +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) +{ +return x_11; +} +else +{ +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_inc(x_13); +lean_dec(x_11); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_13); +lean_ctor_set(x_15, 1, x_14); +return x_15; +} +} +else +{ +uint8_t x_16; +x_16 = !lean_is_exclusive(x_11); +if (x_16 == 0) +{ +return x_11; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_11, 0); +x_18 = lean_ctor_get(x_11, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_11); +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_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__16(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__16___rarg), 9, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_newSubgoal___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_Lean_Meta_SynthInstance_mkGeneratorNode_x3f(x_1, x_2, x_5, x_6, x_7, x_8, x_9); +return x_10; +} +} +static lean_object* _init_l_Lean_Meta_SynthInstance_newSubgoal___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("no instances for ", 17); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_SynthInstance_newSubgoal___lambda__2___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_SynthInstance_newSubgoal___lambda__2___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_SynthInstance_newSubgoal___lambda__2___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; +} +} +static lean_object* _init_l_Lean_Meta_SynthInstance_newSubgoal___lambda__2___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("new goal ", 9); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_SynthInstance_newSubgoal___lambda__2___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_SynthInstance_newSubgoal___lambda__2___closed__4; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_newSubgoal___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: +{ +if (lean_obj_tag(x_3) == 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_dec(x_2); @@ -6858,7 +6927,7 @@ x_18 = lean_alloc_closure((void*)(l_Lean_withTraceNode_x27___at_Lean_Meta_SynthI lean_closure_set(x_18, 0, x_15); lean_closure_set(x_18, 1, x_14); lean_closure_set(x_18, 2, x_17); -x_19 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__18___rarg(x_1, x_18, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_19 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__16___rarg(x_1, x_18, x_5, x_6, x_7, x_8, x_9, x_10, x_11); return x_19; } } @@ -6909,30 +6978,13 @@ lean_dec(x_1); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_getBoolOption___at_Lean_Meta_SynthInstance_newSubgoal___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) { -_start: -{ -uint8_t x_10; lean_object* x_11; -x_10 = lean_unbox(x_2); -lean_dec(x_2); -x_11 = l_Lean_getBoolOption___at_Lean_Meta_SynthInstance_newSubgoal___spec__13(x_1, x_10, 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_1); -return x_11; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Meta_SynthInstance_newSubgoal___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_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Meta_SynthInstance_newSubgoal___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) { _start: { uint8_t x_13; lean_object* x_14; x_13 = lean_unbox(x_5); lean_dec(x_5); -x_14 = l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Meta_SynthInstance_newSubgoal___spec__16(x_1, 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___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___at_Lean_Meta_SynthInstance_newSubgoal___spec__14(x_1, x_2, x_3, x_4, x_13, 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); @@ -6942,13 +6994,13 @@ lean_dec(x_6); return x_14; } } -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___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_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___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, lean_object* x_11, lean_object* x_12) { _start: { uint8_t x_13; lean_object* x_14; x_13 = lean_unbox(x_5); lean_dec(x_5); -x_14 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__15(x_1, 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___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__13(x_1, x_2, x_3, x_4, x_13, 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); @@ -6957,11 +7009,11 @@ lean_dec(x_6); return x_14; } } -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_newSubgoal___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_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_newSubgoal___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) { _start: { lean_object* x_9; -x_9 = l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_newSubgoal___spec__17(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_9 = l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_newSubgoal___spec__15(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); @@ -6988,6 +7040,32 @@ lean_dec(x_5); return x_16; } } +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___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: +{ +uint8_t x_17; uint8_t x_18; double x_19; lean_object* x_20; +x_17 = lean_unbox(x_5); +lean_dec(x_5); +x_18 = lean_unbox(x_7); +lean_dec(x_7); +x_19 = lean_unbox_float(x_8); +lean_dec(x_8); +x_20 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3(x_1, x_2, x_3, x_4, x_17, x_6, x_18, x_19, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +return x_20; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___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: +{ +uint8_t x_15; uint8_t x_16; lean_object* x_17; +x_15 = lean_unbox(x_4); +lean_dec(x_4); +x_16 = lean_unbox(x_6); +lean_dec(x_6); +x_17 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__4(x_1, x_2, x_3, x_15, x_5, x_16, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +return x_17; +} +} LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___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) { _start: { @@ -7197,7 +7275,7 @@ static lean_object* _init_l_Lean_Meta_SynthInstance_getEntry___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_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_getInstances___spec__7___closed__1; x_2 = l_Lean_Meta_SynthInstance_getEntry___closed__1; -x_3 = lean_unsigned_to_nat(263u); +x_3 = lean_unsigned_to_nat(258u); x_4 = lean_unsigned_to_nat(18u); x_5 = l_Lean_Meta_SynthInstance_getEntry___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -7504,7 +7582,7 @@ x_11 = l_Lean_Meta_SynthInstance_mkTableKeyFor___closed__1; x_12 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_SynthInstance_newSubgoal___spec__7___rarg), 9, 2); lean_closure_set(x_12, 0, x_10); lean_closure_set(x_12, 1, x_11); -x_13 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__18___rarg(x_1, x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_13 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__16___rarg(x_1, x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_13; } } @@ -8112,7 +8190,7 @@ x_2 = lean_alloc_closure((void*)(l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstan return x_2; } } -LEAN_EXPORT lean_object* l_Lean_withSeconds___at_Lean_Meta_SynthInstance_tryResolve___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_EXPORT lean_object* l_Lean_withSeconds___at_Lean_Meta_SynthInstance_tryResolve___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; lean_object* x_8; lean_object* x_9; lean_object* x_10; @@ -8144,7 +8222,7 @@ x_17 = 0; x_18 = lean_unsigned_to_nat(0u); x_19 = l_Float_ofScientific(x_16, x_17, x_18); lean_dec(x_16); -x_20 = l_Lean_withSeconds___at_Lean_Meta_SynthInstance_newSubgoal___spec__14___closed__1; +x_20 = l_Lean_withSeconds___at_Lean_Meta_SynthInstance_newSubgoal___spec__12___closed__1; x_21 = lean_float_div(x_19, x_20); x_22 = lean_box_float(x_21); x_23 = lean_alloc_ctor(0, 2, 0); @@ -8168,7 +8246,7 @@ x_27 = 0; x_28 = lean_unsigned_to_nat(0u); x_29 = l_Float_ofScientific(x_26, x_27, x_28); lean_dec(x_26); -x_30 = l_Lean_withSeconds___at_Lean_Meta_SynthInstance_newSubgoal___spec__14___closed__1; +x_30 = l_Lean_withSeconds___at_Lean_Meta_SynthInstance_newSubgoal___spec__12___closed__1; x_31 = lean_float_div(x_29, x_30); x_32 = lean_box_float(x_31); x_33 = lean_alloc_ctor(0, 2, 0); @@ -8205,204 +8283,28 @@ return x_38; } } } -LEAN_EXPORT lean_object* l_Lean_withOptProfile___at_Lean_Meta_SynthInstance_tryResolve___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_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_tryResolve___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) { _start: { -lean_object* x_7; uint8_t x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_7 = l_Lean_withOptProfile___at_Lean_Meta_SynthInstance_newSubgoal___spec__12___closed__2; -x_8 = 0; -x_9 = l_Lean_getBoolOption___at_Lean_Meta_processPostponed___spec__4(x_7, x_8, x_2, x_3, x_4, x_5, x_6); -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -x_11 = lean_unbox(x_10); -lean_dec(x_10); -if (x_11 == 0) -{ -lean_object* x_12; lean_object* x_13; -x_12 = lean_ctor_get(x_9, 1); -lean_inc(x_12); -lean_dec(x_9); -x_13 = lean_apply_5(x_1, x_2, x_3, x_4, x_5, x_12); -if (lean_obj_tag(x_13) == 0) -{ -uint8_t x_14; -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) +if (lean_obj_tag(x_1) == 0) { -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_13, 0); -x_16 = lean_box(0); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_15); -lean_ctor_set(x_17, 1, x_16); -lean_ctor_set(x_13, 0, x_17); -return x_13; +lean_object* x_7; lean_object* x_8; +x_7 = lean_ctor_get(x_1, 0); +lean_inc(x_7); +x_8 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_8, 0, x_7); +lean_ctor_set(x_8, 1, x_6); +return x_8; } else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_18 = lean_ctor_get(x_13, 0); -x_19 = lean_ctor_get(x_13, 1); -lean_inc(x_19); -lean_inc(x_18); -lean_dec(x_13); -x_20 = lean_box(0); -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(0, 2, 0); -lean_ctor_set(x_22, 0, x_21); -lean_ctor_set(x_22, 1, x_19); -return x_22; -} -} -else -{ -uint8_t x_23; -x_23 = !lean_is_exclusive(x_13); -if (x_23 == 0) -{ -return x_13; -} -else -{ -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_13, 0); -x_25 = lean_ctor_get(x_13, 1); -lean_inc(x_25); -lean_inc(x_24); -lean_dec(x_13); -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_27; lean_object* x_28; -x_27 = lean_ctor_get(x_9, 1); -lean_inc(x_27); -lean_dec(x_9); -x_28 = l_Lean_withSeconds___at_Lean_Meta_SynthInstance_tryResolve___spec__4(x_1, x_2, x_3, x_4, x_5, x_27); -if (lean_obj_tag(x_28) == 0) -{ -uint8_t x_29; -x_29 = !lean_is_exclusive(x_28); -if (x_29 == 0) -{ -lean_object* x_30; uint8_t x_31; -x_30 = lean_ctor_get(x_28, 0); -x_31 = !lean_is_exclusive(x_30); -if (x_31 == 0) -{ -lean_object* x_32; lean_object* x_33; -x_32 = lean_ctor_get(x_30, 1); -x_33 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_33, 0, x_32); -lean_ctor_set(x_30, 1, x_33); -return x_28; -} -else -{ -lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_34 = lean_ctor_get(x_30, 0); -x_35 = lean_ctor_get(x_30, 1); -lean_inc(x_35); -lean_inc(x_34); -lean_dec(x_30); -x_36 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_36, 0, x_35); -x_37 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_37, 0, x_34); -lean_ctor_set(x_37, 1, x_36); -lean_ctor_set(x_28, 0, x_37); -return x_28; -} -} -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; -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_ctor_get(x_38, 0); -lean_inc(x_40); -x_41 = lean_ctor_get(x_38, 1); -lean_inc(x_41); -if (lean_is_exclusive(x_38)) { - lean_ctor_release(x_38, 0); - lean_ctor_release(x_38, 1); - x_42 = x_38; -} else { - lean_dec_ref(x_38); - x_42 = lean_box(0); -} -x_43 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_43, 0, x_41); -if (lean_is_scalar(x_42)) { - x_44 = lean_alloc_ctor(0, 2, 0); -} else { - x_44 = x_42; -} -lean_ctor_set(x_44, 0, x_40); -lean_ctor_set(x_44, 1, 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_39); -return x_45; -} -} -else -{ -uint8_t x_46; -x_46 = !lean_is_exclusive(x_28); -if (x_46 == 0) -{ -return x_28; -} -else -{ -lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_47 = lean_ctor_get(x_28, 0); -x_48 = lean_ctor_get(x_28, 1); -lean_inc(x_48); -lean_inc(x_47); -lean_dec(x_28); -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_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_tryResolve___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: -{ -if (lean_obj_tag(x_1) == 0) -{ -lean_object* x_7; lean_object* x_8; -x_7 = lean_ctor_get(x_1, 0); -lean_inc(x_7); -x_8 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_8, 0, x_7); -lean_ctor_set(x_8, 1, x_6); -return x_8; -} -else -{ -lean_object* x_9; lean_object* x_10; -x_9 = lean_ctor_get(x_1, 0); -lean_inc(x_9); -x_10 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_10, 0, x_9); -lean_ctor_set(x_10, 1, x_6); -return x_10; +lean_object* x_9; lean_object* x_10; +x_9 = lean_ctor_get(x_1, 0); +lean_inc(x_9); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_6); +return x_10; } } } @@ -8477,161 +8379,141 @@ LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_tryRe { lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_inc(x_10); -x_13 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_processPostponed___spec__6(x_1, x_2, x_3, x_6, x_4, x_8, x_9, x_10, x_11, x_12); +x_13 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_processPostponed___spec__4(x_1, x_2, x_3, x_6, x_4, x_8, x_9, x_10, x_11, x_12); x_14 = lean_ctor_get(x_13, 1); lean_inc(x_14); lean_dec(x_13); -x_15 = l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_tryResolve___spec__5(x_5, x_8, x_9, x_10, x_11, x_14); +x_15 = l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_tryResolve___spec__4(x_5, x_8, x_9, x_10, x_11, x_14); lean_dec(x_10); return x_15; } } -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_tryResolve___spec__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_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_tryResolve___spec__2___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, uint8_t x_7, double 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_10; lean_object* x_11; uint8_t x_12; -lean_inc(x_1); -x_10 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_1, x_5, x_6, x_7, x_8, x_9); -x_11 = lean_ctor_get(x_10, 0); +lean_object* x_15; lean_object* x_16; +lean_dec(x_9); +x_15 = lean_ctor_get(x_12, 5); +lean_inc(x_15); +lean_inc(x_13); +lean_inc(x_12); lean_inc(x_11); -x_12 = lean_unbox(x_11); -lean_dec(x_11); -if (x_12 == 0) +lean_inc(x_10); +lean_inc(x_2); +x_16 = lean_apply_6(x_1, x_2, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_13; lean_object* x_14; -lean_dec(x_2); -lean_dec(x_1); -x_13 = lean_ctor_get(x_10, 1); -lean_inc(x_13); +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); +x_19 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__1; +x_20 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_6, x_19); +lean_dec(x_6); +if (x_20 == 0) +{ +if (x_7 == 0) +{ +lean_object* x_21; lean_object* x_22; +x_21 = lean_box(0); +x_22 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_tryResolve___spec__2___lambda__2(x_3, x_4, x_15, x_5, x_2, x_17, x_21, x_10, x_11, x_12, x_13, x_18); +lean_dec(x_13); +lean_dec(x_11); lean_dec(x_10); -x_14 = lean_apply_5(x_3, x_5, x_6, x_7, x_8, x_13); -return x_14; +lean_dec(x_2); +return x_22; } 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; -x_15 = lean_ctor_get(x_10, 1); -lean_inc(x_15); +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; 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_23 = lean_float_to_string(x_8); +x_24 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_24, 0, x_23); +x_25 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_25, 0, x_24); +x_26 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__3; +x_27 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_25); +x_28 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__5; +x_29 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +x_30 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_17); +x_31 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; +x_32 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +x_33 = lean_box(0); +x_34 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_tryResolve___spec__2___lambda__2(x_3, x_4, x_15, x_5, x_2, x_32, x_33, x_10, x_11, x_12, x_13, x_18); +lean_dec(x_13); +lean_dec(x_11); lean_dec(x_10); -x_16 = lean_ctor_get(x_7, 5); -lean_inc(x_16); -x_17 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_processPostponed___spec__2___rarg(x_8, 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_alloc_closure((void*)(l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_tryResolve___spec__2___lambda__1), 6, 1); -lean_closure_set(x_20, 0, x_3); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_21 = l_Lean_withOptProfile___at_Lean_Meta_SynthInstance_tryResolve___spec__3(x_20, x_5, x_6, x_7, x_8, 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; -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_22, 0); -lean_inc(x_24); -x_25 = lean_ctor_get(x_22, 1); -lean_inc(x_25); -lean_dec(x_22); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_24); -x_26 = lean_apply_6(x_2, x_24, x_5, x_6, x_7, x_8, x_23); -if (lean_obj_tag(x_26) == 0) -{ -if (lean_obj_tag(x_25) == 0) -{ -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; -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_box(0); -x_30 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_tryResolve___spec__2___lambda__2(x_18, x_1, x_16, x_4, x_24, x_27, x_29, x_5, x_6, x_7, x_8, x_28); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_24); -return x_30; +lean_dec(x_2); +return x_34; +} } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; double x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; 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_31 = lean_ctor_get(x_26, 0); -lean_inc(x_31); -x_32 = lean_ctor_get(x_26, 1); -lean_inc(x_32); -lean_dec(x_26); -x_33 = lean_ctor_get(x_25, 0); -lean_inc(x_33); -lean_dec(x_25); -x_34 = lean_unbox_float(x_33); -lean_dec(x_33); -x_35 = lean_float_to_string(x_34); +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; 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_float_to_string(x_8); x_36 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_36, 0, x_35); x_37 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_37, 0, x_36); -x_38 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___closed__2; +x_38 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__3; x_39 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_39, 0, x_38); lean_ctor_set(x_39, 1, x_37); -x_40 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___closed__4; +x_40 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__5; x_41 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_41, 0, x_39); lean_ctor_set(x_41, 1, x_40); x_42 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_31); +lean_ctor_set(x_42, 1, x_17); x_43 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; x_44 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_44, 0, x_42); lean_ctor_set(x_44, 1, x_43); x_45 = lean_box(0); -x_46 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_tryResolve___spec__2___lambda__2(x_18, x_1, x_16, x_4, x_24, x_44, x_45, x_5, x_6, x_7, x_8, x_32); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_24); +x_46 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_tryResolve___spec__2___lambda__2(x_3, x_4, x_15, x_5, x_2, x_44, x_45, x_10, x_11, x_12, x_13, x_18); +lean_dec(x_13); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_2); return x_46; } } else { uint8_t x_47; -lean_dec(x_25); -lean_dec(x_24); -lean_dec(x_18); -lean_dec(x_16); -lean_dec(x_8); -lean_dec(x_7); +lean_dec(x_15); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_1); -x_47 = !lean_is_exclusive(x_26); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_47 = !lean_is_exclusive(x_16); if (x_47 == 0) { -return x_26; +return x_16; } else { lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_48 = lean_ctor_get(x_26, 0); -x_49 = lean_ctor_get(x_26, 1); +x_48 = lean_ctor_get(x_16, 0); +x_49 = lean_ctor_get(x_16, 1); lean_inc(x_49); lean_inc(x_48); -lean_dec(x_26); +lean_dec(x_16); x_50 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_50, 0, x_48); lean_ctor_set(x_50, 1, x_49); @@ -8639,2941 +8521,2625 @@ return x_50; } } } -else -{ -uint8_t x_51; -lean_dec(x_18); -lean_dec(x_16); -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_51 = !lean_is_exclusive(x_21); -if (x_51 == 0) -{ -return x_21; } -else +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_tryResolve___spec__2___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t 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, lean_object* x_12) { +_start: { -lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_52 = lean_ctor_get(x_21, 0); -x_53 = lean_ctor_get(x_21, 1); -lean_inc(x_53); -lean_inc(x_52); -lean_dec(x_21); -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; -} -} -} -} -} -static lean_object* _init_l_Lean_Meta_SynthInstance_tryResolve___lambda__1___closed__1() { -_start: +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +lean_dec(x_7); +x_13 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_processPostponed___spec__2___rarg(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_alloc_closure((void*)(l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_tryResolve___spec__2___lambda__1), 6, 1); +lean_closure_set(x_16, 0, x_1); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_17 = l_Lean_withSeconds___at_Lean_Meta_SynthInstance_tryResolve___spec__3(x_16, x_8, x_9, x_10, x_11, x_15); +if (lean_obj_tag(x_17) == 0) { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes(" ", 1); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_SynthInstance_tryResolve___lambda__1___closed__2() { -_start: +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; lean_object* x_67; uint8_t x_68; +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_ctor_get(x_18, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_18, 1); +lean_inc(x_21); +lean_dec(x_18); +x_67 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__4___closed__1; +x_68 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_5, x_67); +if (x_68 == 0) { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_SynthInstance_tryResolve___lambda__1___closed__1; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +uint8_t x_69; +x_69 = 0; +x_22 = x_69; +goto block_66; } -} -static lean_object* _init_l_Lean_Meta_SynthInstance_tryResolve___lambda__1___closed__3() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes(" ≟ ", 5); -return x_1; -} +double x_70; double x_71; uint8_t x_72; +x_70 = l_Lean_trace_profiler_threshold_getSecs(x_5); +x_71 = lean_unbox_float(x_21); +x_72 = lean_float_decLt(x_70, x_71); +x_22 = x_72; +goto block_66; } -static lean_object* _init_l_Lean_Meta_SynthInstance_tryResolve___lambda__1___closed__4() { -_start: +block_66: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_SynthInstance_tryResolve___lambda__1___closed__3; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_tryResolve___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: +if (x_6 == 0) { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_9 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_1, 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); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; +lean_dec(x_21); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +x_23 = lean_st_ref_take(x_11, 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_is_exclusive(x_24); +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_24, 3); +x_28 = l_Lean_PersistentArray_append___rarg(x_14, x_27); +lean_ctor_set(x_24, 3, x_28); +x_29 = lean_st_ref_set(x_11, x_24, x_25); +x_30 = lean_ctor_get(x_29, 1); +lean_inc(x_30); +lean_dec(x_29); +x_31 = l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_tryResolve___spec__4(x_20, x_8, x_9, x_10, x_11, x_30); +lean_dec(x_11); +lean_dec(x_10); lean_dec(x_9); -x_12 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_2, x_4, x_5, x_6, x_7, x_11); -x_13 = !lean_is_exclusive(x_12); -if (x_13 == 0) +lean_dec(x_8); +lean_dec(x_20); +if (lean_obj_tag(x_31) == 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; -x_14 = lean_ctor_get(x_12, 0); -x_15 = l_Lean_exceptOptionEmoji___rarg(x_3); -x_16 = l_Lean_stringToMessageData(x_15); -lean_dec(x_15); -x_17 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; -x_18 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_18, 0, x_17); -lean_ctor_set(x_18, 1, x_16); -x_19 = l_Lean_Meta_SynthInstance_tryResolve___lambda__1___closed__2; -x_20 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_20, 0, x_18); -lean_ctor_set(x_20, 1, x_19); -x_21 = l_Lean_MessageData_ofExpr(x_10); -x_22 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_22, 0, x_20); -lean_ctor_set(x_22, 1, x_21); -x_23 = l_Lean_Meta_SynthInstance_tryResolve___lambda__1___closed__4; -x_24 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_24, 0, x_22); -lean_ctor_set(x_24, 1, x_23); -x_25 = l_Lean_MessageData_ofExpr(x_14); -x_26 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_26, 0, x_24); -lean_ctor_set(x_26, 1, x_25); -x_27 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_27, 0, x_26); -lean_ctor_set(x_27, 1, x_17); -lean_ctor_set(x_12, 0, x_27); -return x_12; +uint8_t x_32; +x_32 = !lean_is_exclusive(x_31); +if (x_32 == 0) +{ +return x_31; } 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; -x_28 = lean_ctor_get(x_12, 0); -x_29 = lean_ctor_get(x_12, 1); -lean_inc(x_29); -lean_inc(x_28); -lean_dec(x_12); -x_30 = l_Lean_exceptOptionEmoji___rarg(x_3); -x_31 = l_Lean_stringToMessageData(x_30); -lean_dec(x_30); -x_32 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; -x_33 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_33, 0, x_32); -lean_ctor_set(x_33, 1, x_31); -x_34 = l_Lean_Meta_SynthInstance_tryResolve___lambda__1___closed__2; -x_35 = lean_alloc_ctor(7, 2, 0); +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(0, 2, 0); lean_ctor_set(x_35, 0, x_33); lean_ctor_set(x_35, 1, x_34); -x_36 = l_Lean_MessageData_ofExpr(x_10); -x_37 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_37, 0, x_35); -lean_ctor_set(x_37, 1, x_36); -x_38 = l_Lean_Meta_SynthInstance_tryResolve___lambda__1___closed__4; -x_39 = lean_alloc_ctor(7, 2, 0); +return x_35; +} +} +else +{ +uint8_t x_36; +x_36 = !lean_is_exclusive(x_31); +if (x_36 == 0) +{ +return x_31; +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_31, 0); +x_38 = lean_ctor_get(x_31, 1); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_31); +x_39 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_39, 0, x_37); lean_ctor_set(x_39, 1, x_38); -x_40 = l_Lean_MessageData_ofExpr(x_28); -x_41 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_41, 0, x_39); -lean_ctor_set(x_41, 1, x_40); -x_42 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_32); -x_43 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_43, 0, x_42); -lean_ctor_set(x_43, 1, x_29); -return x_43; +return x_39; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_tryResolve___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) { -_start: +else { -lean_object* x_10; lean_object* x_11; -x_10 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_tryResolve___lambda__1___boxed), 8, 3); -lean_closure_set(x_10, 0, x_1); -lean_closure_set(x_10, 1, x_2); -lean_closure_set(x_10, 2, x_4); -x_11 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_tryResolve___spec__1___rarg(x_3, x_10, x_5, x_6, x_7, x_8, x_9); -return x_11; +lean_object* x_40; lean_object* x_41; 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_40 = lean_ctor_get(x_24, 0); +x_41 = lean_ctor_get(x_24, 1); +x_42 = lean_ctor_get(x_24, 2); +x_43 = lean_ctor_get(x_24, 3); +x_44 = lean_ctor_get(x_24, 4); +x_45 = lean_ctor_get(x_24, 5); +x_46 = lean_ctor_get(x_24, 6); +lean_inc(x_46); +lean_inc(x_45); +lean_inc(x_44); +lean_inc(x_43); +lean_inc(x_42); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_24); +x_47 = l_Lean_PersistentArray_append___rarg(x_14, x_43); +x_48 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_48, 0, x_40); +lean_ctor_set(x_48, 1, x_41); +lean_ctor_set(x_48, 2, x_42); +lean_ctor_set(x_48, 3, x_47); +lean_ctor_set(x_48, 4, x_44); +lean_ctor_set(x_48, 5, x_45); +lean_ctor_set(x_48, 6, x_46); +x_49 = lean_st_ref_set(x_11, x_48, x_25); +x_50 = lean_ctor_get(x_49, 1); +lean_inc(x_50); +lean_dec(x_49); +x_51 = l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_tryResolve___spec__4(x_20, x_8, x_9, x_10, x_11, x_50); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_20); +if (lean_obj_tag(x_51) == 0) +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_52 = lean_ctor_get(x_51, 0); +lean_inc(x_52); +x_53 = lean_ctor_get(x_51, 1); +lean_inc(x_53); +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_51); + x_54 = lean_box(0); } +if (lean_is_scalar(x_54)) { + x_55 = lean_alloc_ctor(0, 2, 0); +} else { + x_55 = x_54; } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_tryResolve___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) { -_start: +lean_ctor_set(x_55, 0, x_52); +lean_ctor_set(x_55, 1, x_53); +return x_55; +} +else { -lean_object* x_7; lean_object* x_8; -x_7 = lean_box(0); -x_8 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_8, 0, x_7); -lean_ctor_set(x_8, 1, x_6); -return x_8; +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_56 = lean_ctor_get(x_51, 0); +lean_inc(x_56); +x_57 = lean_ctor_get(x_51, 1); +lean_inc(x_57); +if (lean_is_exclusive(x_51)) { + lean_ctor_release(x_51, 0); + lean_ctor_release(x_51, 1); + x_58 = x_51; +} else { + lean_dec_ref(x_51); + x_58 = lean_box(0); +} +if (lean_is_scalar(x_58)) { + x_59 = lean_alloc_ctor(1, 2, 0); +} else { + x_59 = x_58; } +lean_ctor_set(x_59, 0, x_56); +lean_ctor_set(x_59, 1, x_57); +return x_59; } -static lean_object* _init_l_Lean_Meta_SynthInstance_tryResolve___lambda__4___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_tryResolve___lambda__3___boxed), 6, 0); -return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_tryResolve___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) { -_start: +else { -lean_object* x_12; -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -x_12 = l_Lean_Meta_isExprDefEq(x_1, x_2, 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; 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 = l_Lean_Meta_SynthInstance_tryResolve___lambda__4___closed__1; -x_16 = lean_unbox(x_13); -lean_dec(x_13); -if (x_16 == 0) -{ -lean_object* x_17; lean_object* x_18; -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -x_17 = lean_box(0); -x_18 = lean_apply_6(x_15, x_17, x_7, x_8, x_9, x_10, x_14); -return x_18; +lean_object* x_60; double x_61; lean_object* x_62; +x_60 = lean_box(0); +x_61 = lean_unbox_float(x_21); +lean_dec(x_21); +x_62 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_tryResolve___spec__2___lambda__3(x_2, x_20, x_14, x_3, x_4, x_5, x_22, x_61, x_60, x_8, x_9, x_10, x_11, x_19); +return x_62; +} } else { -uint8_t x_19; uint8_t x_20; uint8_t x_21; lean_object* x_22; -x_19 = 0; -x_20 = 1; -x_21 = 1; -x_22 = l_Lean_Meta_mkLambdaFVars(x_3, x_4, x_19, x_20, x_21, x_7, x_8, x_9, x_10, x_14); -if (lean_obj_tag(x_22) == 0) -{ -lean_object* x_23; lean_object* x_24; lean_object* 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); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -x_25 = l_Lean_Meta_isExprDefEq(x_5, x_23, x_7, x_8, x_9, x_10, x_24); -if (lean_obj_tag(x_25) == 0) -{ -lean_object* x_26; uint8_t x_27; -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_dec(x_6); -x_28 = lean_ctor_get(x_25, 1); -lean_inc(x_28); -lean_dec(x_25); -x_29 = lean_box(0); -x_30 = lean_apply_6(x_15, x_29, x_7, x_8, x_9, x_10, x_28); -return x_30; +lean_object* x_63; double x_64; lean_object* x_65; +x_63 = lean_box(0); +x_64 = lean_unbox_float(x_21); +lean_dec(x_21); +x_65 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_tryResolve___spec__2___lambda__3(x_2, x_20, x_14, x_3, x_4, x_5, x_22, x_64, x_63, x_8, x_9, x_10, x_11, x_19); +return x_65; +} +} } else { -lean_object* x_31; lean_object* x_32; uint8_t x_33; +uint8_t x_73; +lean_dec(x_14); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_7); -x_31 = lean_ctor_get(x_25, 1); -lean_inc(x_31); -lean_dec(x_25); -x_32 = lean_st_ref_get(x_8, x_31); lean_dec(x_8); -x_33 = !lean_is_exclusive(x_32); -if (x_33 == 0) +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +x_73 = !lean_is_exclusive(x_17); +if (x_73 == 0) { -lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; -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_alloc_ctor(0, 2, 0); -lean_ctor_set(x_36, 0, x_35); -lean_ctor_set(x_36, 1, x_6); -x_37 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_37, 0, x_36); -lean_ctor_set(x_32, 0, x_37); -return x_32; +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; -x_38 = lean_ctor_get(x_32, 0); -x_39 = lean_ctor_get(x_32, 1); -lean_inc(x_39); -lean_inc(x_38); -lean_dec(x_32); -x_40 = lean_ctor_get(x_38, 0); -lean_inc(x_40); -lean_dec(x_38); -x_41 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_41, 0, x_40); -lean_ctor_set(x_41, 1, x_6); -x_42 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_42, 0, x_41); -x_43 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_43, 0, x_42); -lean_ctor_set(x_43, 1, x_39); -return x_43; +lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_74 = lean_ctor_get(x_17, 0); +x_75 = lean_ctor_get(x_17, 1); +lean_inc(x_75); +lean_inc(x_74); +lean_dec(x_17); +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; } } } -else +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_tryResolve___spec__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) { +_start: { -uint8_t x_44; +lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_10 = lean_ctor_get(x_7, 2); +lean_inc(x_10); +lean_inc(x_1); +x_11 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___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_unbox(x_12); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); +lean_dec(x_11); +x_15 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__4___closed__1; +x_16 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_10, x_15); +if (x_16 == 0) +{ +lean_object* x_17; +lean_dec(x_12); lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -x_44 = !lean_is_exclusive(x_25); -if (x_44 == 0) +lean_dec(x_2); +lean_dec(x_1); +x_17 = lean_apply_5(x_3, x_5, x_6, x_7, x_8, x_14); +if (lean_obj_tag(x_17) == 0) { -return x_25; +uint8_t x_18; +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) +{ +return x_17; } else { -lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_45 = lean_ctor_get(x_25, 0); -x_46 = lean_ctor_get(x_25, 1); -lean_inc(x_46); -lean_inc(x_45); -lean_dec(x_25); -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_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_17, 0); +x_20 = lean_ctor_get(x_17, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_17); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); +return x_21; } } else { -uint8_t x_48; -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_48 = !lean_is_exclusive(x_22); -if (x_48 == 0) +uint8_t x_22; +x_22 = !lean_is_exclusive(x_17); +if (x_22 == 0) { -return x_22; +return x_17; } else { -lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_49 = lean_ctor_get(x_22, 0); -x_50 = lean_ctor_get(x_22, 1); -lean_inc(x_50); -lean_inc(x_49); -lean_dec(x_22); -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_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_17, 0); +x_24 = lean_ctor_get(x_17, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_17); +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 { -uint8_t x_52; -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_52 = !lean_is_exclusive(x_12); -if (x_52 == 0) -{ -return x_12; +lean_object* x_26; uint8_t x_27; lean_object* x_28; +x_26 = lean_box(0); +x_27 = lean_unbox(x_12); +lean_dec(x_12); +x_28 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_tryResolve___spec__2___lambda__4(x_3, x_2, x_1, x_4, x_10, x_27, x_26, x_5, x_6, x_7, x_8, x_14); +return x_28; +} } else { -lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_53 = lean_ctor_get(x_12, 0); -x_54 = lean_ctor_get(x_12, 1); -lean_inc(x_54); -lean_inc(x_53); +lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; +x_29 = lean_ctor_get(x_11, 1); +lean_inc(x_29); +lean_dec(x_11); +x_30 = lean_box(0); +x_31 = lean_unbox(x_12); lean_dec(x_12); -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; -} +x_32 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_tryResolve___spec__2___lambda__4(x_3, x_2, x_1, x_4, x_10, x_31, x_30, x_5, x_6, x_7, x_8, x_29); +return x_32; } } } -static lean_object* _init_l_Lean_Meta_SynthInstance_tryResolve___lambda__5___closed__1() { +static lean_object* _init_l_Lean_Meta_SynthInstance_tryResolve___lambda__1___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("tryResolve", 10); +x_1 = lean_mk_string_from_bytes(" ", 1); return x_1; } } -static lean_object* _init_l_Lean_Meta_SynthInstance_tryResolve___lambda__5___closed__2() { +static lean_object* _init_l_Lean_Meta_SynthInstance_tryResolve___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_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_6____closed__8; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_6____closed__1; -x_3 = l_Lean_Meta_SynthInstance_tryResolve___lambda__5___closed__1; -x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); -return x_4; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_SynthInstance_tryResolve___lambda__1___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_tryResolve___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) { +static lean_object* _init_l_Lean_Meta_SynthInstance_tryResolve___lambda__1___closed__3() { _start: { -lean_object* x_12; -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_5); -x_12 = l_Lean_Meta_SynthInstance_getSubgoals(x_1, x_2, x_5, x_3, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_12) == 0) +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(" ≟ ", 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_SynthInstance_tryResolve___lambda__1___closed__4() { +_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; uint8_t x_25; lean_object* x_26; -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_ctor_get(x_13, 0); -lean_inc(x_15); -x_16 = lean_ctor_get(x_13, 1); -lean_inc(x_16); -x_17 = lean_ctor_get(x_13, 2); -lean_inc(x_17); -lean_dec(x_13); -x_18 = lean_st_ref_get(x_8, x_14); -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); -lean_inc(x_6); -x_22 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_tryResolve___lambda__2), 9, 3); -lean_closure_set(x_22, 0, x_6); -lean_closure_set(x_22, 1, x_17); -lean_closure_set(x_22, 2, x_21); -x_23 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_tryResolve___lambda__4), 11, 6); -lean_closure_set(x_23, 0, x_6); -lean_closure_set(x_23, 1, x_17); -lean_closure_set(x_23, 2, x_5); -lean_closure_set(x_23, 3, x_16); -lean_closure_set(x_23, 4, x_4); -lean_closure_set(x_23, 5, x_15); -x_24 = l_Lean_Meta_SynthInstance_tryResolve___lambda__5___closed__2; -x_25 = 1; -x_26 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_tryResolve___spec__2(x_24, x_22, x_23, x_25, x_7, x_8, x_9, x_10, x_20); -return x_26; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_SynthInstance_tryResolve___lambda__1___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } -else +} +LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_tryResolve___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: { -uint8_t x_27; -lean_dec(x_10); +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_9 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_1, 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); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -x_27 = !lean_is_exclusive(x_12); -if (x_27 == 0) +x_12 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_2, x_4, x_5, x_6, x_7, x_11); +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; +x_14 = lean_ctor_get(x_12, 0); +x_15 = l_Lean_exceptOptionEmoji___rarg(x_3); +x_16 = l_Lean_stringToMessageData(x_15); +lean_dec(x_15); +x_17 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; +x_18 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_16); +x_19 = l_Lean_Meta_SynthInstance_tryResolve___lambda__1___closed__2; +x_20 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_20, 0, x_18); +lean_ctor_set(x_20, 1, x_19); +x_21 = l_Lean_MessageData_ofExpr(x_10); +x_22 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +x_23 = l_Lean_Meta_SynthInstance_tryResolve___lambda__1___closed__4; +x_24 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_24, 0, x_22); +lean_ctor_set(x_24, 1, x_23); +x_25 = l_Lean_MessageData_ofExpr(x_14); +x_26 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +x_27 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_17); +lean_ctor_set(x_12, 0, x_27); return x_12; } else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_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_28 = lean_ctor_get(x_12, 0); x_29 = lean_ctor_get(x_12, 1); lean_inc(x_29); lean_inc(x_28); lean_dec(x_12); -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; +x_30 = l_Lean_exceptOptionEmoji___rarg(x_3); +x_31 = l_Lean_stringToMessageData(x_30); +lean_dec(x_30); +x_32 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; +x_33 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_31); +x_34 = l_Lean_Meta_SynthInstance_tryResolve___lambda__1___closed__2; +x_35 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +x_36 = l_Lean_MessageData_ofExpr(x_10); +x_37 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_37, 0, x_35); +lean_ctor_set(x_37, 1, x_36); +x_38 = l_Lean_Meta_SynthInstance_tryResolve___lambda__1___closed__4; +x_39 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +x_40 = l_Lean_MessageData_ofExpr(x_28); +x_41 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_40); +x_42 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_32); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_29); +return x_43; } } } +LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_tryResolve___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) { +_start: +{ +lean_object* x_10; lean_object* x_11; +x_10 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_tryResolve___lambda__1___boxed), 8, 3); +lean_closure_set(x_10, 0, x_1); +lean_closure_set(x_10, 1, x_2); +lean_closure_set(x_10, 2, x_4); +x_11 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_tryResolve___spec__1___rarg(x_3, x_10, x_5, x_6, x_7, x_8, x_9); +return x_11; +} } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_tryResolve(lean_object* x_1, lean_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_Meta_SynthInstance_tryResolve___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) { _start: { -lean_object* x_8; -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -lean_inc(x_1); -x_8 = lean_infer_type(x_1, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_8) == 0) +lean_object* x_7; lean_object* x_8; +x_7 = lean_box(0); +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_7); +lean_ctor_set(x_8, 1, x_6); +return x_8; +} +} +static lean_object* _init_l_Lean_Meta_SynthInstance_tryResolve___lambda__4___closed__1() { +_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; -x_9 = lean_ctor_get(x_8, 0); -lean_inc(x_9); -x_10 = lean_ctor_get(x_8, 1); +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_tryResolve___lambda__3___boxed), 6, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_tryResolve___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) { +_start: +{ +lean_object* x_12; lean_inc(x_10); -lean_dec(x_8); -x_11 = lean_ctor_get(x_3, 1); -lean_inc(x_11); -x_12 = l_Lean_Meta_getLocalInstances(x_3, x_4, x_5, x_6, x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +x_12 = l_Lean_Meta_isExprDefEq(x_1, x_2, 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; 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_alloc_closure((void*)(l_Lean_Meta_SynthInstance_tryResolve___lambda__5), 11, 4); -lean_closure_set(x_15, 0, x_11); -lean_closure_set(x_15, 1, x_13); -lean_closure_set(x_15, 2, x_2); -lean_closure_set(x_15, 3, x_1); -x_16 = l_Lean_Meta_forallTelescopeReducing___at_Lean_Meta_getParamNames___spec__2___rarg(x_9, x_15, x_3, x_4, x_5, x_6, x_14); -return x_16; -} -else +x_15 = l_Lean_Meta_SynthInstance_tryResolve___lambda__4___closed__1; +x_16 = lean_unbox(x_13); +lean_dec(x_13); +if (x_16 == 0) { -uint8_t x_17; +lean_object* x_17; lean_object* x_18; 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_is_exclusive(x_8); -if (x_17 == 0) -{ -return x_8; +x_17 = lean_box(0); +x_18 = lean_apply_6(x_15, x_17, x_7, x_8, x_9, x_10, x_14); +return x_18; } else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_ctor_get(x_8, 0); -x_19 = lean_ctor_get(x_8, 1); -lean_inc(x_19); -lean_inc(x_18); -lean_dec(x_8); -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_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_tryResolve___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_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_tryResolve___spec__5(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_withTraceNode___at_Lean_Meta_SynthInstance_tryResolve___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) { -_start: +uint8_t x_19; uint8_t x_20; uint8_t x_21; lean_object* x_22; +x_19 = 0; +x_20 = 1; +x_21 = 1; +x_22 = l_Lean_Meta_mkLambdaFVars(x_3, x_4, x_19, x_20, x_21, x_7, x_8, x_9, x_10, x_14); +if (lean_obj_tag(x_22) == 0) { -uint8_t x_13; lean_object* x_14; -x_13 = lean_unbox(x_4); -lean_dec(x_4); -x_14 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_tryResolve___spec__2___lambda__2(x_1, x_2, x_3, x_13, 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_5); -return x_14; -} -} -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_tryResolve___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: +lean_object* x_23; lean_object* x_24; lean_object* 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); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +x_25 = l_Lean_Meta_isExprDefEq(x_5, x_23, x_7, x_8, x_9, x_10, x_24); +if (lean_obj_tag(x_25) == 0) { -uint8_t x_10; lean_object* x_11; -x_10 = lean_unbox(x_4); -lean_dec(x_4); -x_11 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_tryResolve___spec__2(x_1, x_2, x_3, x_10, x_5, x_6, x_7, x_8, x_9); -return x_11; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_tryResolve___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_26; uint8_t x_27; +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_9; -x_9 = l_Lean_Meta_SynthInstance_tryResolve___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_7); +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_9; -} +x_28 = lean_ctor_get(x_25, 1); +lean_inc(x_28); +lean_dec(x_25); +x_29 = lean_box(0); +x_30 = lean_apply_6(x_15, x_29, x_7, x_8, x_9, x_10, x_28); +return x_30; } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_tryResolve___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) { -_start: +else { -lean_object* x_7; -x_7 = l_Lean_Meta_SynthInstance_tryResolve___lambda__3(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_object* x_31; lean_object* x_32; uint8_t x_33; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_7); +x_31 = lean_ctor_get(x_25, 1); +lean_inc(x_31); +lean_dec(x_25); +x_32 = lean_st_ref_get(x_8, x_31); +lean_dec(x_8); +x_33 = !lean_is_exclusive(x_32); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +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_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_6); +x_37 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set(x_32, 0, x_37); +return x_32; } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_tryAnswer___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: +else { -lean_object* x_9; -x_9 = l_Lean_Meta_openAbstractMVarsResult(x_1, x_4, x_5, x_6, x_7, x_8); -return x_9; +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_32, 0); +x_39 = lean_ctor_get(x_32, 1); +lean_inc(x_39); +lean_inc(x_38); +lean_dec(x_32); +x_40 = lean_ctor_get(x_38, 0); +lean_inc(x_40); +lean_dec(x_38); +x_41 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_6); +x_42 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_42, 0, x_41); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_39); +return x_43; } } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_tryAnswer___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) { -_start: +} +else { -lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = lean_ctor_get(x_2, 1); -lean_inc(x_10); -lean_dec(x_2); -x_11 = lean_ctor_get(x_10, 1); -lean_inc(x_11); +uint8_t x_44; lean_dec(x_10); -lean_inc(x_6); -x_12 = l_Lean_Meta_isExprDefEq(x_1, x_11, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_12) == 0) -{ -lean_object* x_13; uint8_t x_14; -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_unbox(x_13); -lean_dec(x_13); -if (x_14 == 0) -{ -uint8_t x_15; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); lean_dec(x_6); -x_15 = !lean_is_exclusive(x_12); -if (x_15 == 0) +x_44 = !lean_is_exclusive(x_25); +if (x_44 == 0) { -lean_object* x_16; lean_object* x_17; -x_16 = lean_ctor_get(x_12, 0); -lean_dec(x_16); -x_17 = lean_box(0); -lean_ctor_set(x_12, 0, x_17); -return x_12; +return x_25; } else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_ctor_get(x_12, 1); -lean_inc(x_18); -lean_dec(x_12); -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_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_25, 0); +x_46 = lean_ctor_get(x_25, 1); +lean_inc(x_46); +lean_inc(x_45); +lean_dec(x_25); +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_21; lean_object* x_22; uint8_t x_23; -x_21 = lean_ctor_get(x_12, 1); -lean_inc(x_21); -lean_dec(x_12); -x_22 = lean_st_ref_get(x_6, x_21); +uint8_t x_48; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); lean_dec(x_6); -x_23 = !lean_is_exclusive(x_22); -if (x_23 == 0) +lean_dec(x_5); +x_48 = !lean_is_exclusive(x_22); +if (x_48 == 0) { -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_24, 0); -lean_inc(x_25); -lean_dec(x_24); -x_26 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_22, 0, x_26); return x_22; } else { -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_22, 0); -x_28 = lean_ctor_get(x_22, 1); -lean_inc(x_28); -lean_inc(x_27); +lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_49 = lean_ctor_get(x_22, 0); +x_50 = lean_ctor_get(x_22, 1); +lean_inc(x_50); +lean_inc(x_49); lean_dec(x_22); -x_29 = lean_ctor_get(x_27, 0); -lean_inc(x_29); -lean_dec(x_27); -x_30 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_30, 0, x_29); -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; +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 { -uint8_t x_32; +uint8_t x_52; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); lean_dec(x_6); -x_32 = !lean_is_exclusive(x_12); -if (x_32 == 0) +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_52 = !lean_is_exclusive(x_12); +if (x_52 == 0) { return x_12; } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_12, 0); -x_34 = lean_ctor_get(x_12, 1); -lean_inc(x_34); -lean_inc(x_33); +lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_53 = lean_ctor_get(x_12, 0); +x_54 = lean_ctor_get(x_12, 1); +lean_inc(x_54); +lean_inc(x_53); lean_dec(x_12); -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; +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_EXPORT lean_object* l_Lean_Meta_SynthInstance_tryAnswer(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, 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_Meta_SynthInstance_tryResolve___lambda__5___closed__1() { _start: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_11 = lean_ctor_get(x_3, 0); -lean_inc(x_11); -lean_dec(x_3); -x_12 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_tryAnswer___lambda__1___boxed), 8, 1); -lean_closure_set(x_12, 0, x_11); -x_13 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_tryAnswer___lambda__2___boxed), 9, 1); -lean_closure_set(x_13, 0, x_2); -x_14 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_SynthInstance_newSubgoal___spec__7___rarg), 9, 2); -lean_closure_set(x_14, 0, x_12); -lean_closure_set(x_14, 1, x_13); -x_15 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__18___rarg(x_1, x_14, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_15; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("tryResolve", 10); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_tryAnswer___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) { +static lean_object* _init_l_Lean_Meta_SynthInstance_tryResolve___lambda__5___closed__2() { _start: { -lean_object* x_9; -x_9 = l_Lean_Meta_SynthInstance_tryAnswer___lambda__1(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_3); -lean_dec(x_2); -return x_9; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_6____closed__8; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_6____closed__1; +x_3 = l_Lean_Meta_SynthInstance_tryResolve___lambda__5___closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_tryAnswer___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_Meta_SynthInstance_tryResolve___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) { _start: { -lean_object* x_10; -x_10 = l_Lean_Meta_SynthInstance_tryAnswer___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_4); -lean_dec(x_3); -return x_10; -} -} -LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Meta_SynthInstance_wakeUp___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_12; +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_5); +x_12 = l_Lean_Meta_SynthInstance_getSubgoals(x_1, x_2, x_5, x_3, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_12) == 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; uint8_t x_17; -x_10 = lean_ctor_get(x_7, 5); -x_11 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_2, 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_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; lean_object* x_26; +x_13 = lean_ctor_get(x_12, 0); lean_inc(x_13); -lean_dec(x_11); -x_14 = lean_st_ref_take(x_8, x_13); -x_15 = lean_ctor_get(x_14, 0); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = lean_ctor_get(x_13, 0); lean_inc(x_15); -x_16 = lean_ctor_get(x_14, 1); +x_16 = lean_ctor_get(x_13, 1); lean_inc(x_16); -lean_dec(x_14); -x_17 = !lean_is_exclusive(x_15); -if (x_17 == 0) +x_17 = lean_ctor_get(x_13, 2); +lean_inc(x_17); +lean_dec(x_13); +x_18 = lean_st_ref_get(x_8, x_14); +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); +lean_inc(x_6); +x_22 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_tryResolve___lambda__2), 9, 3); +lean_closure_set(x_22, 0, x_6); +lean_closure_set(x_22, 1, x_17); +lean_closure_set(x_22, 2, x_21); +x_23 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_tryResolve___lambda__4), 11, 6); +lean_closure_set(x_23, 0, x_6); +lean_closure_set(x_23, 1, x_17); +lean_closure_set(x_23, 2, x_5); +lean_closure_set(x_23, 3, x_16); +lean_closure_set(x_23, 4, x_4); +lean_closure_set(x_23, 5, x_15); +x_24 = l_Lean_Meta_SynthInstance_tryResolve___lambda__5___closed__2; +x_25 = 1; +x_26 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_tryResolve___spec__2(x_24, x_22, x_23, x_25, x_7, x_8, x_9, x_10, x_20); +return x_26; +} +else { -lean_object* x_18; lean_object* x_19; uint8_t 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_15, 3); -x_19 = l_Lean_Meta_SynthInstance_instInhabitedInstance___closed__2; -x_20 = 0; -x_21 = lean_alloc_ctor(9, 3, 1); -lean_ctor_set(x_21, 0, x_1); -lean_ctor_set(x_21, 1, x_12); -lean_ctor_set(x_21, 2, x_19); -lean_ctor_set_uint8(x_21, sizeof(void*)*3, x_20); -lean_inc(x_10); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_10); -lean_ctor_set(x_22, 1, x_21); -x_23 = l_Lean_PersistentArray_push___rarg(x_18, x_22); -lean_ctor_set(x_15, 3, x_23); -x_24 = lean_st_ref_set(x_8, x_15, x_16); -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) +uint8_t x_27; +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_27 = !lean_is_exclusive(x_12); +if (x_27 == 0) { -lean_object* x_26; lean_object* x_27; -x_26 = lean_ctor_get(x_24, 0); -lean_dec(x_26); -x_27 = lean_box(0); -lean_ctor_set(x_24, 0, x_27); -return x_24; +return x_12; } else { lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_24, 1); +x_28 = lean_ctor_get(x_12, 0); +x_29 = lean_ctor_get(x_12, 1); +lean_inc(x_29); lean_inc(x_28); -lean_dec(x_24); -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); +lean_dec(x_12); +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; } } -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; uint8_t x_39; lean_object* x_40; lean_object* x_41; 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_31 = lean_ctor_get(x_15, 0); -x_32 = lean_ctor_get(x_15, 1); -x_33 = lean_ctor_get(x_15, 2); -x_34 = lean_ctor_get(x_15, 3); -x_35 = lean_ctor_get(x_15, 4); -x_36 = lean_ctor_get(x_15, 5); -x_37 = lean_ctor_get(x_15, 6); -lean_inc(x_37); -lean_inc(x_36); -lean_inc(x_35); -lean_inc(x_34); -lean_inc(x_33); -lean_inc(x_32); -lean_inc(x_31); -lean_dec(x_15); -x_38 = l_Lean_Meta_SynthInstance_instInhabitedInstance___closed__2; -x_39 = 0; -x_40 = lean_alloc_ctor(9, 3, 1); -lean_ctor_set(x_40, 0, x_1); -lean_ctor_set(x_40, 1, x_12); -lean_ctor_set(x_40, 2, x_38); -lean_ctor_set_uint8(x_40, sizeof(void*)*3, x_39); -lean_inc(x_10); -x_41 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_41, 0, x_10); -lean_ctor_set(x_41, 1, x_40); -x_42 = l_Lean_PersistentArray_push___rarg(x_34, x_41); -x_43 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_43, 0, x_31); -lean_ctor_set(x_43, 1, x_32); -lean_ctor_set(x_43, 2, x_33); -lean_ctor_set(x_43, 3, x_42); -lean_ctor_set(x_43, 4, x_35); -lean_ctor_set(x_43, 5, x_36); -lean_ctor_set(x_43, 6, x_37); -x_44 = lean_st_ref_set(x_8, x_43, x_16); -x_45 = lean_ctor_get(x_44, 1); -lean_inc(x_45); -if (lean_is_exclusive(x_44)) { - lean_ctor_release(x_44, 0); - lean_ctor_release(x_44, 1); - x_46 = x_44; -} else { - lean_dec_ref(x_44); - x_46 = lean_box(0); -} -x_47 = lean_box(0); -if (lean_is_scalar(x_46)) { - x_48 = lean_alloc_ctor(0, 2, 0); -} else { - x_48 = x_46; -} -lean_ctor_set(x_48, 0, x_47); -lean_ctor_set(x_48, 1, x_45); -return x_48; -} } } -static lean_object* _init_l_Lean_Meta_SynthInstance_wakeUp___closed__1() { +LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_tryResolve(lean_object* x_1, 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_1; -x_1 = lean_mk_string_from_bytes("skip answer containing metavariables ", 37); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_SynthInstance_wakeUp___closed__2() { -_start: +lean_object* x_8; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_1); +x_8 = lean_infer_type(x_1, x_3, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_8) == 0) { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_SynthInstance_wakeUp___closed__1; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} +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_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_8, 1); +lean_inc(x_10); +lean_dec(x_8); +x_11 = lean_ctor_get(x_3, 1); +lean_inc(x_11); +x_12 = l_Lean_Meta_getLocalInstances(x_3, x_4, x_5, x_6, x_10); +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_alloc_closure((void*)(l_Lean_Meta_SynthInstance_tryResolve___lambda__5), 11, 4); +lean_closure_set(x_15, 0, x_11); +lean_closure_set(x_15, 1, x_13); +lean_closure_set(x_15, 2, x_2); +lean_closure_set(x_15, 3, x_1); +x_16 = l_Lean_Meta_forallTelescopeReducing___at_Lean_Meta_getParamNames___spec__2___rarg(x_9, x_15, x_3, x_4, x_5, x_6, x_14); +return x_16; } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_wakeUp(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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_2) == 0) +else { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +uint8_t x_17; +lean_dec(x_6); lean_dec(x_5); -x_10 = lean_ctor_get(x_2, 0); -lean_inc(x_10); +lean_dec(x_4); +lean_dec(x_3); lean_dec(x_2); -x_11 = lean_st_ref_take(x_4, 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); -x_14 = !lean_is_exclusive(x_12); -if (x_14 == 0) -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_15 = lean_ctor_get(x_12, 2); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_10); -lean_ctor_set(x_16, 1, x_1); -x_17 = lean_array_push(x_15, x_16); -lean_ctor_set(x_12, 2, x_17); -x_18 = lean_st_ref_set(x_4, x_12, x_13); -x_19 = !lean_is_exclusive(x_18); -if (x_19 == 0) -{ -lean_object* x_20; lean_object* x_21; -x_20 = lean_ctor_get(x_18, 0); -lean_dec(x_20); -x_21 = lean_box(0); -lean_ctor_set(x_18, 0, x_21); -return x_18; -} -else +lean_dec(x_1); +x_17 = !lean_is_exclusive(x_8); +if (x_17 == 0) { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_18, 1); -lean_inc(x_22); -lean_dec(x_18); -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; -} +return x_8; } 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; -x_25 = lean_ctor_get(x_12, 0); -x_26 = lean_ctor_get(x_12, 1); -x_27 = lean_ctor_get(x_12, 2); -x_28 = lean_ctor_get(x_12, 3); -lean_inc(x_28); -lean_inc(x_27); -lean_inc(x_26); -lean_inc(x_25); -lean_dec(x_12); -x_29 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_29, 0, x_10); -lean_ctor_set(x_29, 1, x_1); -x_30 = lean_array_push(x_27, x_29); -x_31 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_31, 0, x_25); -lean_ctor_set(x_31, 1, x_26); -lean_ctor_set(x_31, 2, x_30); -lean_ctor_set(x_31, 3, x_28); -x_32 = lean_st_ref_set(x_4, x_31, x_13); -x_33 = lean_ctor_get(x_32, 1); -lean_inc(x_33); -if (lean_is_exclusive(x_32)) { - lean_ctor_release(x_32, 0); - lean_ctor_release(x_32, 1); - x_34 = x_32; -} else { - lean_dec_ref(x_32); - x_34 = lean_box(0); +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_8, 0); +x_19 = lean_ctor_get(x_8, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_8); +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; } -x_35 = lean_box(0); -if (lean_is_scalar(x_34)) { - x_36 = lean_alloc_ctor(0, 2, 0); -} else { - x_36 = x_34; } -lean_ctor_set(x_36, 0, x_35); -lean_ctor_set(x_36, 1, x_33); -return x_36; } } -else +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_tryResolve___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: { -lean_object* x_37; lean_object* x_38; lean_object* x_39; uint8_t x_40; -x_37 = lean_ctor_get(x_1, 0); -lean_inc(x_37); +lean_object* x_7; +x_7 = l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_tryResolve___spec__4(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_38 = lean_ctor_get(x_37, 1); -lean_inc(x_38); -x_39 = lean_unsigned_to_nat(0u); -x_40 = lean_nat_dec_eq(x_38, x_39); -lean_dec(x_38); -if (x_40 == 0) +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_tryResolve___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) { +_start: { -lean_object* x_41; 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; uint8_t x_49; -lean_inc(x_5); -x_41 = l_Lean_Meta_openAbstractMVarsResult(x_37, x_5, x_6, x_7, x_8, x_9); -x_42 = lean_ctor_get(x_41, 0); -lean_inc(x_42); -x_43 = lean_ctor_get(x_42, 1); -lean_inc(x_43); -lean_dec(x_42); -x_44 = lean_ctor_get(x_41, 1); -lean_inc(x_44); -lean_dec(x_41); -x_45 = lean_ctor_get(x_43, 1); -lean_inc(x_45); -lean_dec(x_43); -x_46 = l_Lean_Meta_SynthInstance_newSubgoal___closed__1; -x_47 = l_Lean_isTracingEnabledFor___at_Lean_Meta_SynthInstance_newSubgoal___spec__10(x_46, x_3, x_4, x_5, x_6, x_7, x_8, x_44); -x_48 = lean_ctor_get(x_47, 0); -lean_inc(x_48); -x_49 = lean_unbox(x_48); -lean_dec(x_48); -if (x_49 == 0) +uint8_t x_13; lean_object* x_14; +x_13 = lean_unbox(x_4); +lean_dec(x_4); +x_14 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_tryResolve___spec__2___lambda__2(x_1, x_2, x_3, x_13, 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_5); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_tryResolve___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_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: { -uint8_t x_50; -lean_dec(x_45); +uint8_t x_15; uint8_t x_16; double x_17; lean_object* x_18; +x_15 = lean_unbox(x_5); lean_dec(x_5); -x_50 = !lean_is_exclusive(x_47); -if (x_50 == 0) +x_16 = lean_unbox(x_7); +lean_dec(x_7); +x_17 = lean_unbox_float(x_8); +lean_dec(x_8); +x_18 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_tryResolve___spec__2___lambda__3(x_1, x_2, x_3, x_4, x_15, x_6, x_16, x_17, x_9, x_10, x_11, x_12, x_13, x_14); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_tryResolve___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, lean_object* x_11, lean_object* x_12) { +_start: { -lean_object* x_51; lean_object* x_52; -x_51 = lean_ctor_get(x_47, 0); -lean_dec(x_51); -x_52 = lean_box(0); -lean_ctor_set(x_47, 0, x_52); -return x_47; +uint8_t x_13; uint8_t x_14; lean_object* x_15; +x_13 = lean_unbox(x_4); +lean_dec(x_4); +x_14 = lean_unbox(x_6); +lean_dec(x_6); +x_15 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_tryResolve___spec__2___lambda__4(x_1, x_2, x_3, x_13, x_5, x_14, x_7, x_8, x_9, x_10, x_11, x_12); +return x_15; } -else +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_tryResolve___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: { -lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_53 = lean_ctor_get(x_47, 1); -lean_inc(x_53); -lean_dec(x_47); -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_53); -return x_55; +uint8_t x_10; lean_object* x_11; +x_10 = lean_unbox(x_4); +lean_dec(x_4); +x_11 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_tryResolve___spec__2(x_1, x_2, x_3, x_10, x_5, x_6, x_7, x_8, x_9); +return x_11; } } -else +LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_tryResolve___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_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_56 = lean_ctor_get(x_47, 1); -lean_inc(x_56); -lean_dec(x_47); -x_57 = l_Lean_MessageData_ofExpr(x_45); -x_58 = l_Lean_Meta_SynthInstance_wakeUp___closed__2; -x_59 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_59, 0, x_58); -lean_ctor_set(x_59, 1, x_57); -x_60 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; -x_61 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_61, 0, x_59); -lean_ctor_set(x_61, 1, x_60); -x_62 = l_Lean_addTrace___at_Lean_Meta_SynthInstance_wakeUp___spec__1(x_46, x_61, x_3, x_4, x_5, x_6, x_7, x_8, x_56); +lean_object* x_9; +x_9 = l_Lean_Meta_SynthInstance_tryResolve___lambda__1(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); -return x_62; +lean_dec(x_4); +lean_dec(x_3); +return x_9; } } -else +LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_tryResolve___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) { +_start: { -lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; +lean_object* x_7; +x_7 = l_Lean_Meta_SynthInstance_tryResolve___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6); lean_dec(x_5); -x_63 = lean_st_ref_take(x_4, x_9); -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_alloc_ctor(1, 1, 0); -lean_ctor_set(x_66, 0, x_37); -x_67 = !lean_is_exclusive(x_64); -if (x_67 == 0) -{ -lean_object* x_68; lean_object* x_69; uint8_t x_70; -x_68 = lean_ctor_get(x_64, 0); -lean_dec(x_68); -lean_ctor_set(x_64, 0, x_66); -x_69 = lean_st_ref_set(x_4, x_64, x_65); -x_70 = !lean_is_exclusive(x_69); -if (x_70 == 0) +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_Meta_SynthInstance_tryAnswer___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_71; lean_object* x_72; -x_71 = lean_ctor_get(x_69, 0); -lean_dec(x_71); -x_72 = lean_box(0); -lean_ctor_set(x_69, 0, x_72); -return x_69; +lean_object* x_9; +x_9 = l_Lean_Meta_openAbstractMVarsResult(x_1, x_4, x_5, x_6, x_7, x_8); +return x_9; } -else +} +LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_tryAnswer___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) { +_start: { -lean_object* x_73; lean_object* x_74; lean_object* x_75; -x_73 = lean_ctor_get(x_69, 1); -lean_inc(x_73); -lean_dec(x_69); -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_73); -return x_75; -} -} -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; -x_76 = lean_ctor_get(x_64, 1); -x_77 = lean_ctor_get(x_64, 2); -x_78 = lean_ctor_get(x_64, 3); -lean_inc(x_78); -lean_inc(x_77); -lean_inc(x_76); -lean_dec(x_64); -x_79 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_79, 0, x_66); -lean_ctor_set(x_79, 1, x_76); -lean_ctor_set(x_79, 2, x_77); -lean_ctor_set(x_79, 3, x_78); -x_80 = lean_st_ref_set(x_4, x_79, x_65); -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_EXPORT lean_object* l_Lean_addTrace___at_Lean_Meta_SynthInstance_wakeUp___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) { -_start: -{ -lean_object* x_10; -x_10 = l_Lean_addTrace___at_Lean_Meta_SynthInstance_wakeUp___spec__1(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); -return x_10; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_wakeUp___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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_Meta_SynthInstance_wakeUp(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_4); -lean_dec(x_3); -return x_10; -} -} -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_SynthInstance_isNewAnswer___spec__1(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4) { -_start: +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = lean_ctor_get(x_2, 1); +lean_inc(x_10); +lean_dec(x_2); +x_11 = lean_ctor_get(x_10, 1); +lean_inc(x_11); +lean_dec(x_10); +lean_inc(x_6); +x_12 = l_Lean_Meta_isExprDefEq(x_1, x_11, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_12) == 0) { -uint8_t x_5; -x_5 = lean_usize_dec_eq(x_3, x_4); -if (x_5 == 0) +lean_object* x_13; uint8_t x_14; +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_unbox(x_13); +lean_dec(x_13); +if (x_14 == 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 = lean_ctor_get(x_6, 1); -lean_inc(x_7); +uint8_t x_15; lean_dec(x_6); -x_8 = lean_ctor_get(x_1, 1); -x_9 = lean_expr_eqv(x_7, x_8); -lean_dec(x_7); -if (x_9 == 0) +x_15 = !lean_is_exclusive(x_12); +if (x_15 == 0) { -size_t x_10; size_t x_11; -x_10 = 1; -x_11 = lean_usize_add(x_3, x_10); -x_3 = x_11; -goto _start; +lean_object* x_16; lean_object* x_17; +x_16 = lean_ctor_get(x_12, 0); +lean_dec(x_16); +x_17 = lean_box(0); +lean_ctor_set(x_12, 0, x_17); +return x_12; } else { -uint8_t x_13; -x_13 = 1; -return x_13; +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_12, 1); +lean_inc(x_18); +lean_dec(x_12); +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 { -uint8_t x_14; -x_14 = 0; -return x_14; -} -} -} -LEAN_EXPORT uint8_t l_Lean_Meta_SynthInstance_isNewAnswer(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; lean_object* x_4; uint8_t x_5; -x_3 = lean_array_get_size(x_1); -x_4 = lean_unsigned_to_nat(0u); -x_5 = lean_nat_dec_lt(x_4, x_3); -if (x_5 == 0) +lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_21 = lean_ctor_get(x_12, 1); +lean_inc(x_21); +lean_dec(x_12); +x_22 = lean_st_ref_get(x_6, x_21); +lean_dec(x_6); +x_23 = !lean_is_exclusive(x_22); +if (x_23 == 0) { -uint8_t x_6; -lean_dec(x_3); -x_6 = 1; -return x_6; +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_24, 0); +lean_inc(x_25); +lean_dec(x_24); +x_26 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_22, 0, x_26); +return x_22; } else { -uint8_t x_7; -x_7 = lean_nat_dec_le(x_3, x_3); -if (x_7 == 0) -{ -uint8_t x_8; -lean_dec(x_3); -x_8 = 1; -return x_8; +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_22, 0); +x_28 = lean_ctor_get(x_22, 1); +lean_inc(x_28); +lean_inc(x_27); +lean_dec(x_22); +x_29 = lean_ctor_get(x_27, 0); +lean_inc(x_29); +lean_dec(x_27); +x_30 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_30, 0, x_29); +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 { -size_t x_9; size_t x_10; uint8_t x_11; -x_9 = 0; -x_10 = lean_usize_of_nat(x_3); -lean_dec(x_3); -x_11 = l_Array_anyMUnsafe_any___at_Lean_Meta_SynthInstance_isNewAnswer___spec__1(x_2, x_1, x_9, x_10); -if (x_11 == 0) +uint8_t x_32; +lean_dec(x_6); +x_32 = !lean_is_exclusive(x_12); +if (x_32 == 0) { -uint8_t x_12; -x_12 = 1; return x_12; } else { -uint8_t x_13; -x_13 = 0; -return x_13; -} +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_12, 0); +x_34 = lean_ctor_get(x_12, 1); +lean_inc(x_34); +lean_inc(x_33); +lean_dec(x_12); +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_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Meta_SynthInstance_isNewAnswer___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_Meta_SynthInstance_tryAnswer(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, 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_5; size_t x_6; uint8_t x_7; lean_object* x_8; -x_5 = lean_unbox_usize(x_3); +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_11 = lean_ctor_get(x_3, 0); +lean_inc(x_11); lean_dec(x_3); -x_6 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_7 = l_Array_anyMUnsafe_any___at_Lean_Meta_SynthInstance_isNewAnswer___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; +x_12 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_tryAnswer___lambda__1___boxed), 8, 1); +lean_closure_set(x_12, 0, x_11); +x_13 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_tryAnswer___lambda__2___boxed), 9, 1); +lean_closure_set(x_13, 0, x_2); +x_14 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_SynthInstance_newSubgoal___spec__7___rarg), 9, 2); +lean_closure_set(x_14, 0, x_12); +lean_closure_set(x_14, 1, x_13); +x_15 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__16___rarg(x_1, x_14, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_15; } } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_isNewAnswer___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_tryAnswer___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: { -uint8_t x_3; lean_object* x_4; -x_3 = l_Lean_Meta_SynthInstance_isNewAnswer(x_1, x_2); +lean_object* x_9; +x_9 = l_Lean_Meta_SynthInstance_tryAnswer___lambda__1(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_3); lean_dec(x_2); -lean_dec(x_1); -x_4 = lean_box(x_3); -return x_4; +return x_9; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___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_Meta_SynthInstance_tryAnswer___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) { _start: { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_9 = l_Lean_Meta_abstractMVars(x_1, 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 = lean_ctor_get(x_10, 2); -lean_inc(x_12); -x_13 = lean_infer_type(x_12, x_4, x_5, x_6, x_7, x_11); -if (lean_obj_tag(x_13) == 0) -{ -uint8_t x_14; -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; -x_15 = lean_ctor_get(x_13, 0); -x_16 = lean_ctor_get(x_2, 4); -x_17 = lean_unsigned_to_nat(1u); -x_18 = lean_nat_add(x_16, x_17); -x_19 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_19, 0, x_10); -lean_ctor_set(x_19, 1, x_15); -lean_ctor_set(x_19, 2, x_18); -lean_ctor_set(x_13, 0, x_19); -return x_13; -} -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_13, 0); -x_21 = lean_ctor_get(x_13, 1); -lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_13); -x_22 = lean_ctor_get(x_2, 4); -x_23 = lean_unsigned_to_nat(1u); -x_24 = lean_nat_add(x_22, x_23); -x_25 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_25, 0, x_10); -lean_ctor_set(x_25, 1, x_20); -lean_ctor_set(x_25, 2, x_24); -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_21); -return x_26; +lean_object* x_10; +x_10 = l_Lean_Meta_SynthInstance_tryAnswer___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_4); +lean_dec(x_3); +return x_10; } } -else +LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Meta_SynthInstance_wakeUp___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_27; -lean_dec(x_10); -x_27 = !lean_is_exclusive(x_13); -if (x_27 == 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; uint8_t x_17; +x_10 = lean_ctor_get(x_7, 5); +x_11 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_2, 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); +x_14 = lean_st_ref_take(x_8, x_13); +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_is_exclusive(x_15); +if (x_17 == 0) { -return x_13; +lean_object* x_18; lean_object* x_19; uint8_t 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_15, 3); +x_19 = l_Lean_Meta_SynthInstance_instInhabitedInstance___closed__2; +x_20 = 0; +x_21 = lean_alloc_ctor(9, 3, 1); +lean_ctor_set(x_21, 0, x_1); +lean_ctor_set(x_21, 1, x_12); +lean_ctor_set(x_21, 2, x_19); +lean_ctor_set_uint8(x_21, sizeof(void*)*3, x_20); +lean_inc(x_10); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_10); +lean_ctor_set(x_22, 1, x_21); +x_23 = l_Lean_PersistentArray_push___rarg(x_18, x_22); +lean_ctor_set(x_15, 3, x_23); +x_24 = lean_st_ref_set(x_8, x_15, x_16); +x_25 = !lean_is_exclusive(x_24); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; +x_26 = lean_ctor_get(x_24, 0); +lean_dec(x_26); +x_27 = lean_box(0); +lean_ctor_set(x_24, 0, x_27); +return x_24; } 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); +x_28 = lean_ctor_get(x_24, 1); 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); +lean_dec(x_24); +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; } } -} -} -static lean_object* _init_l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___lambda__2___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("newAnswer", 9); -return x_1; -} -} -static lean_object* _init_l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___lambda__2___closed__2() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_6____closed__8; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_6____closed__1; -x_3 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___lambda__2___closed__1; -x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); -return x_4; -} +lean_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; lean_object* x_40; lean_object* x_41; 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_31 = lean_ctor_get(x_15, 0); +x_32 = lean_ctor_get(x_15, 1); +x_33 = lean_ctor_get(x_15, 2); +x_34 = lean_ctor_get(x_15, 3); +x_35 = lean_ctor_get(x_15, 4); +x_36 = lean_ctor_get(x_15, 5); +x_37 = lean_ctor_get(x_15, 6); +lean_inc(x_37); +lean_inc(x_36); +lean_inc(x_35); +lean_inc(x_34); +lean_inc(x_33); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_15); +x_38 = l_Lean_Meta_SynthInstance_instInhabitedInstance___closed__2; +x_39 = 0; +x_40 = lean_alloc_ctor(9, 3, 1); +lean_ctor_set(x_40, 0, x_1); +lean_ctor_set(x_40, 1, x_12); +lean_ctor_set(x_40, 2, x_38); +lean_ctor_set_uint8(x_40, sizeof(void*)*3, x_39); +lean_inc(x_10); +x_41 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_41, 0, x_10); +lean_ctor_set(x_41, 1, x_40); +x_42 = l_Lean_PersistentArray_push___rarg(x_34, x_41); +x_43 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_43, 0, x_31); +lean_ctor_set(x_43, 1, x_32); +lean_ctor_set(x_43, 2, x_33); +lean_ctor_set(x_43, 3, x_42); +lean_ctor_set(x_43, 4, x_35); +lean_ctor_set(x_43, 5, x_36); +lean_ctor_set(x_43, 6, x_37); +x_44 = lean_st_ref_set(x_8, x_43, x_16); +x_45 = lean_ctor_get(x_44, 1); +lean_inc(x_45); +if (lean_is_exclusive(x_44)) { + lean_ctor_release(x_44, 0); + lean_ctor_release(x_44, 1); + x_46 = x_44; +} else { + lean_dec_ref(x_44); + x_46 = lean_box(0); } -static lean_object* _init_l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___lambda__2___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("size: ", 6); -return x_1; +x_47 = lean_box(0); +if (lean_is_scalar(x_46)) { + x_48 = lean_alloc_ctor(0, 2, 0); +} else { + x_48 = x_46; } +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set(x_48, 1, x_45); +return x_48; } -static lean_object* _init_l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___lambda__2___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___lambda__2___closed__3; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; } } -static lean_object* _init_l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___lambda__2___closed__5() { +static lean_object* _init_l_Lean_Meta_SynthInstance_wakeUp___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes(", val: ", 7); +x_1 = lean_mk_string_from_bytes("skip answer containing metavariables ", 37); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___lambda__2___closed__6() { +static lean_object* _init_l_Lean_Meta_SynthInstance_wakeUp___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___lambda__2___closed__5; +x_1 = l_Lean_Meta_SynthInstance_wakeUp___closed__1; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___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_Meta_SynthInstance_wakeUp(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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_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_8 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_1, x_3, x_4, x_5, x_6, x_7); -x_9 = lean_ctor_get(x_8, 0); -lean_inc(x_9); -x_10 = lean_ctor_get(x_8, 1); +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +lean_dec(x_5); +x_10 = lean_ctor_get(x_2, 0); lean_inc(x_10); -lean_dec(x_8); -x_11 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___lambda__2___closed__2; -x_12 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_11, x_3, x_4, x_5, x_6, x_10); -x_13 = lean_ctor_get(x_12, 0); +lean_dec(x_2); +x_11 = lean_st_ref_take(x_4, 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); -x_14 = lean_unbox(x_13); -lean_dec(x_13); +lean_dec(x_11); +x_14 = !lean_is_exclusive(x_12); if (x_14 == 0) { -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_12, 1); -lean_inc(x_15); -lean_dec(x_12); -x_16 = lean_box(0); -x_17 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___lambda__1(x_9, x_2, x_16, x_3, x_4, x_5, x_6, x_15); -lean_dec(x_2); -return x_17; -} -else +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_15 = lean_ctor_get(x_12, 2); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_10); +lean_ctor_set(x_16, 1, x_1); +x_17 = lean_array_push(x_15, x_16); +lean_ctor_set(x_12, 2, x_17); +x_18 = lean_st_ref_set(x_4, x_12, x_13); +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 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; -x_18 = lean_ctor_get(x_12, 1); -lean_inc(x_18); -lean_dec(x_12); -x_19 = lean_ctor_get(x_2, 4); -lean_inc(x_19); -x_20 = l_Nat_repr(x_19); -x_21 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_21, 0, x_20); -x_22 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_22, 0, x_21); -x_23 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___lambda__2___closed__4; -x_24 = lean_alloc_ctor(7, 2, 0); +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_18, 0); +lean_dec(x_20); +x_21 = lean_box(0); +lean_ctor_set(x_18, 0, x_21); +return x_18; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_18, 1); +lean_inc(x_22); +lean_dec(x_18); +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); -x_25 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___lambda__2___closed__6; -x_26 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_26, 0, x_24); -lean_ctor_set(x_26, 1, x_25); -lean_inc(x_9); -x_27 = l_Lean_MessageData_ofExpr(x_9); -x_28 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -x_29 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; -x_30 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_30, 0, x_28); -lean_ctor_set(x_30, 1, x_29); -x_31 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_11, x_30, x_3, x_4, x_5, x_6, x_18); -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_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___lambda__1(x_9, x_2, x_32, x_3, x_4, x_5, x_6, x_33); -lean_dec(x_32); -lean_dec(x_2); -return x_34; -} +return x_24; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer(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: +else { -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_7 = lean_ctor_get(x_1, 2); -lean_inc(x_7); -x_8 = lean_ctor_get(x_1, 0); -lean_inc(x_8); -x_9 = lean_alloc_closure((void*)(l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___lambda__2), 7, 2); -lean_closure_set(x_9, 0, x_8); -lean_closure_set(x_9, 1, x_1); -x_10 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_tryResolve___spec__1___rarg(x_7, x_9, x_2, x_3, x_4, x_5, x_6); -return x_10; +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; 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_25 = lean_ctor_get(x_12, 0); +x_26 = lean_ctor_get(x_12, 1); +x_27 = lean_ctor_get(x_12, 2); +x_28 = lean_ctor_get(x_12, 3); +lean_inc(x_28); +lean_inc(x_27); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_12); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_10); +lean_ctor_set(x_29, 1, x_1); +x_30 = lean_array_push(x_27, x_29); +x_31 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_31, 0, x_25); +lean_ctor_set(x_31, 1, x_26); +lean_ctor_set(x_31, 2, x_30); +lean_ctor_set(x_31, 3, x_28); +x_32 = lean_st_ref_set(x_4, x_31, x_13); +x_33 = lean_ctor_get(x_32, 1); +lean_inc(x_33); +if (lean_is_exclusive(x_32)) { + lean_ctor_release(x_32, 0); + lean_ctor_release(x_32, 1); + x_34 = x_32; +} else { + lean_dec_ref(x_32); + x_34 = lean_box(0); } +x_35 = lean_box(0); +if (lean_is_scalar(x_34)) { + x_36 = lean_alloc_ctor(0, 2, 0); +} else { + x_36 = x_34; } -LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___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___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_3); -lean_dec(x_2); -return x_9; +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_33); +return x_36; } } -LEAN_EXPORT lean_object* l_ReaderT_read___at_Lean_Meta_SynthInstance_addAnswer___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) { -_start: +else { -lean_object* x_8; -x_8 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_8, 0, x_1); -lean_ctor_set(x_8, 1, x_7); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_addAnswer___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, lean_object* x_11, lean_object* x_12) { -_start: +lean_object* x_37; lean_object* x_38; lean_object* x_39; uint8_t x_40; +x_37 = lean_ctor_get(x_1, 0); +lean_inc(x_37); +lean_dec(x_1); +x_38 = lean_ctor_get(x_37, 1); +lean_inc(x_38); +x_39 = lean_unsigned_to_nat(0u); +x_40 = lean_nat_dec_eq(x_38, x_39); +lean_dec(x_38); +if (x_40 == 0) { -uint8_t x_13; -x_13 = lean_usize_dec_eq(x_3, x_4); -if (x_13 == 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; uint8_t x_49; +lean_inc(x_5); +x_41 = l_Lean_Meta_openAbstractMVarsResult(x_37, x_5, x_6, x_7, x_8, x_9); +x_42 = lean_ctor_get(x_41, 0); +lean_inc(x_42); +x_43 = lean_ctor_get(x_42, 1); +lean_inc(x_43); +lean_dec(x_42); +x_44 = lean_ctor_get(x_41, 1); +lean_inc(x_44); +lean_dec(x_41); +x_45 = lean_ctor_get(x_43, 1); +lean_inc(x_45); +lean_dec(x_43); +x_46 = l_Lean_Meta_SynthInstance_newSubgoal___closed__1; +x_47 = l_Lean_isTracingEnabledFor___at_Lean_Meta_SynthInstance_newSubgoal___spec__10(x_46, x_3, x_4, x_5, x_6, x_7, x_8, x_44); +x_48 = lean_ctor_get(x_47, 0); +lean_inc(x_48); +x_49 = lean_unbox(x_48); +lean_dec(x_48); +if (x_49 == 0) { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; size_t x_18; size_t x_19; +uint8_t x_50; +lean_dec(x_45); lean_dec(x_5); -x_14 = lean_array_uget(x_2, x_3); -lean_inc(x_8); -lean_inc(x_1); -x_15 = l_Lean_Meta_SynthInstance_wakeUp(x_1, x_14, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -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 = 1; -x_19 = lean_usize_add(x_3, x_18); -x_3 = x_19; -x_5 = x_16; -x_12 = x_17; -goto _start; +x_50 = !lean_is_exclusive(x_47); +if (x_50 == 0) +{ +lean_object* x_51; lean_object* x_52; +x_51 = lean_ctor_get(x_47, 0); +lean_dec(x_51); +x_52 = lean_box(0); +lean_ctor_set(x_47, 0, x_52); +return x_47; } else { -lean_object* x_21; -lean_dec(x_8); -lean_dec(x_1); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_5); -lean_ctor_set(x_21, 1, x_12); -return x_21; -} +lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_53 = lean_ctor_get(x_47, 1); +lean_inc(x_53); +lean_dec(x_47); +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_53); +return x_55; } } -LEAN_EXPORT lean_object* l_Lean_withSeconds___at_Lean_Meta_SynthInstance_addAnswer___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) { -_start: -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_9 = lean_io_mono_nanos_now(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_apply_7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_11); -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_io_mono_nanos_now(x_14); -x_16 = !lean_is_exclusive(x_15); -if (x_16 == 0) +else { -lean_object* x_17; lean_object* x_18; uint8_t x_19; lean_object* x_20; double x_21; double x_22; double x_23; lean_object* x_24; lean_object* x_25; -x_17 = lean_ctor_get(x_15, 0); -x_18 = lean_nat_sub(x_17, x_10); -lean_dec(x_10); -lean_dec(x_17); -x_19 = 0; -x_20 = lean_unsigned_to_nat(0u); -x_21 = l_Float_ofScientific(x_18, x_19, x_20); -lean_dec(x_18); -x_22 = l_Lean_withSeconds___at_Lean_Meta_SynthInstance_newSubgoal___spec__14___closed__1; -x_23 = lean_float_div(x_21, x_22); -x_24 = lean_box_float(x_23); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_13); -lean_ctor_set(x_25, 1, x_24); -lean_ctor_set(x_15, 0, x_25); -return x_15; +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_56 = lean_ctor_get(x_47, 1); +lean_inc(x_56); +lean_dec(x_47); +x_57 = l_Lean_MessageData_ofExpr(x_45); +x_58 = l_Lean_Meta_SynthInstance_wakeUp___closed__2; +x_59 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_59, 0, x_58); +lean_ctor_set(x_59, 1, x_57); +x_60 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; +x_61 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_61, 0, x_59); +lean_ctor_set(x_61, 1, x_60); +x_62 = l_Lean_addTrace___at_Lean_Meta_SynthInstance_wakeUp___spec__1(x_46, x_61, x_3, x_4, x_5, x_6, x_7, x_8, x_56); +lean_dec(x_5); +return x_62; +} } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; lean_object* x_30; double x_31; double x_32; double x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_26 = lean_ctor_get(x_15, 0); -x_27 = lean_ctor_get(x_15, 1); -lean_inc(x_27); -lean_inc(x_26); -lean_dec(x_15); -x_28 = lean_nat_sub(x_26, x_10); -lean_dec(x_10); -lean_dec(x_26); -x_29 = 0; -x_30 = lean_unsigned_to_nat(0u); -x_31 = l_Float_ofScientific(x_28, x_29, x_30); -lean_dec(x_28); -x_32 = l_Lean_withSeconds___at_Lean_Meta_SynthInstance_newSubgoal___spec__14___closed__1; -x_33 = lean_float_div(x_31, x_32); -x_34 = lean_box_float(x_33); -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_13); -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_27); -return x_36; -} +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; +lean_dec(x_5); +x_63 = lean_st_ref_take(x_4, x_9); +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_alloc_ctor(1, 1, 0); +lean_ctor_set(x_66, 0, x_37); +x_67 = !lean_is_exclusive(x_64); +if (x_67 == 0) +{ +lean_object* x_68; lean_object* x_69; uint8_t x_70; +x_68 = lean_ctor_get(x_64, 0); +lean_dec(x_68); +lean_ctor_set(x_64, 0, x_66); +x_69 = lean_st_ref_set(x_4, x_64, x_65); +x_70 = !lean_is_exclusive(x_69); +if (x_70 == 0) +{ +lean_object* x_71; lean_object* x_72; +x_71 = lean_ctor_get(x_69, 0); +lean_dec(x_71); +x_72 = lean_box(0); +lean_ctor_set(x_69, 0, x_72); +return x_69; } else { -uint8_t x_37; -lean_dec(x_10); -x_37 = !lean_is_exclusive(x_12); -if (x_37 == 0) -{ -return x_12; +lean_object* x_73; lean_object* x_74; lean_object* x_75; +x_73 = lean_ctor_get(x_69, 1); +lean_inc(x_73); +lean_dec(x_69); +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_73); +return x_75; +} } 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_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; +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_76 = lean_ctor_get(x_64, 1); +x_77 = lean_ctor_get(x_64, 2); +x_78 = lean_ctor_get(x_64, 3); +lean_inc(x_78); +lean_inc(x_77); +lean_inc(x_76); +lean_dec(x_64); +x_79 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_79, 0, x_66); +lean_ctor_set(x_79, 1, x_76); +lean_ctor_set(x_79, 2, x_77); +lean_ctor_set(x_79, 3, x_78); +x_80 = lean_st_ref_set(x_4, x_79, x_65); +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_EXPORT lean_object* l_Lean_withOptProfile___at_Lean_Meta_SynthInstance_addAnswer___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; uint8_t x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_9 = l_Lean_withOptProfile___at_Lean_Meta_SynthInstance_newSubgoal___spec__12___closed__2; -x_10 = 0; -x_11 = l_Lean_getBoolOption___at_Lean_Meta_SynthInstance_newSubgoal___spec__13(x_9, x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_unbox(x_12); -lean_dec(x_12); -if (x_13 == 0) -{ -lean_object* x_14; lean_object* x_15; -x_14 = lean_ctor_get(x_11, 1); -lean_inc(x_14); -lean_dec(x_11); -x_15 = lean_apply_7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_14); -if (lean_obj_tag(x_15) == 0) -{ -uint8_t x_16; -x_16 = !lean_is_exclusive(x_15); -if (x_16 == 0) -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_17 = lean_ctor_get(x_15, 0); -x_18 = lean_box(0); -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_17); -lean_ctor_set(x_19, 1, x_18); -lean_ctor_set(x_15, 0, x_19); -return x_15; } -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_box(0); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_20); -lean_ctor_set(x_23, 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_21); -return x_24; } } -else -{ -uint8_t x_25; -x_25 = !lean_is_exclusive(x_15); -if (x_25 == 0) +LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Meta_SynthInstance_wakeUp___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) { +_start: { -return x_15; +lean_object* x_10; +x_10 = l_Lean_addTrace___at_Lean_Meta_SynthInstance_wakeUp___spec__1(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); +return x_10; } -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_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_EXPORT lean_object* l_Lean_Meta_SynthInstance_wakeUp___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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_Meta_SynthInstance_wakeUp(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_4); +lean_dec(x_3); +return x_10; } } -else -{ -lean_object* x_29; lean_object* x_30; -x_29 = lean_ctor_get(x_11, 1); -lean_inc(x_29); -lean_dec(x_11); -x_30 = l_Lean_withSeconds___at_Lean_Meta_SynthInstance_addAnswer___spec__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_29); -if (lean_obj_tag(x_30) == 0) +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_SynthInstance_isNewAnswer___spec__1(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4) { +_start: { -uint8_t x_31; -x_31 = !lean_is_exclusive(x_30); -if (x_31 == 0) +uint8_t x_5; +x_5 = lean_usize_dec_eq(x_3, x_4); +if (x_5 == 0) { -lean_object* x_32; uint8_t x_33; -x_32 = lean_ctor_get(x_30, 0); -x_33 = !lean_is_exclusive(x_32); -if (x_33 == 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 = lean_ctor_get(x_6, 1); +lean_inc(x_7); +lean_dec(x_6); +x_8 = lean_ctor_get(x_1, 1); +x_9 = lean_expr_eqv(x_7, x_8); +lean_dec(x_7); +if (x_9 == 0) { -lean_object* x_34; lean_object* x_35; -x_34 = lean_ctor_get(x_32, 1); -x_35 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_35, 0, x_34); -lean_ctor_set(x_32, 1, x_35); -return x_30; +size_t x_10; size_t x_11; +x_10 = 1; +x_11 = lean_usize_add(x_3, x_10); +x_3 = x_11; +goto _start; } 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 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_38, 0, x_37); -x_39 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_39, 0, x_36); -lean_ctor_set(x_39, 1, x_38); -lean_ctor_set(x_30, 0, x_39); -return x_30; +uint8_t x_13; +x_13 = 1; +return x_13; } } 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; -x_40 = lean_ctor_get(x_30, 0); -x_41 = lean_ctor_get(x_30, 1); -lean_inc(x_41); -lean_inc(x_40); -lean_dec(x_30); -x_42 = lean_ctor_get(x_40, 0); -lean_inc(x_42); -x_43 = lean_ctor_get(x_40, 1); -lean_inc(x_43); -if (lean_is_exclusive(x_40)) { - lean_ctor_release(x_40, 0); - lean_ctor_release(x_40, 1); - x_44 = x_40; -} else { - lean_dec_ref(x_40); - x_44 = lean_box(0); +uint8_t x_14; +x_14 = 0; +return x_14; } -x_45 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_45, 0, x_43); -if (lean_is_scalar(x_44)) { - x_46 = lean_alloc_ctor(0, 2, 0); -} else { - x_46 = x_44; } -lean_ctor_set(x_46, 0, x_42); -lean_ctor_set(x_46, 1, 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_41); -return x_47; } +LEAN_EXPORT uint8_t l_Lean_Meta_SynthInstance_isNewAnswer(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; uint8_t x_5; +x_3 = lean_array_get_size(x_1); +x_4 = lean_unsigned_to_nat(0u); +x_5 = lean_nat_dec_lt(x_4, x_3); +if (x_5 == 0) +{ +uint8_t x_6; +lean_dec(x_3); +x_6 = 1; +return x_6; } else { -uint8_t x_48; -x_48 = !lean_is_exclusive(x_30); -if (x_48 == 0) +uint8_t x_7; +x_7 = lean_nat_dec_le(x_3, x_3); +if (x_7 == 0) { -return x_30; +uint8_t x_8; +lean_dec(x_3); +x_8 = 1; +return x_8; } else { -lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_49 = lean_ctor_get(x_30, 0); -x_50 = lean_ctor_get(x_30, 1); -lean_inc(x_50); -lean_inc(x_49); -lean_dec(x_30); -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_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_addAnswer___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: -{ -if (lean_obj_tag(x_1) == 0) +size_t x_9; size_t x_10; uint8_t x_11; +x_9 = 0; +x_10 = lean_usize_of_nat(x_3); +lean_dec(x_3); +x_11 = l_Array_anyMUnsafe_any___at_Lean_Meta_SynthInstance_isNewAnswer___spec__1(x_2, x_1, x_9, x_10); +if (x_11 == 0) { -lean_object* x_9; lean_object* x_10; -x_9 = lean_ctor_get(x_1, 0); -lean_inc(x_9); -x_10 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_10, 0, x_9); -lean_ctor_set(x_10, 1, x_8); -return x_10; +uint8_t x_12; +x_12 = 1; +return x_12; } else { -lean_object* x_11; lean_object* x_12; -x_11 = lean_ctor_get(x_1, 0); -lean_inc(x_11); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_11); -lean_ctor_set(x_12, 1, x_8); -return x_12; +uint8_t x_13; +x_13 = 0; +return x_13; } } } -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_addAnswer___spec__3___lambda__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_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Meta_SynthInstance_isNewAnswer___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_15; lean_object* x_16; lean_object* x_17; -lean_inc(x_12); -x_15 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__15(x_1, x_2, x_3, x_6, x_4, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -x_16 = lean_ctor_get(x_15, 1); -lean_inc(x_16); -lean_dec(x_15); -x_17 = l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_addAnswer___spec__6(x_5, x_8, x_9, x_10, x_11, x_12, x_13, x_16); -lean_dec(x_12); -lean_dec(x_8); -return x_17; +size_t x_5; size_t x_6; uint8_t x_7; lean_object* x_8; +x_5 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_6 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_7 = l_Array_anyMUnsafe_any___at_Lean_Meta_SynthInstance_isNewAnswer___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_Lean_withTraceNode___at_Lean_Meta_SynthInstance_addAnswer___spec__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_EXPORT lean_object* l_Lean_Meta_SynthInstance_isNewAnswer___boxed(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_12; lean_object* x_13; uint8_t x_14; -lean_inc(x_1); -x_12 = l_Lean_isTracingEnabledFor___at_Lean_Meta_SynthInstance_newSubgoal___spec__10(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_unbox(x_13); -lean_dec(x_13); -if (x_14 == 0) -{ -lean_object* x_15; lean_object* x_16; +uint8_t x_3; lean_object* x_4; +x_3 = l_Lean_Meta_SynthInstance_isNewAnswer(x_1, x_2); lean_dec(x_2); lean_dec(x_1); -x_15 = lean_ctor_get(x_12, 1); -lean_inc(x_15); -lean_dec(x_12); -x_16 = lean_apply_7(x_3, x_5, x_6, x_7, x_8, x_9, x_10, x_15); -return x_16; +x_4 = lean_box(x_3); +return x_4; } -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_12, 1); -lean_inc(x_17); -lean_dec(x_12); -x_18 = lean_ctor_get(x_9, 5); -lean_inc(x_18); -x_19 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_SynthInstance_newSubgoal___spec__11___rarg(x_10, x_17); -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_alloc_closure((void*)(l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__1), 8, 1); -lean_closure_set(x_22, 0, x_3); -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_23 = l_Lean_withOptProfile___at_Lean_Meta_SynthInstance_addAnswer___spec__4(x_22, x_5, x_6, x_7, x_8, x_9, x_10, x_21); -if (lean_obj_tag(x_23) == 0) +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___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_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; -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_24, 0); -lean_inc(x_26); -x_27 = lean_ctor_get(x_24, 1); -lean_inc(x_27); -lean_dec(x_24); +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_9 = l_Lean_Meta_abstractMVars(x_1, x_4, x_5, x_6, x_7, x_8); +x_10 = lean_ctor_get(x_9, 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_26); -x_28 = lean_apply_8(x_2, x_26, x_5, x_6, x_7, x_8, x_9, x_10, x_25); -if (lean_obj_tag(x_28) == 0) +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +lean_dec(x_9); +x_12 = lean_ctor_get(x_10, 2); +lean_inc(x_12); +x_13 = lean_infer_type(x_12, x_4, x_5, x_6, x_7, x_11); +if (lean_obj_tag(x_13) == 0) { -if (lean_obj_tag(x_27) == 0) +uint8_t x_14; +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) { -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; -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_box(0); -x_32 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_addAnswer___spec__3___lambda__1(x_20, x_1, x_18, x_4, x_26, x_29, x_31, x_5, x_6, x_7, x_8, x_9, x_10, x_30); -lean_dec(x_10); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_26); -return x_32; +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 = lean_ctor_get(x_2, 4); +x_17 = lean_unsigned_to_nat(1u); +x_18 = lean_nat_add(x_16, x_17); +x_19 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_19, 0, x_10); +lean_ctor_set(x_19, 1, x_15); +lean_ctor_set(x_19, 2, x_18); +lean_ctor_set(x_13, 0, x_19); +return x_13; } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; double x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; 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_28, 0); -lean_inc(x_33); -x_34 = lean_ctor_get(x_28, 1); -lean_inc(x_34); -lean_dec(x_28); -x_35 = lean_ctor_get(x_27, 0); -lean_inc(x_35); -lean_dec(x_27); -x_36 = lean_unbox_float(x_35); -lean_dec(x_35); -x_37 = lean_float_to_string(x_36); -x_38 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_38, 0, x_37); -x_39 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_39, 0, x_38); -x_40 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___closed__2; -x_41 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_41, 0, x_40); -lean_ctor_set(x_41, 1, x_39); -x_42 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___closed__4; -x_43 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_43, 0, x_41); -lean_ctor_set(x_43, 1, x_42); -x_44 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_44, 0, x_43); -lean_ctor_set(x_44, 1, x_33); -x_45 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; -x_46 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -x_47 = lean_box(0); -x_48 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_addAnswer___spec__3___lambda__1(x_20, x_1, x_18, x_4, x_26, x_46, x_47, x_5, x_6, x_7, x_8, x_9, x_10, x_34); -lean_dec(x_10); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_26); -return x_48; +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_13, 0); +x_21 = lean_ctor_get(x_13, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_13); +x_22 = lean_ctor_get(x_2, 4); +x_23 = lean_unsigned_to_nat(1u); +x_24 = lean_nat_add(x_22, x_23); +x_25 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_25, 0, x_10); +lean_ctor_set(x_25, 1, x_20); +lean_ctor_set(x_25, 2, x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_21); +return x_26; } } else { -uint8_t x_49; -lean_dec(x_27); -lean_dec(x_26); -lean_dec(x_20); -lean_dec(x_18); +uint8_t x_27; 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_1); -x_49 = !lean_is_exclusive(x_28); -if (x_49 == 0) +x_27 = !lean_is_exclusive(x_13); +if (x_27 == 0) { -return x_28; +return x_13; } else { -lean_object* x_50; lean_object* x_51; lean_object* x_52; -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, 2, 0); -lean_ctor_set(x_52, 0, x_50); -lean_ctor_set(x_52, 1, x_51); -return x_52; +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; } } } -else +} +static lean_object* _init_l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___lambda__2___closed__1() { +_start: { -uint8_t x_53; -lean_dec(x_20); -lean_dec(x_18); -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_53 = !lean_is_exclusive(x_23); -if (x_53 == 0) -{ -return x_23; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("newAnswer", 9); +return x_1; } -else -{ -lean_object* x_54; lean_object* x_55; lean_object* x_56; -x_54 = lean_ctor_get(x_23, 0); -x_55 = lean_ctor_get(x_23, 1); -lean_inc(x_55); -lean_inc(x_54); -lean_dec(x_23); -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; } +static lean_object* _init_l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___lambda__2___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_6____closed__8; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_6____closed__1; +x_3 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___lambda__2___closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; } } +static lean_object* _init_l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___lambda__2___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("size: ", 6); +return x_1; } } -static lean_object* _init_l_Lean_Meta_SynthInstance_addAnswer___lambda__1___closed__1() { +static lean_object* _init_l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___lambda__2___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_checkEmoji; +x_1 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___lambda__2___closed__3; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_Meta_SynthInstance_addAnswer___lambda__1___closed__2() { +static lean_object* _init_l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___lambda__2___closed__5() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; -x_2 = l_Lean_Meta_SynthInstance_addAnswer___lambda__1___closed__1; -x_3 = lean_alloc_ctor(7, 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; +x_1 = lean_mk_string_from_bytes(", val: ", 7); +return x_1; } } -static lean_object* _init_l_Lean_Meta_SynthInstance_addAnswer___lambda__1___closed__3() { +static lean_object* _init_l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___lambda__2___closed__6() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_SynthInstance_addAnswer___lambda__1___closed__2; -x_2 = l_Lean_Meta_SynthInstance_tryResolve___lambda__1___closed__2; -x_3 = lean_alloc_ctor(7, 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 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___lambda__2___closed__5; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_addAnswer___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_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___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; -x_10 = lean_ctor_get(x_1, 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; uint8_t x_14; +x_8 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_1, x_3, x_4, x_5, x_6, 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_dec(x_1); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_11 = lean_infer_type(x_10, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_11) == 0) -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; -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 = l_Lean_instantiateMVars___at_Lean_Meta_SynthInstance_mkTableKeyFor___spec__1(x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_13); lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -x_15 = !lean_is_exclusive(x_14); -if (x_15 == 0) +x_11 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___lambda__2___closed__2; +x_12 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_11, x_3, x_4, x_5, x_6, x_10); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_unbox(x_13); +lean_dec(x_13); +if (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_14, 0); -x_17 = l_Lean_MessageData_ofExpr(x_16); -x_18 = l_Lean_Meta_SynthInstance_addAnswer___lambda__1___closed__3; -x_19 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_19, 0, x_18); -lean_ctor_set(x_19, 1, x_17); -x_20 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; -x_21 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_21, 0, x_19); -lean_ctor_set(x_21, 1, x_20); -lean_ctor_set(x_14, 0, x_21); -return x_14; +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_12, 1); +lean_inc(x_15); +lean_dec(x_12); +x_16 = lean_box(0); +x_17 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___lambda__1(x_9, x_2, x_16, x_3, x_4, x_5, x_6, x_15); +lean_dec(x_2); +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; lean_object* x_28; lean_object* x_29; -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 = l_Lean_MessageData_ofExpr(x_22); -x_25 = l_Lean_Meta_SynthInstance_addAnswer___lambda__1___closed__3; +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_18 = lean_ctor_get(x_12, 1); +lean_inc(x_18); +lean_dec(x_12); +x_19 = lean_ctor_get(x_2, 4); +lean_inc(x_19); +x_20 = l_Nat_repr(x_19); +x_21 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_21, 0, x_20); +x_22 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_22, 0, x_21); +x_23 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___lambda__2___closed__4; +x_24 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_22); +x_25 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___lambda__2___closed__6; x_26 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_24); -x_27 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +lean_inc(x_9); +x_27 = l_Lean_MessageData_ofExpr(x_9); x_28 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_28, 0, x_26); lean_ctor_set(x_28, 1, x_27); -x_29 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_29, 0, x_28); -lean_ctor_set(x_29, 1, x_23); -return x_29; +x_29 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; +x_30 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +x_31 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_11, x_30, x_3, x_4, x_5, x_6, x_18); +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_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___lambda__1(x_9, x_2, x_32, x_3, x_4, x_5, x_6, x_33); +lean_dec(x_32); +lean_dec(x_2); +return x_34; } } -else -{ -uint8_t x_30; -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -x_30 = !lean_is_exclusive(x_11); -if (x_30 == 0) -{ -return x_11; } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer(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_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_11, 0); -x_32 = lean_ctor_get(x_11, 1); -lean_inc(x_32); -lean_inc(x_31); -lean_dec(x_11); -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_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_7 = lean_ctor_get(x_1, 2); +lean_inc(x_7); +x_8 = lean_ctor_get(x_1, 0); +lean_inc(x_8); +x_9 = lean_alloc_closure((void*)(l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___lambda__2), 7, 2); +lean_closure_set(x_9, 0, x_8); +lean_closure_set(x_9, 1, x_1); +x_10 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_tryResolve___spec__1___rarg(x_7, x_9, x_2, x_3, x_4, x_5, x_6); +return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_addAnswer___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___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___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___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer(x_1, x_4, x_5, x_6, x_7, x_8); +x_9 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_3); +lean_dec(x_2); return x_9; } } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_addAnswer___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_ReaderT_read___at_Lean_Meta_SynthInstance_addAnswer___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) { _start: { -lean_object* x_10; lean_object* x_11; -x_10 = lean_ctor_get(x_1, 1); -lean_inc(x_10); -lean_dec(x_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); -lean_inc(x_10); -x_11 = l_Lean_Meta_SynthInstance_getEntry(x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -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; uint8_t x_14; -x_13 = lean_ctor_get(x_11, 0); -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) +lean_object* x_8; +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_1); +lean_ctor_set(x_8, 1, x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_addAnswer___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, lean_object* x_11, lean_object* x_12) { +_start: { -lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_15 = lean_ctor_get(x_11, 1); -x_16 = lean_ctor_get(x_13, 0); -x_17 = lean_ctor_get(x_13, 1); -x_18 = l_Lean_Meta_SynthInstance_isNewAnswer(x_17, x_2); -if (x_18 == 0) +uint8_t x_13; +x_13 = lean_usize_dec_eq(x_3, x_4); +if (x_13 == 0) { -lean_object* x_19; -lean_free_object(x_13); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_10); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; size_t x_18; size_t x_19; lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_19 = lean_box(0); -lean_ctor_set(x_11, 0, x_19); -return x_11; +x_14 = lean_array_uget(x_2, x_3); +lean_inc(x_8); +lean_inc(x_1); +x_15 = l_Lean_Meta_SynthInstance_wakeUp(x_1, x_14, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +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 = 1; +x_19 = lean_usize_add(x_3, x_18); +x_3 = x_19; +x_5 = x_16; +x_12 = x_17; +goto _start; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; -lean_free_object(x_11); -lean_inc(x_2); -x_20 = lean_array_push(x_17, x_2); -lean_inc(x_16); -lean_ctor_set(x_13, 1, x_20); -x_21 = lean_st_ref_take(x_4, x_15); -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_is_exclusive(x_22); -if (x_24 == 0) +lean_object* x_21; +lean_dec(x_8); +lean_dec(x_1); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_5); +lean_ctor_set(x_21, 1, x_12); +return x_21; +} +} +} +LEAN_EXPORT lean_object* l_Lean_withSeconds___at_Lean_Meta_SynthInstance_addAnswer___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_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; -x_25 = lean_ctor_get(x_22, 3); -x_26 = l_Lean_HashMap_insert___at_Lean_Meta_SynthInstance_newSubgoal___spec__1(x_25, x_10, x_13); -lean_ctor_set(x_22, 3, x_26); -x_27 = lean_st_ref_set(x_4, x_22, x_23); -x_28 = !lean_is_exclusive(x_27); -if (x_28 == 0) +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_9 = lean_io_mono_nanos_now(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_apply_7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_11); +if (lean_obj_tag(x_12) == 0) { -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; -x_29 = lean_ctor_get(x_27, 1); -x_30 = lean_ctor_get(x_27, 0); -lean_dec(x_30); -x_31 = lean_array_get_size(x_16); -x_32 = lean_unsigned_to_nat(0u); -x_33 = lean_nat_dec_lt(x_32, x_31); -if (x_33 == 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_io_mono_nanos_now(x_14); +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 0) { -lean_object* x_34; -lean_dec(x_31); -lean_dec(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_34 = lean_box(0); -lean_ctor_set(x_27, 0, x_34); -return x_27; +lean_object* x_17; lean_object* x_18; uint8_t x_19; lean_object* x_20; double x_21; double x_22; double x_23; lean_object* x_24; lean_object* x_25; +x_17 = lean_ctor_get(x_15, 0); +x_18 = lean_nat_sub(x_17, x_10); +lean_dec(x_10); +lean_dec(x_17); +x_19 = 0; +x_20 = lean_unsigned_to_nat(0u); +x_21 = l_Float_ofScientific(x_18, x_19, x_20); +lean_dec(x_18); +x_22 = l_Lean_withSeconds___at_Lean_Meta_SynthInstance_newSubgoal___spec__12___closed__1; +x_23 = lean_float_div(x_21, x_22); +x_24 = lean_box_float(x_23); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_13); +lean_ctor_set(x_25, 1, x_24); +lean_ctor_set(x_15, 0, x_25); +return x_15; } else { -uint8_t x_35; -x_35 = lean_nat_dec_le(x_31, x_31); -if (x_35 == 0) +lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; lean_object* x_30; double x_31; double x_32; double x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_26 = lean_ctor_get(x_15, 0); +x_27 = lean_ctor_get(x_15, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_15); +x_28 = lean_nat_sub(x_26, x_10); +lean_dec(x_10); +lean_dec(x_26); +x_29 = 0; +x_30 = lean_unsigned_to_nat(0u); +x_31 = l_Float_ofScientific(x_28, x_29, x_30); +lean_dec(x_28); +x_32 = l_Lean_withSeconds___at_Lean_Meta_SynthInstance_newSubgoal___spec__12___closed__1; +x_33 = lean_float_div(x_31, x_32); +x_34 = lean_box_float(x_33); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_13); +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_27); +return x_36; +} +} +else { -lean_object* x_36; -lean_dec(x_31); -lean_dec(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_36 = lean_box(0); -lean_ctor_set(x_27, 0, x_36); -return x_27; +uint8_t x_37; +lean_dec(x_10); +x_37 = !lean_is_exclusive(x_12); +if (x_37 == 0) +{ +return x_12; } else { -size_t x_37; size_t x_38; lean_object* x_39; lean_object* x_40; -lean_free_object(x_27); -x_37 = 0; -x_38 = lean_usize_of_nat(x_31); -lean_dec(x_31); -x_39 = lean_box(0); -x_40 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_addAnswer___spec__2(x_2, x_16, x_37, x_38, x_39, x_3, x_4, x_5, x_6, x_7, x_8, x_29); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_16); +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_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; } } } -else -{ -lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; -x_41 = lean_ctor_get(x_27, 1); -lean_inc(x_41); -lean_dec(x_27); -x_42 = lean_array_get_size(x_16); -x_43 = lean_unsigned_to_nat(0u); -x_44 = lean_nat_dec_lt(x_43, x_42); -if (x_44 == 0) -{ -lean_object* x_45; lean_object* x_46; -lean_dec(x_42); -lean_dec(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_45 = lean_box(0); -x_46 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_46, 0, x_45); -lean_ctor_set(x_46, 1, x_41); -return x_46; } -else +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_addAnswer___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) { +_start: { -uint8_t x_47; -x_47 = lean_nat_dec_le(x_42, x_42); -if (x_47 == 0) +if (lean_obj_tag(x_1) == 0) { -lean_object* x_48; lean_object* x_49; -lean_dec(x_42); -lean_dec(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_48 = lean_box(0); -x_49 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_49, 0, x_48); -lean_ctor_set(x_49, 1, x_41); -return x_49; +lean_object* x_9; lean_object* x_10; +x_9 = lean_ctor_get(x_1, 0); +lean_inc(x_9); +x_10 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_8); +return x_10; } else { -size_t x_50; size_t x_51; lean_object* x_52; lean_object* x_53; -x_50 = 0; -x_51 = lean_usize_of_nat(x_42); -lean_dec(x_42); -x_52 = lean_box(0); -x_53 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_addAnswer___spec__2(x_2, x_16, x_50, x_51, x_52, x_3, x_4, x_5, x_6, x_7, x_8, x_41); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_16); -return x_53; -} +lean_object* x_11; lean_object* x_12; +x_11 = lean_ctor_get(x_1, 0); +lean_inc(x_11); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_8); +return x_12; } } } -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; uint8_t x_65; -x_54 = lean_ctor_get(x_22, 0); -x_55 = lean_ctor_get(x_22, 1); -x_56 = lean_ctor_get(x_22, 2); -x_57 = lean_ctor_get(x_22, 3); -lean_inc(x_57); -lean_inc(x_56); -lean_inc(x_55); -lean_inc(x_54); -lean_dec(x_22); -x_58 = l_Lean_HashMap_insert___at_Lean_Meta_SynthInstance_newSubgoal___spec__1(x_57, x_10, x_13); -x_59 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_59, 0, x_54); -lean_ctor_set(x_59, 1, x_55); -lean_ctor_set(x_59, 2, x_56); -lean_ctor_set(x_59, 3, x_58); -x_60 = lean_st_ref_set(x_4, x_59, x_23); -x_61 = lean_ctor_get(x_60, 1); -lean_inc(x_61); -if (lean_is_exclusive(x_60)) { - lean_ctor_release(x_60, 0); - lean_ctor_release(x_60, 1); - x_62 = x_60; -} else { - lean_dec_ref(x_60); - x_62 = lean_box(0); -} -x_63 = lean_array_get_size(x_16); -x_64 = lean_unsigned_to_nat(0u); -x_65 = lean_nat_dec_lt(x_64, x_63); -if (x_65 == 0) +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_addAnswer___spec__3___lambda__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) { +_start: { -lean_object* x_66; lean_object* x_67; -lean_dec(x_63); -lean_dec(x_16); +lean_object* x_15; lean_object* x_16; lean_object* x_17; +lean_inc(x_12); +x_15 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__13(x_1, x_2, x_3, x_6, x_4, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +x_16 = lean_ctor_get(x_15, 1); +lean_inc(x_16); +lean_dec(x_15); +x_17 = l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_addAnswer___spec__5(x_5, x_8, x_9, x_10, x_11, x_12, x_13, x_16); +lean_dec(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); -x_66 = lean_box(0); -if (lean_is_scalar(x_62)) { - x_67 = lean_alloc_ctor(0, 2, 0); -} else { - x_67 = x_62; +return x_17; } -lean_ctor_set(x_67, 0, x_66); -lean_ctor_set(x_67, 1, x_61); -return x_67; } -else +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_addAnswer___spec__3___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, uint8_t x_7, double 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_68; -x_68 = lean_nat_dec_le(x_63, x_63); -if (x_68 == 0) +lean_object* x_17; lean_object* x_18; +lean_dec(x_9); +x_17 = lean_ctor_get(x_14, 5); +lean_inc(x_17); +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_2); +x_18 = lean_apply_8(x_1, x_2, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_18) == 0) { -lean_object* x_69; lean_object* x_70; -lean_dec(x_63); -lean_dec(x_16); -lean_dec(x_8); -lean_dec(x_7); +lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +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_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__1; +x_22 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_6, x_21); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); +if (x_22 == 0) +{ +if (x_7 == 0) +{ +lean_object* x_23; lean_object* x_24; +x_23 = lean_box(0); +x_24 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_addAnswer___spec__3___lambda__1(x_3, x_4, x_17, x_5, x_2, x_19, x_23, x_10, x_11, x_12, x_13, x_14, x_15, x_20); +lean_dec(x_15); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_2); -x_69 = lean_box(0); -if (lean_is_scalar(x_62)) { - x_70 = lean_alloc_ctor(0, 2, 0); -} else { - x_70 = x_62; -} -lean_ctor_set(x_70, 0, x_69); -lean_ctor_set(x_70, 1, x_61); -return x_70; +return x_24; } else { -size_t x_71; size_t x_72; lean_object* x_73; lean_object* x_74; -lean_dec(x_62); -x_71 = 0; -x_72 = lean_usize_of_nat(x_63); -lean_dec(x_63); -x_73 = lean_box(0); -x_74 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_addAnswer___spec__2(x_2, x_16, x_71, x_72, x_73, x_3, x_4, x_5, x_6, x_7, x_8, x_61); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_16); -return x_74; -} +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; 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_25 = lean_float_to_string(x_8); +x_26 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_26, 0, x_25); +x_27 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_27, 0, x_26); +x_28 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__3; +x_29 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_27); +x_30 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__5; +x_31 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_30); +x_32 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_19); +x_33 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; +x_34 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +x_35 = lean_box(0); +x_36 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_addAnswer___spec__3___lambda__1(x_3, x_4, x_17, x_5, x_2, x_34, x_35, x_10, x_11, x_12, x_13, x_14, x_15, x_20); +lean_dec(x_15); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_2); +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; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_37 = lean_float_to_string(x_8); +x_38 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_38, 0, x_37); +x_39 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_39, 0, x_38); +x_40 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__3; +x_41 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_39); +x_42 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__5; +x_43 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_43, 0, x_41); +lean_ctor_set(x_43, 1, x_42); +x_44 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_44, 0, x_43); +lean_ctor_set(x_44, 1, x_19); +x_45 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; +x_46 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +x_47 = lean_box(0); +x_48 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_addAnswer___spec__3___lambda__1(x_3, x_4, x_17, x_5, x_2, x_46, x_47, x_10, x_11, x_12, x_13, x_14, x_15, x_20); +lean_dec(x_15); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_2); +return x_48; } } else { -lean_object* x_75; lean_object* x_76; lean_object* x_77; uint8_t x_78; -x_75 = lean_ctor_get(x_11, 1); -x_76 = lean_ctor_get(x_13, 0); -x_77 = lean_ctor_get(x_13, 1); -lean_inc(x_77); -lean_inc(x_76); +uint8_t x_49; +lean_dec(x_17); +lean_dec(x_15); +lean_dec(x_14); lean_dec(x_13); -x_78 = l_Lean_Meta_SynthInstance_isNewAnswer(x_77, x_2); -if (x_78 == 0) -{ -lean_object* x_79; -lean_dec(x_77); -lean_dec(x_76); +lean_dec(x_12); +lean_dec(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); lean_dec(x_3); lean_dec(x_2); -x_79 = lean_box(0); -lean_ctor_set(x_11, 0, x_79); -return x_11; +x_49 = !lean_is_exclusive(x_18); +if (x_49 == 0) +{ +return x_18; } 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; uint8_t x_97; -lean_free_object(x_11); -lean_inc(x_2); -x_80 = lean_array_push(x_77, x_2); -lean_inc(x_76); -x_81 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_81, 0, x_76); -lean_ctor_set(x_81, 1, x_80); -x_82 = lean_st_ref_take(x_4, x_75); -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_ctor_get(x_83, 0); -lean_inc(x_85); -x_86 = lean_ctor_get(x_83, 1); -lean_inc(x_86); -x_87 = lean_ctor_get(x_83, 2); -lean_inc(x_87); -x_88 = lean_ctor_get(x_83, 3); -lean_inc(x_88); -if (lean_is_exclusive(x_83)) { - lean_ctor_release(x_83, 0); - lean_ctor_release(x_83, 1); - lean_ctor_release(x_83, 2); - lean_ctor_release(x_83, 3); - x_89 = x_83; -} else { - lean_dec_ref(x_83); - x_89 = lean_box(0); +lean_object* x_50; lean_object* x_51; lean_object* x_52; +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_alloc_ctor(1, 2, 0); +lean_ctor_set(x_52, 0, x_50); +lean_ctor_set(x_52, 1, x_51); +return x_52; } -x_90 = l_Lean_HashMap_insert___at_Lean_Meta_SynthInstance_newSubgoal___spec__1(x_88, x_10, x_81); -if (lean_is_scalar(x_89)) { - x_91 = lean_alloc_ctor(0, 4, 0); -} else { - x_91 = x_89; } -lean_ctor_set(x_91, 0, x_85); -lean_ctor_set(x_91, 1, x_86); -lean_ctor_set(x_91, 2, x_87); -lean_ctor_set(x_91, 3, x_90); -x_92 = lean_st_ref_set(x_4, x_91, x_84); -x_93 = lean_ctor_get(x_92, 1); -lean_inc(x_93); -if (lean_is_exclusive(x_92)) { - lean_ctor_release(x_92, 0); - lean_ctor_release(x_92, 1); - x_94 = x_92; -} else { - lean_dec_ref(x_92); - x_94 = lean_box(0); } -x_95 = lean_array_get_size(x_76); -x_96 = lean_unsigned_to_nat(0u); -x_97 = lean_nat_dec_lt(x_96, x_95); -if (x_97 == 0) +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_addAnswer___spec__3___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t 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, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: { -lean_object* x_98; lean_object* x_99; -lean_dec(x_95); -lean_dec(x_76); -lean_dec(x_8); +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_dec(x_7); -lean_dec(x_6); +x_15 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_SynthInstance_newSubgoal___spec__11___rarg(x_13, x_14); +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_closure((void*)(l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__1), 8, 1); +lean_closure_set(x_18, 0, x_1); +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_19 = l_Lean_withSeconds___at_Lean_Meta_SynthInstance_addAnswer___spec__4(x_18, x_8, x_9, x_10, x_11, x_12, x_13, x_17); +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_69; uint8_t x_70; +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_20, 0); +lean_inc(x_22); +x_23 = lean_ctor_get(x_20, 1); +lean_inc(x_23); +lean_dec(x_20); +x_69 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__4___closed__1; +x_70 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_5, x_69); +if (x_70 == 0) +{ +uint8_t x_71; +x_71 = 0; +x_24 = x_71; +goto block_68; +} +else +{ +double x_72; double x_73; uint8_t x_74; +x_72 = l_Lean_trace_profiler_threshold_getSecs(x_5); +x_73 = lean_unbox_float(x_23); +x_74 = lean_float_decLt(x_72, x_73); +x_24 = x_74; +goto block_68; +} +block_68: +{ +if (x_6 == 0) +{ +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; +lean_dec(x_23); lean_dec(x_5); -lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_98 = lean_box(0); -if (lean_is_scalar(x_94)) { - x_99 = lean_alloc_ctor(0, 2, 0); -} else { - x_99 = x_94; +x_25 = lean_st_ref_take(x_13, x_21); +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_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; +x_29 = lean_ctor_get(x_26, 3); +x_30 = l_Lean_PersistentArray_append___rarg(x_16, x_29); +lean_ctor_set(x_26, 3, x_30); +x_31 = lean_st_ref_set(x_13, x_26, x_27); +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +lean_dec(x_31); +x_33 = l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_addAnswer___spec__5(x_22, x_8, x_9, x_10, x_11, x_12, x_13, x_32); +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_22); +if (lean_obj_tag(x_33) == 0) +{ +uint8_t x_34; +x_34 = !lean_is_exclusive(x_33); +if (x_34 == 0) +{ +return x_33; +} +else +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_33, 0); +x_36 = lean_ctor_get(x_33, 1); +lean_inc(x_36); +lean_inc(x_35); +lean_dec(x_33); +x_37 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_37, 0, x_35); +lean_ctor_set(x_37, 1, x_36); +return x_37; } -lean_ctor_set(x_99, 0, x_98); -lean_ctor_set(x_99, 1, x_93); -return x_99; } else { -uint8_t x_100; -x_100 = lean_nat_dec_le(x_95, x_95); -if (x_100 == 0) +uint8_t x_38; +x_38 = !lean_is_exclusive(x_33); +if (x_38 == 0) { -lean_object* x_101; lean_object* x_102; -lean_dec(x_95); -lean_dec(x_76); +return x_33; +} +else +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_33, 0); +x_40 = lean_ctor_get(x_33, 1); +lean_inc(x_40); +lean_inc(x_39); +lean_dec(x_33); +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 +{ +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_42 = lean_ctor_get(x_26, 0); +x_43 = lean_ctor_get(x_26, 1); +x_44 = lean_ctor_get(x_26, 2); +x_45 = lean_ctor_get(x_26, 3); +x_46 = lean_ctor_get(x_26, 4); +x_47 = lean_ctor_get(x_26, 5); +x_48 = lean_ctor_get(x_26, 6); +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_inc(x_42); +lean_dec(x_26); +x_49 = l_Lean_PersistentArray_append___rarg(x_16, x_45); +x_50 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_50, 0, x_42); +lean_ctor_set(x_50, 1, x_43); +lean_ctor_set(x_50, 2, x_44); +lean_ctor_set(x_50, 3, x_49); +lean_ctor_set(x_50, 4, x_46); +lean_ctor_set(x_50, 5, x_47); +lean_ctor_set(x_50, 6, x_48); +x_51 = lean_st_ref_set(x_13, x_50, x_27); +x_52 = lean_ctor_get(x_51, 1); +lean_inc(x_52); +lean_dec(x_51); +x_53 = l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_addAnswer___spec__5(x_22, x_8, x_9, x_10, x_11, x_12, x_13, x_52); +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); -lean_dec(x_3); -lean_dec(x_2); -x_101 = lean_box(0); -if (lean_is_scalar(x_94)) { - x_102 = lean_alloc_ctor(0, 2, 0); +lean_dec(x_22); +if (lean_obj_tag(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_53, 0); +lean_inc(x_54); +x_55 = lean_ctor_get(x_53, 1); +lean_inc(x_55); +if (lean_is_exclusive(x_53)) { + lean_ctor_release(x_53, 0); + lean_ctor_release(x_53, 1); + x_56 = x_53; } else { - x_102 = x_94; + lean_dec_ref(x_53); + x_56 = lean_box(0); } -lean_ctor_set(x_102, 0, x_101); -lean_ctor_set(x_102, 1, x_93); -return x_102; +if (lean_is_scalar(x_56)) { + x_57 = lean_alloc_ctor(0, 2, 0); +} else { + x_57 = x_56; +} +lean_ctor_set(x_57, 0, x_54); +lean_ctor_set(x_57, 1, x_55); +return x_57; } else { -size_t x_103; size_t x_104; lean_object* x_105; lean_object* x_106; -lean_dec(x_94); -x_103 = 0; -x_104 = lean_usize_of_nat(x_95); -lean_dec(x_95); -x_105 = lean_box(0); -x_106 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_addAnswer___spec__2(x_2, x_76, x_103, x_104, x_105, x_3, x_4, x_5, x_6, x_7, x_8, x_93); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_76); -return x_106; +lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_58 = lean_ctor_get(x_53, 0); +lean_inc(x_58); +x_59 = lean_ctor_get(x_53, 1); +lean_inc(x_59); +if (lean_is_exclusive(x_53)) { + lean_ctor_release(x_53, 0); + lean_ctor_release(x_53, 1); + x_60 = x_53; +} else { + lean_dec_ref(x_53); + x_60 = lean_box(0); } +if (lean_is_scalar(x_60)) { + x_61 = lean_alloc_ctor(1, 2, 0); +} else { + x_61 = x_60; } +lean_ctor_set(x_61, 0, x_58); +lean_ctor_set(x_61, 1, x_59); +return x_61; } } } else { -lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; uint8_t x_112; -x_107 = lean_ctor_get(x_11, 0); -x_108 = lean_ctor_get(x_11, 1); -lean_inc(x_108); -lean_inc(x_107); -lean_dec(x_11); -x_109 = lean_ctor_get(x_107, 0); -lean_inc(x_109); -x_110 = lean_ctor_get(x_107, 1); -lean_inc(x_110); -if (lean_is_exclusive(x_107)) { - lean_ctor_release(x_107, 0); - lean_ctor_release(x_107, 1); - x_111 = x_107; -} else { - lean_dec_ref(x_107); - x_111 = lean_box(0); +lean_object* x_62; double x_63; lean_object* x_64; +x_62 = lean_box(0); +x_63 = lean_unbox_float(x_23); +lean_dec(x_23); +x_64 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_addAnswer___spec__3___lambda__2(x_2, x_22, x_16, x_3, x_4, x_5, x_24, x_63, x_62, x_8, x_9, x_10, x_11, x_12, x_13, x_21); +return x_64; } -x_112 = l_Lean_Meta_SynthInstance_isNewAnswer(x_110, x_2); -if (x_112 == 0) +} +else { -lean_object* x_113; lean_object* x_114; -lean_dec(x_111); -lean_dec(x_110); -lean_dec(x_109); +lean_object* x_65; double x_66; lean_object* x_67; +x_65 = lean_box(0); +x_66 = lean_unbox_float(x_23); +lean_dec(x_23); +x_67 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_addAnswer___spec__3___lambda__2(x_2, x_22, x_16, x_3, x_4, x_5, x_24, x_66, x_65, x_8, x_9, x_10, x_11, x_12, x_13, x_21); +return x_67; +} +} +} +else +{ +uint8_t x_75; +lean_dec(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_4); lean_dec(x_3); lean_dec(x_2); -x_113 = lean_box(0); -x_114 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_114, 0, x_113); -lean_ctor_set(x_114, 1, x_108); -return x_114; +x_75 = !lean_is_exclusive(x_19); +if (x_75 == 0) +{ +return x_19; } else { -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; uint8_t x_132; -lean_inc(x_2); -x_115 = lean_array_push(x_110, x_2); -lean_inc(x_109); -if (lean_is_scalar(x_111)) { - x_116 = lean_alloc_ctor(0, 2, 0); -} else { - x_116 = x_111; +lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_76 = lean_ctor_get(x_19, 0); +x_77 = lean_ctor_get(x_19, 1); +lean_inc(x_77); +lean_inc(x_76); +lean_dec(x_19); +x_78 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_78, 0, x_76); +lean_ctor_set(x_78, 1, x_77); +return x_78; } -lean_ctor_set(x_116, 0, x_109); -lean_ctor_set(x_116, 1, x_115); -x_117 = lean_st_ref_take(x_4, x_108); -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); -x_120 = lean_ctor_get(x_118, 0); -lean_inc(x_120); -x_121 = lean_ctor_get(x_118, 1); -lean_inc(x_121); -x_122 = lean_ctor_get(x_118, 2); -lean_inc(x_122); -x_123 = lean_ctor_get(x_118, 3); -lean_inc(x_123); -if (lean_is_exclusive(x_118)) { - lean_ctor_release(x_118, 0); - lean_ctor_release(x_118, 1); - lean_ctor_release(x_118, 2); - lean_ctor_release(x_118, 3); - x_124 = x_118; -} else { - lean_dec_ref(x_118); - x_124 = lean_box(0); } -x_125 = l_Lean_HashMap_insert___at_Lean_Meta_SynthInstance_newSubgoal___spec__1(x_123, x_10, x_116); -if (lean_is_scalar(x_124)) { - x_126 = lean_alloc_ctor(0, 4, 0); -} else { - x_126 = x_124; } -lean_ctor_set(x_126, 0, x_120); -lean_ctor_set(x_126, 1, x_121); -lean_ctor_set(x_126, 2, x_122); -lean_ctor_set(x_126, 3, x_125); -x_127 = lean_st_ref_set(x_4, x_126, x_119); -x_128 = lean_ctor_get(x_127, 1); -lean_inc(x_128); -if (lean_is_exclusive(x_127)) { - lean_ctor_release(x_127, 0); - lean_ctor_release(x_127, 1); - x_129 = x_127; -} else { - lean_dec_ref(x_127); - x_129 = lean_box(0); } -x_130 = lean_array_get_size(x_109); -x_131 = lean_unsigned_to_nat(0u); -x_132 = lean_nat_dec_lt(x_131, x_130); -if (x_132 == 0) +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_addAnswer___spec__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) { +_start: { -lean_object* x_133; lean_object* x_134; -lean_dec(x_130); -lean_dec(x_109); -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_133 = lean_box(0); -if (lean_is_scalar(x_129)) { - x_134 = lean_alloc_ctor(0, 2, 0); -} else { - x_134 = x_129; -} -lean_ctor_set(x_134, 0, x_133); -lean_ctor_set(x_134, 1, x_128); -return x_134; -} -else +lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_12 = lean_ctor_get(x_9, 2); +lean_inc(x_12); +lean_inc(x_1); +x_13 = l_Lean_isTracingEnabledFor___at_Lean_Meta_SynthInstance_newSubgoal___spec__10(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_unbox(x_14); +if (x_15 == 0) { -uint8_t x_135; -x_135 = lean_nat_dec_le(x_130, x_130); -if (x_135 == 0) +lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); +lean_dec(x_13); +x_17 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__4___closed__1; +x_18 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_12, x_17); +if (x_18 == 0) { -lean_object* x_136; lean_object* x_137; -lean_dec(x_130); -lean_dec(x_109); -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_object* x_19; +lean_dec(x_14); +lean_dec(x_12); lean_dec(x_2); -x_136 = lean_box(0); -if (lean_is_scalar(x_129)) { - x_137 = lean_alloc_ctor(0, 2, 0); -} else { - x_137 = x_129; -} -lean_ctor_set(x_137, 0, x_136); -lean_ctor_set(x_137, 1, x_128); -return x_137; +lean_dec(x_1); +x_19 = lean_apply_7(x_3, x_5, x_6, x_7, x_8, x_9, x_10, x_16); +if (lean_obj_tag(x_19) == 0) +{ +uint8_t x_20; +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) +{ +return x_19; } else { -size_t x_138; size_t x_139; lean_object* x_140; lean_object* x_141; -lean_dec(x_129); -x_138 = 0; -x_139 = lean_usize_of_nat(x_130); -lean_dec(x_130); -x_140 = lean_box(0); -x_141 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_addAnswer___spec__2(x_2, x_109, x_138, x_139, x_140, x_3, x_4, x_5, x_6, x_7, x_8, x_128); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_109); -return x_141; -} -} -} +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(0, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +return x_23; } } else { -uint8_t x_142; -lean_dec(x_10); -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_142 = !lean_is_exclusive(x_11); -if (x_142 == 0) +uint8_t x_24; +x_24 = !lean_is_exclusive(x_19); +if (x_24 == 0) { -return x_11; +return x_19; } else { -lean_object* x_143; lean_object* x_144; lean_object* x_145; -x_143 = lean_ctor_get(x_11, 0); -x_144 = lean_ctor_get(x_11, 1); -lean_inc(x_144); -lean_inc(x_143); -lean_dec(x_11); -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_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_19, 0); +x_26 = lean_ctor_get(x_19, 1); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_19); +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; } } } -static lean_object* _init_l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("answer", 6); -return x_1; +lean_object* x_28; uint8_t x_29; lean_object* x_30; +x_28 = lean_box(0); +x_29 = lean_unbox(x_14); +lean_dec(x_14); +x_30 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_addAnswer___spec__3___lambda__3(x_3, x_2, x_1, x_4, x_12, x_29, x_28, x_5, x_6, x_7, x_8, x_9, x_10, x_16); +return x_30; } } -static lean_object* _init_l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__2() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_6____closed__8; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_6____closed__1; -x_3 = l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__1; -x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); -return x_4; +lean_object* x_31; lean_object* x_32; uint8_t x_33; lean_object* x_34; +x_31 = lean_ctor_get(x_13, 1); +lean_inc(x_31); +lean_dec(x_13); +x_32 = lean_box(0); +x_33 = lean_unbox(x_14); +lean_dec(x_14); +x_34 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_addAnswer___spec__3___lambda__3(x_3, x_2, x_1, x_4, x_12, x_33, x_32, x_5, x_6, x_7, x_8, x_9, x_10, x_31); +return x_34; } } -static lean_object* _init_l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__3() { +} +static lean_object* _init_l_Lean_Meta_SynthInstance_addAnswer___lambda__1___closed__1() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_crossEmoji; +x_1 = l_Lean_checkEmoji; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__4() { +static lean_object* _init_l_Lean_Meta_SynthInstance_addAnswer___lambda__1___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; -x_2 = l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__3; +x_2 = l_Lean_Meta_SynthInstance_addAnswer___lambda__1___closed__1; x_3 = lean_alloc_ctor(7, 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_Meta_SynthInstance_addAnswer___lambda__4___closed__5() { +static lean_object* _init_l_Lean_Meta_SynthInstance_addAnswer___lambda__1___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__4; +x_1 = l_Lean_Meta_SynthInstance_addAnswer___lambda__1___closed__2; x_2 = l_Lean_Meta_SynthInstance_tryResolve___lambda__1___closed__2; x_3 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11581,1098 +11147,1129 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__6() { +LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_addAnswer___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_1; lean_object* x_2; -x_1 = lean_box(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_Meta_SynthInstance_addAnswer___lambda__4___closed__7() { -_start: +lean_object* x_10; lean_object* x_11; +x_10 = lean_ctor_get(x_1, 0); +lean_inc(x_10); +lean_dec(x_1); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_11 = lean_infer_type(x_10, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_11) == 0) { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("(size: ", 7); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__8() { -_start: +lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +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 = l_Lean_instantiateMVars___at_Lean_Meta_SynthInstance_mkTableKeyFor___spec__1(x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_13); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__7; -x_2 = l_Lean_stringToMessageData(x_1); -return x_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; +x_16 = lean_ctor_get(x_14, 0); +x_17 = l_Lean_MessageData_ofExpr(x_16); +x_18 = l_Lean_Meta_SynthInstance_addAnswer___lambda__1___closed__3; +x_19 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_17); +x_20 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; +x_21 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); +lean_ctor_set(x_14, 0, x_21); +return x_14; } -static lean_object* _init_l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__9() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes(" ≥ ", 5); -return x_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; +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 = l_Lean_MessageData_ofExpr(x_22); +x_25 = l_Lean_Meta_SynthInstance_addAnswer___lambda__1___closed__3; +x_26 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_24); +x_27 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; +x_28 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 1, x_27); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_23); +return x_29; } } -static lean_object* _init_l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__10() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__9; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} +uint8_t x_30; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +x_30 = !lean_is_exclusive(x_11); +if (x_30 == 0) +{ +return x_11; } -static lean_object* _init_l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__11() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes(")", 1); -return x_1; +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_11, 0); +x_32 = lean_ctor_get(x_11, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_11); +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_Meta_SynthInstance_addAnswer___lambda__4___closed__12() { +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_addAnswer___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_1; lean_object* x_2; -x_1 = l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__11; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +lean_object* x_9; +x_9 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer(x_1, x_4, x_5, x_6, x_7, x_8); +return x_9; } } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_addAnswer___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_EXPORT lean_object* l_Lean_Meta_SynthInstance_addAnswer___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; uint8_t x_12; -x_10 = lean_ctor_get(x_2, 0); +lean_object* x_10; lean_object* x_11; +x_10 = lean_ctor_get(x_1, 1); lean_inc(x_10); -lean_dec(x_2); -x_11 = lean_ctor_get(x_1, 4); -lean_inc(x_11); -x_12 = lean_nat_dec_le(x_10, x_11); -lean_dec(x_10); +lean_dec(x_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); +lean_inc(x_10); +x_11 = l_Lean_Meta_SynthInstance_getEntry(x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +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; uint8_t x_18; lean_object* x_19; -lean_dec(x_11); -lean_inc(x_1); -x_13 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_addAnswer___lambda__1___boxed), 9, 1); -lean_closure_set(x_13, 0, x_1); -lean_inc(x_1); -x_14 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_addAnswer___lambda__2___boxed), 8, 1); -lean_closure_set(x_14, 0, x_1); -x_15 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_addAnswer___lambda__3), 9, 1); -lean_closure_set(x_15, 0, x_1); -x_16 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_SynthInstance_newSubgoal___spec__7___rarg), 9, 2); -lean_closure_set(x_16, 0, x_14); -lean_closure_set(x_16, 1, x_15); -x_17 = l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__2; -x_18 = 1; -x_19 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_addAnswer___spec__3(x_17, x_13, x_16, x_18, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_19; +lean_object* x_13; uint8_t x_14; +x_13 = lean_ctor_get(x_11, 0); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_15 = lean_ctor_get(x_11, 1); +x_16 = lean_ctor_get(x_13, 0); +x_17 = lean_ctor_get(x_13, 1); +x_18 = l_Lean_Meta_SynthInstance_isNewAnswer(x_17, x_2); +if (x_18 == 0) +{ +lean_object* x_19; +lean_free_object(x_13); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_10); +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_19 = lean_box(0); +lean_ctor_set(x_11, 0, x_19); +return x_11; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; -x_20 = l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__2; -x_21 = l_Lean_isTracingEnabledFor___at_Lean_Meta_SynthInstance_newSubgoal___spec__10(x_20, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +lean_free_object(x_11); +lean_inc(x_2); +x_20 = lean_array_push(x_17, x_2); +lean_inc(x_16); +lean_ctor_set(x_13, 1, x_20); +x_21 = lean_st_ref_take(x_4, x_15); x_22 = lean_ctor_get(x_21, 0); lean_inc(x_22); -x_23 = lean_unbox(x_22); -lean_dec(x_22); -if (x_23 == 0) +x_23 = lean_ctor_get(x_21, 1); +lean_inc(x_23); +lean_dec(x_21); +x_24 = !lean_is_exclusive(x_22); +if (x_24 == 0) { -uint8_t x_24; -lean_dec(x_11); +lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_25 = lean_ctor_get(x_22, 3); +x_26 = l_Lean_HashMap_insert___at_Lean_Meta_SynthInstance_newSubgoal___spec__1(x_25, x_10, x_13); +lean_ctor_set(x_22, 3, x_26); +x_27 = lean_st_ref_set(x_4, x_22, x_23); +x_28 = !lean_is_exclusive(x_27); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; +x_29 = lean_ctor_get(x_27, 1); +x_30 = lean_ctor_get(x_27, 0); +lean_dec(x_30); +x_31 = lean_array_get_size(x_16); +x_32 = lean_unsigned_to_nat(0u); +x_33 = lean_nat_dec_lt(x_32, x_31); +if (x_33 == 0) +{ +lean_object* x_34; +lean_dec(x_31); +lean_dec(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_1); -x_24 = !lean_is_exclusive(x_21); -if (x_24 == 0) +lean_dec(x_2); +x_34 = lean_box(0); +lean_ctor_set(x_27, 0, x_34); +return x_27; +} +else { -lean_object* x_25; lean_object* x_26; -x_25 = lean_ctor_get(x_21, 0); -lean_dec(x_25); -x_26 = lean_box(0); -lean_ctor_set(x_21, 0, x_26); -return x_21; +uint8_t x_35; +x_35 = lean_nat_dec_le(x_31, x_31); +if (x_35 == 0) +{ +lean_object* x_36; +lean_dec(x_31); +lean_dec(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_36 = lean_box(0); +lean_ctor_set(x_27, 0, x_36); +return x_27; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_21, 1); -lean_inc(x_27); -lean_dec(x_21); -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; +size_t x_37; size_t x_38; lean_object* x_39; lean_object* x_40; +lean_free_object(x_27); +x_37 = 0; +x_38 = lean_usize_of_nat(x_31); +lean_dec(x_31); +x_39 = lean_box(0); +x_40 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_addAnswer___spec__2(x_2, x_16, x_37, x_38, x_39, x_3, x_4, x_5, x_6, x_7, x_8, x_29); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_16); +return x_40; +} } } else { -lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_30 = lean_ctor_get(x_21, 1); -lean_inc(x_30); -lean_dec(x_21); -x_31 = lean_ctor_get(x_1, 0); -lean_inc(x_31); -lean_dec(x_1); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_32 = lean_infer_type(x_31, x_5, x_6, x_7, x_8, x_30); -if (lean_obj_tag(x_32) == 0) +lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; +x_41 = lean_ctor_get(x_27, 1); +lean_inc(x_41); +lean_dec(x_27); +x_42 = lean_array_get_size(x_16); +x_43 = lean_unsigned_to_nat(0u); +x_44 = lean_nat_dec_lt(x_43, x_42); +if (x_44 == 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; -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_instantiateMVars___at_Lean_Meta_SynthInstance_mkTableKeyFor___spec__1(x_33, x_3, x_4, x_5, x_6, x_7, x_8, 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); -lean_dec(x_35); -x_38 = l_Lean_MessageData_ofExpr(x_36); -x_39 = l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__5; -x_40 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_40, 0, x_39); -lean_ctor_set(x_40, 1, x_38); -x_41 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; -x_42 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_42, 0, x_40); -lean_ctor_set(x_42, 1, x_41); -x_43 = l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__6; -x_44 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_44, 0, x_42); -lean_ctor_set(x_44, 1, x_43); -x_45 = l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__8; -x_46 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -x_47 = l_Nat_repr(x_11); -x_48 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_48, 0, x_47); -x_49 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_49, 0, x_48); -x_50 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_50, 0, x_46); -lean_ctor_set(x_50, 1, x_49); -x_51 = l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__10; -x_52 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_52, 0, x_50); -lean_ctor_set(x_52, 1, x_51); -x_53 = lean_ctor_get(x_3, 0); -lean_inc(x_53); -x_54 = l_Nat_repr(x_53); -x_55 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_55, 0, x_54); -x_56 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_56, 0, x_55); -x_57 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_57, 0, x_52); -lean_ctor_set(x_57, 1, x_56); -x_58 = l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__12; -x_59 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_59, 0, x_57); -lean_ctor_set(x_59, 1, x_58); -x_60 = l_Lean_addTrace___at_Lean_Meta_SynthInstance_wakeUp___spec__1(x_20, x_59, x_3, x_4, x_5, x_6, x_7, x_8, x_37); +lean_object* x_45; lean_object* x_46; +lean_dec(x_42); +lean_dec(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); -return x_60; +lean_dec(x_2); +x_45 = lean_box(0); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_41); +return x_46; } else { -uint8_t x_61; -lean_dec(x_11); +uint8_t x_47; +x_47 = lean_nat_dec_le(x_42, x_42); +if (x_47 == 0) +{ +lean_object* x_48; lean_object* x_49; +lean_dec(x_42); +lean_dec(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); -x_61 = !lean_is_exclusive(x_32); -if (x_61 == 0) -{ -return x_32; +lean_dec(x_2); +x_48 = lean_box(0); +x_49 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_49, 0, x_48); +lean_ctor_set(x_49, 1, x_41); +return x_49; } else { -lean_object* x_62; lean_object* x_63; lean_object* x_64; -x_62 = lean_ctor_get(x_32, 0); -x_63 = lean_ctor_get(x_32, 1); -lean_inc(x_63); -lean_inc(x_62); -lean_dec(x_32); -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; -} -} -} -} +size_t x_50; size_t x_51; lean_object* x_52; lean_object* x_53; +x_50 = 0; +x_51 = lean_usize_of_nat(x_42); +lean_dec(x_42); +x_52 = lean_box(0); +x_53 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_addAnswer___spec__2(x_2, x_16, x_50, x_51, x_52, x_3, x_4, x_5, x_6, x_7, x_8, x_41); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_16); +return x_53; } } -static lean_object* _init_l_Lean_Meta_SynthInstance_addAnswer___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_ReaderT_read___at_Lean_Meta_SynthInstance_addAnswer___spec__1___boxed), 7, 0); -return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_addAnswer(lean_object* x_1, lean_object* x_2, 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; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_9 = lean_ctor_get(x_1, 2); -lean_inc(x_9); -x_10 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_addAnswer___lambda__4), 9, 1); -lean_closure_set(x_10, 0, x_1); -x_11 = l_Lean_Meta_SynthInstance_addAnswer___closed__1; -x_12 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_SynthInstance_newSubgoal___spec__7___rarg), 9, 2); -lean_closure_set(x_12, 0, x_11); -lean_closure_set(x_12, 1, x_10); -x_13 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__18___rarg(x_9, x_12, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -return x_13; -} +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; +x_54 = lean_ctor_get(x_22, 0); +x_55 = lean_ctor_get(x_22, 1); +x_56 = lean_ctor_get(x_22, 2); +x_57 = lean_ctor_get(x_22, 3); +lean_inc(x_57); +lean_inc(x_56); +lean_inc(x_55); +lean_inc(x_54); +lean_dec(x_22); +x_58 = l_Lean_HashMap_insert___at_Lean_Meta_SynthInstance_newSubgoal___spec__1(x_57, x_10, x_13); +x_59 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_59, 0, x_54); +lean_ctor_set(x_59, 1, x_55); +lean_ctor_set(x_59, 2, x_56); +lean_ctor_set(x_59, 3, x_58); +x_60 = lean_st_ref_set(x_4, x_59, x_23); +x_61 = lean_ctor_get(x_60, 1); +lean_inc(x_61); +if (lean_is_exclusive(x_60)) { + lean_ctor_release(x_60, 0); + lean_ctor_release(x_60, 1); + x_62 = x_60; +} else { + lean_dec_ref(x_60); + x_62 = lean_box(0); } -LEAN_EXPORT lean_object* l_ReaderT_read___at_Lean_Meta_SynthInstance_addAnswer___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) { -_start: +x_63 = lean_array_get_size(x_16); +x_64 = lean_unsigned_to_nat(0u); +x_65 = lean_nat_dec_lt(x_64, x_63); +if (x_65 == 0) { -lean_object* x_8; -x_8 = l_ReaderT_read___at_Lean_Meta_SynthInstance_addAnswer___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_object* x_66; lean_object* x_67; +lean_dec(x_63); +lean_dec(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); -return x_8; +x_66 = lean_box(0); +if (lean_is_scalar(x_62)) { + x_67 = lean_alloc_ctor(0, 2, 0); +} else { + x_67 = x_62; } +lean_ctor_set(x_67, 0, x_66); +lean_ctor_set(x_67, 1, x_61); +return x_67; } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_addAnswer___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: +else { -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 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_addAnswer___spec__2(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_10); -lean_dec(x_9); +uint8_t x_68; +x_68 = lean_nat_dec_le(x_63, x_63); +if (x_68 == 0) +{ +lean_object* x_69; lean_object* x_70; +lean_dec(x_63); +lean_dec(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); -return x_15; +x_69 = lean_box(0); +if (lean_is_scalar(x_62)) { + x_70 = lean_alloc_ctor(0, 2, 0); +} else { + x_70 = x_62; } +lean_ctor_set(x_70, 0, x_69); +lean_ctor_set(x_70, 1, x_61); +return x_70; } -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_addAnswer___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: +else { -lean_object* x_9; -x_9 = l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_addAnswer___spec__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +size_t x_71; size_t x_72; lean_object* x_73; lean_object* x_74; +lean_dec(x_62); +x_71 = 0; +x_72 = lean_usize_of_nat(x_63); +lean_dec(x_63); +x_73 = lean_box(0); +x_74 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_addAnswer___spec__2(x_2, x_16, x_71, x_72, x_73, x_3, x_4, x_5, x_6, x_7, x_8, x_61); +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_9; +lean_dec(x_16); +return x_74; } } -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_addAnswer___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: +} +} +} +else { -uint8_t x_15; lean_object* x_16; -x_15 = lean_unbox(x_4); -lean_dec(x_4); -x_16 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_addAnswer___spec__3___lambda__1(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); +lean_object* x_75; lean_object* x_76; lean_object* x_77; uint8_t x_78; +x_75 = lean_ctor_get(x_11, 1); +x_76 = lean_ctor_get(x_13, 0); +x_77 = lean_ctor_get(x_13, 1); +lean_inc(x_77); +lean_inc(x_76); lean_dec(x_13); -lean_dec(x_11); +x_78 = l_Lean_Meta_SynthInstance_isNewAnswer(x_77, x_2); +if (x_78 == 0) +{ +lean_object* x_79; +lean_dec(x_77); +lean_dec(x_76); 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_16; +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_79 = lean_box(0); +lean_ctor_set(x_11, 0, x_79); +return x_11; } -} -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_addAnswer___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) { -_start: +else { -uint8_t x_12; lean_object* x_13; -x_12 = lean_unbox(x_4); -lean_dec(x_4); -x_13 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_addAnswer___spec__3(x_1, x_2, x_3, x_12, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_13; -} +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; uint8_t x_97; +lean_free_object(x_11); +lean_inc(x_2); +x_80 = lean_array_push(x_77, x_2); +lean_inc(x_76); +x_81 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_81, 0, x_76); +lean_ctor_set(x_81, 1, x_80); +x_82 = lean_st_ref_take(x_4, x_75); +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_ctor_get(x_83, 0); +lean_inc(x_85); +x_86 = lean_ctor_get(x_83, 1); +lean_inc(x_86); +x_87 = lean_ctor_get(x_83, 2); +lean_inc(x_87); +x_88 = lean_ctor_get(x_83, 3); +lean_inc(x_88); +if (lean_is_exclusive(x_83)) { + lean_ctor_release(x_83, 0); + lean_ctor_release(x_83, 1); + lean_ctor_release(x_83, 2); + lean_ctor_release(x_83, 3); + x_89 = x_83; +} else { + lean_dec_ref(x_83); + x_89 = lean_box(0); } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_addAnswer___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_10; -x_10 = l_Lean_Meta_SynthInstance_addAnswer___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_10; +x_90 = l_Lean_HashMap_insert___at_Lean_Meta_SynthInstance_newSubgoal___spec__1(x_88, x_10, x_81); +if (lean_is_scalar(x_89)) { + x_91 = lean_alloc_ctor(0, 4, 0); +} else { + x_91 = x_89; } +lean_ctor_set(x_91, 0, x_85); +lean_ctor_set(x_91, 1, x_86); +lean_ctor_set(x_91, 2, x_87); +lean_ctor_set(x_91, 3, x_90); +x_92 = lean_st_ref_set(x_4, x_91, x_84); +x_93 = lean_ctor_get(x_92, 1); +lean_inc(x_93); +if (lean_is_exclusive(x_92)) { + lean_ctor_release(x_92, 0); + lean_ctor_release(x_92, 1); + x_94 = x_92; +} else { + lean_dec_ref(x_92); + x_94 = lean_box(0); } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_addAnswer___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) { -_start: +x_95 = lean_array_get_size(x_76); +x_96 = lean_unsigned_to_nat(0u); +x_97 = lean_nat_dec_lt(x_96, x_95); +if (x_97 == 0) { -lean_object* x_9; -x_9 = l_Lean_Meta_SynthInstance_addAnswer___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_object* x_98; lean_object* x_99; +lean_dec(x_95); +lean_dec(x_76); +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; -} -} -LEAN_EXPORT uint8_t l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_hasUnusedArguments(lean_object* x_1) { -_start: -{ -if (lean_obj_tag(x_1) == 7) -{ -lean_object* x_2; lean_object* x_3; uint8_t x_4; -x_2 = lean_ctor_get(x_1, 2); -x_3 = lean_unsigned_to_nat(0u); -x_4 = lean_expr_has_loose_bvar(x_2, x_3); -if (x_4 == 0) -{ -uint8_t x_5; -x_5 = 1; -return x_5; -} -else -{ -x_1 = x_2; -goto _start; +x_98 = lean_box(0); +if (lean_is_scalar(x_94)) { + x_99 = lean_alloc_ctor(0, 2, 0); +} else { + x_99 = x_94; } +lean_ctor_set(x_99, 0, x_98); +lean_ctor_set(x_99, 1, x_93); +return x_99; } else { -uint8_t x_7; -x_7 = 0; -return x_7; -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_hasUnusedArguments___boxed(lean_object* x_1) { -_start: -{ -uint8_t x_2; lean_object* x_3; -x_2 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_hasUnusedArguments(x_1); -lean_dec(x_1); -x_3 = lean_box(x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l_List_anyM___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___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) { -_start: -{ -if (lean_obj_tag(x_2) == 0) +uint8_t x_100; +x_100 = lean_nat_dec_le(x_95, x_95); +if (x_100 == 0) { -uint8_t x_8; lean_object* x_9; lean_object* x_10; +lean_object* x_101; lean_object* x_102; +lean_dec(x_95); +lean_dec(x_76); +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_8 = 0; -x_9 = lean_box(x_8); -x_10 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_10, 0, x_9); -lean_ctor_set(x_10, 1, x_7); -return x_10; -} -else -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_11 = lean_ctor_get(x_2, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_2, 1); -lean_inc(x_12); lean_dec(x_2); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -x_13 = lean_infer_type(x_11, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_13) == 0) -{ -uint8_t x_14; -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) -{ -lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_15 = lean_ctor_get(x_13, 0); -x_16 = lean_ctor_get(x_13, 1); -x_17 = l_Lean_Expr_hasAnyFVar_visit___at_Lean_Expr_containsFVar___spec__1(x_1, x_15); -lean_dec(x_15); -if (x_17 == 0) -{ -lean_free_object(x_13); -x_2 = x_12; -x_7 = x_16; -goto _start; +x_101 = lean_box(0); +if (lean_is_scalar(x_94)) { + x_102 = lean_alloc_ctor(0, 2, 0); +} else { + x_102 = x_94; +} +lean_ctor_set(x_102, 0, x_101); +lean_ctor_set(x_102, 1, x_93); +return x_102; } else { -uint8_t x_19; lean_object* x_20; -lean_dec(x_12); +size_t x_103; size_t x_104; lean_object* x_105; lean_object* x_106; +lean_dec(x_94); +x_103 = 0; +x_104 = lean_usize_of_nat(x_95); +lean_dec(x_95); +x_105 = lean_box(0); +x_106 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_addAnswer___spec__2(x_2, x_76, x_103, x_104, x_105, x_3, x_4, x_5, x_6, x_7, x_8, x_93); +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_19 = 1; -x_20 = lean_box(x_19); -lean_ctor_set(x_13, 0, x_20); -return x_13; +lean_dec(x_76); +return x_106; +} +} +} } } else { -lean_object* x_21; lean_object* x_22; uint8_t x_23; -x_21 = lean_ctor_get(x_13, 0); -x_22 = lean_ctor_get(x_13, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_13); -x_23 = l_Lean_Expr_hasAnyFVar_visit___at_Lean_Expr_containsFVar___spec__1(x_1, x_21); -lean_dec(x_21); -if (x_23 == 0) -{ -x_2 = x_12; -x_7 = x_22; -goto _start; +lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; uint8_t x_112; +x_107 = lean_ctor_get(x_11, 0); +x_108 = lean_ctor_get(x_11, 1); +lean_inc(x_108); +lean_inc(x_107); +lean_dec(x_11); +x_109 = lean_ctor_get(x_107, 0); +lean_inc(x_109); +x_110 = lean_ctor_get(x_107, 1); +lean_inc(x_110); +if (lean_is_exclusive(x_107)) { + lean_ctor_release(x_107, 0); + lean_ctor_release(x_107, 1); + x_111 = x_107; +} else { + lean_dec_ref(x_107); + x_111 = lean_box(0); } -else +x_112 = l_Lean_Meta_SynthInstance_isNewAnswer(x_110, x_2); +if (x_112 == 0) { -uint8_t x_25; lean_object* x_26; lean_object* x_27; -lean_dec(x_12); +lean_object* x_113; lean_object* x_114; +lean_dec(x_111); +lean_dec(x_110); +lean_dec(x_109); +lean_dec(x_10); +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_25 = 1; -x_26 = lean_box(x_25); -x_27 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_27, 0, x_26); -lean_ctor_set(x_27, 1, x_22); -return x_27; +lean_dec(x_2); +x_113 = lean_box(0); +x_114 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_114, 0, x_113); +lean_ctor_set(x_114, 1, x_108); +return x_114; +} +else +{ +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; uint8_t x_132; +lean_inc(x_2); +x_115 = lean_array_push(x_110, x_2); +lean_inc(x_109); +if (lean_is_scalar(x_111)) { + x_116 = lean_alloc_ctor(0, 2, 0); +} else { + x_116 = x_111; } +lean_ctor_set(x_116, 0, x_109); +lean_ctor_set(x_116, 1, x_115); +x_117 = lean_st_ref_take(x_4, x_108); +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); +x_120 = lean_ctor_get(x_118, 0); +lean_inc(x_120); +x_121 = lean_ctor_get(x_118, 1); +lean_inc(x_121); +x_122 = lean_ctor_get(x_118, 2); +lean_inc(x_122); +x_123 = lean_ctor_get(x_118, 3); +lean_inc(x_123); +if (lean_is_exclusive(x_118)) { + lean_ctor_release(x_118, 0); + lean_ctor_release(x_118, 1); + lean_ctor_release(x_118, 2); + lean_ctor_release(x_118, 3); + x_124 = x_118; +} else { + lean_dec_ref(x_118); + x_124 = lean_box(0); } +x_125 = l_Lean_HashMap_insert___at_Lean_Meta_SynthInstance_newSubgoal___spec__1(x_123, x_10, x_116); +if (lean_is_scalar(x_124)) { + x_126 = lean_alloc_ctor(0, 4, 0); +} else { + x_126 = x_124; } -else +lean_ctor_set(x_126, 0, x_120); +lean_ctor_set(x_126, 1, x_121); +lean_ctor_set(x_126, 2, x_122); +lean_ctor_set(x_126, 3, x_125); +x_127 = lean_st_ref_set(x_4, x_126, x_119); +x_128 = lean_ctor_get(x_127, 1); +lean_inc(x_128); +if (lean_is_exclusive(x_127)) { + lean_ctor_release(x_127, 0); + lean_ctor_release(x_127, 1); + x_129 = x_127; +} else { + lean_dec_ref(x_127); + x_129 = lean_box(0); +} +x_130 = lean_array_get_size(x_109); +x_131 = lean_unsigned_to_nat(0u); +x_132 = lean_nat_dec_lt(x_131, x_130); +if (x_132 == 0) { -uint8_t x_28; -lean_dec(x_12); +lean_object* x_133; lean_object* x_134; +lean_dec(x_130); +lean_dec(x_109); +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_28 = !lean_is_exclusive(x_13); -if (x_28 == 0) +lean_dec(x_2); +x_133 = lean_box(0); +if (lean_is_scalar(x_129)) { + x_134 = lean_alloc_ctor(0, 2, 0); +} else { + x_134 = x_129; +} +lean_ctor_set(x_134, 0, x_133); +lean_ctor_set(x_134, 1, x_128); +return x_134; +} +else { -return x_13; +uint8_t x_135; +x_135 = lean_nat_dec_le(x_130, x_130); +if (x_135 == 0) +{ +lean_object* x_136; lean_object* x_137; +lean_dec(x_130); +lean_dec(x_109); +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_136 = lean_box(0); +if (lean_is_scalar(x_129)) { + x_137 = lean_alloc_ctor(0, 2, 0); +} else { + x_137 = x_129; +} +lean_ctor_set(x_137, 0, x_136); +lean_ctor_set(x_137, 1, x_128); +return x_137; } else { -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; +size_t x_138; size_t x_139; lean_object* x_140; lean_object* x_141; +lean_dec(x_129); +x_138 = 0; +x_139 = lean_usize_of_nat(x_130); +lean_dec(x_130); +x_140 = lean_box(0); +x_141 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_addAnswer___spec__2(x_2, x_109, x_138, x_139, x_140, x_3, x_4, x_5, x_6, x_7, x_8, x_128); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_109); +return x_141; } } } } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___spec__2___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) { -_start: -{ -lean_object* x_11; -x_11 = l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDeclImp___rarg(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) +else { -uint8_t x_12; -x_12 = !lean_is_exclusive(x_11); -if (x_12 == 0) +uint8_t x_142; +lean_dec(x_10); +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_142 = !lean_is_exclusive(x_11); +if (x_142 == 0) { return x_11; } else { -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_inc(x_13); +lean_object* x_143; lean_object* x_144; lean_object* x_145; +x_143 = lean_ctor_get(x_11, 0); +x_144 = lean_ctor_get(x_11, 1); +lean_inc(x_144); +lean_inc(x_143); lean_dec(x_11); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_13); -lean_ctor_set(x_15, 1, x_14); -return x_15; +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; } } -else -{ -uint8_t x_16; -x_16 = !lean_is_exclusive(x_11); -if (x_16 == 0) -{ -return x_11; } -else +} +static lean_object* _init_l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__1() { +_start: { -lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_17 = lean_ctor_get(x_11, 0); -x_18 = lean_ctor_get(x_11, 1); -lean_inc(x_18); -lean_inc(x_17); -lean_dec(x_11); -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_object* x_1; +x_1 = lean_mk_string_from_bytes("answer", 6); +return x_1; } } +static lean_object* _init_l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_6____closed__8; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_6____closed__1; +x_3 = l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___spec__2(lean_object* x_1) { +static lean_object* _init_l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__3() { _start: { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___spec__2___rarg___boxed), 10, 0); +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_crossEmoji; +x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_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, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +static lean_object* _init_l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__4() { _start: { -uint8_t x_11; -x_11 = lean_usize_dec_eq(x_3, x_4); -if (x_11 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; +x_2 = l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__3; +x_3 = lean_alloc_ctor(7, 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_Meta_SynthInstance_addAnswer___lambda__4___closed__5() { +_start: { -size_t x_12; size_t x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_12 = 1; -x_13 = lean_usize_sub(x_3, x_12); -x_14 = lean_array_uget(x_2, x_13); -lean_inc(x_14); -x_15 = l_Lean_Expr_fvarId_x21(x_14); -x_16 = l_Lean_Expr_hasAnyFVar_visit___at_Lean_Expr_containsFVar___spec__1(x_15, x_1); -if (x_16 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__4; +x_2 = l_Lean_Meta_SynthInstance_tryResolve___lambda__1___closed__2; +x_3 = lean_alloc_ctor(7, 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_Meta_SynthInstance_addAnswer___lambda__4___closed__6() { +_start: { -lean_object* x_17; -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_17 = l_List_anyM___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___spec__1(x_15, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_15); -if (lean_obj_tag(x_17) == 0) -{ -lean_object* x_18; uint8_t x_19; -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -x_19 = lean_unbox(x_18); -lean_dec(x_18); -if (x_19 == 0) -{ -lean_object* x_20; -lean_dec(x_14); -x_20 = lean_ctor_get(x_17, 1); -lean_inc(x_20); -lean_dec(x_17); -x_3 = x_13; -x_10 = x_20; -goto _start; -} -else -{ -lean_object* x_22; lean_object* x_23; -x_22 = lean_ctor_get(x_17, 1); -lean_inc(x_22); -lean_dec(x_17); -x_23 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_23, 0, x_14); -lean_ctor_set(x_23, 1, x_5); -x_3 = x_13; -x_5 = x_23; -x_10 = x_22; -goto _start; -} -} -else -{ -uint8_t x_25; -lean_dec(x_14); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -x_25 = !lean_is_exclusive(x_17); -if (x_25 == 0) -{ -return x_17; -} -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_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; -} -} -} -else -{ -lean_object* x_29; -lean_dec(x_15); -x_29 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_29, 0, x_14); -lean_ctor_set(x_29, 1, x_5); -x_3 = x_13; -x_5 = x_29; -goto _start; -} -} -else -{ -lean_object* x_31; -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -x_31 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_31, 0, x_5); -lean_ctor_set(x_31, 1, x_10); -return x_31; -} -} -} -LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___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) { -_start: -{ -uint8_t x_11; -x_11 = lean_usize_dec_eq(x_3, x_4); -if (x_11 == 0) -{ -size_t x_12; size_t x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_12 = 1; -x_13 = lean_usize_sub(x_3, x_12); -x_14 = lean_array_uget(x_2, x_13); -lean_inc(x_14); -x_15 = l_Lean_Expr_fvarId_x21(x_14); -x_16 = l_Lean_Expr_hasAnyFVar_visit___at_Lean_Expr_containsFVar___spec__1(x_15, x_1); -if (x_16 == 0) -{ -lean_object* x_17; -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_17 = l_List_anyM___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___spec__1(x_15, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_15); -if (lean_obj_tag(x_17) == 0) -{ -lean_object* x_18; uint8_t x_19; -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -x_19 = lean_unbox(x_18); -lean_dec(x_18); -if (x_19 == 0) -{ -lean_object* x_20; -lean_dec(x_14); -x_20 = lean_ctor_get(x_17, 1); -lean_inc(x_20); -lean_dec(x_17); -x_3 = x_13; -x_10 = x_20; -goto _start; -} -else -{ -lean_object* x_22; lean_object* x_23; -x_22 = lean_ctor_get(x_17, 1); -lean_inc(x_22); -lean_dec(x_17); -x_23 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_23, 0, x_14); -lean_ctor_set(x_23, 1, x_5); -x_3 = x_13; -x_5 = x_23; -x_10 = x_22; -goto _start; -} -} -else -{ -uint8_t x_25; -lean_dec(x_14); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -x_25 = !lean_is_exclusive(x_17); -if (x_25 == 0) -{ -return x_17; -} -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_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; -} -} -} -else -{ -lean_object* x_29; -lean_dec(x_15); -x_29 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_29, 0, x_14); -lean_ctor_set(x_29, 1, x_5); -x_3 = x_13; -x_5 = x_29; -goto _start; -} -} -else -{ -lean_object* x_31; -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -x_31 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_31, 0, x_5); -lean_ctor_set(x_31, 1, x_10); -return x_31; -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___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; -x_9 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_9, 0, x_1); -lean_ctor_set(x_9, 1, x_2); -x_10 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_10, 0, x_9); -x_11 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_11, 0, x_10); -lean_ctor_set(x_11, 1, x_8); -return x_11; +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(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___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__2___closed__1() { +static lean_object* _init_l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__7() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("unusedArgs", 10); +x_1 = lean_mk_string_from_bytes("(size: ", 7); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__2___closed__2() { +static lean_object* _init_l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__8() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_6____closed__8; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_6____closed__1; -x_3 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__2___closed__1; -x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); -return x_4; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__7; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } -static lean_object* _init_l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__2___closed__3() { +static lean_object* _init_l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__9() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("\nhas unused arguments, reduced type", 35); +x_1 = lean_mk_string_from_bytes(" ≥ ", 5); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__2___closed__4() { +static lean_object* _init_l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__10() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__2___closed__3; +x_1 = l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__9; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__2___closed__5() { +static lean_object* _init_l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__11() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("\nTransformer", 12); +x_1 = lean_mk_string_from_bytes(")", 1); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__2___closed__6() { +static lean_object* _init_l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__12() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__2___closed__5; +x_1 = l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__11; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_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) { +LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_addAnswer___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_11; uint8_t x_12; uint8_t x_13; uint8_t x_14; lean_object* x_15; -lean_inc(x_5); -x_11 = l_Lean_mkAppN(x_5, x_1); -x_12 = 0; -x_13 = 1; -x_14 = 1; -x_15 = l_Lean_Meta_mkLambdaFVars(x_2, x_11, x_12, x_13, x_14, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_15) == 0) +lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_10 = lean_ctor_get(x_2, 0); +lean_inc(x_10); +lean_dec(x_2); +x_11 = lean_ctor_get(x_1, 4); +lean_inc(x_11); +x_12 = lean_nat_dec_le(x_10, x_11); +lean_dec(x_10); +if (x_12 == 0) { -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, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); -lean_inc(x_17); -lean_dec(x_15); -x_18 = l_Lean_Meta_SynthInstance_newSubgoal___lambda__2___closed__3; -x_19 = lean_array_push(x_18, x_5); -x_20 = l_Lean_Meta_mkLambdaFVars(x_19, x_16, x_12, x_13, x_14, x_6, x_7, x_8, x_9, x_17); -if (lean_obj_tag(x_20) == 0) +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_dec(x_11); +lean_inc(x_1); +x_13 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_addAnswer___lambda__1___boxed), 9, 1); +lean_closure_set(x_13, 0, x_1); +lean_inc(x_1); +x_14 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_addAnswer___lambda__2___boxed), 8, 1); +lean_closure_set(x_14, 0, x_1); +x_15 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_addAnswer___lambda__3), 9, 1); +lean_closure_set(x_15, 0, x_1); +x_16 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_SynthInstance_newSubgoal___spec__7___rarg), 9, 2); +lean_closure_set(x_16, 0, x_14); +lean_closure_set(x_16, 1, x_15); +x_17 = l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__2; +x_18 = 1; +x_19 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_addAnswer___spec__3(x_17, x_13, x_16, x_18, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_19; +} +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_20, 0); -lean_inc(x_21); -x_22 = lean_ctor_get(x_20, 1); +lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_20 = l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__2; +x_21 = l_Lean_isTracingEnabledFor___at_Lean_Meta_SynthInstance_newSubgoal___spec__10(x_20, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_22 = lean_ctor_get(x_21, 0); lean_inc(x_22); -lean_dec(x_20); -x_23 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__2___closed__2; -x_24 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_23, x_6, x_7, x_8, x_9, x_22); -x_25 = lean_ctor_get(x_24, 0); -lean_inc(x_25); -x_26 = lean_unbox(x_25); +x_23 = lean_unbox(x_22); +lean_dec(x_22); +if (x_23 == 0) +{ +uint8_t x_24; +lean_dec(x_11); +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_is_exclusive(x_21); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; +x_25 = lean_ctor_get(x_21, 0); lean_dec(x_25); -if (x_26 == 0) +x_26 = lean_box(0); +lean_ctor_set(x_21, 0, x_26); +return x_21; +} +else { lean_object* x_27; lean_object* x_28; lean_object* x_29; -lean_dec(x_4); -x_27 = lean_ctor_get(x_24, 1); +x_27 = lean_ctor_get(x_21, 1); lean_inc(x_27); -lean_dec(x_24); +lean_dec(x_21); x_28 = lean_box(0); -x_29 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__1(x_3, x_21, x_28, x_6, x_7, x_8, x_9, x_27); -lean_dec(x_8); -lean_dec(x_6); +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_30; lean_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_30 = lean_ctor_get(x_24, 1); +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_21, 1); lean_inc(x_30); -lean_dec(x_24); -x_31 = l_Lean_MessageData_ofExpr(x_4); -x_32 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; -x_33 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_33, 0, x_32); -lean_ctor_set(x_33, 1, x_31); -x_34 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__2___closed__4; -x_35 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -lean_inc(x_3); -x_36 = l_Lean_indentExpr(x_3); -x_37 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_37, 0, x_35); -lean_ctor_set(x_37, 1, x_36); -x_38 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__2___closed__6; -x_39 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_39, 0, x_37); -lean_ctor_set(x_39, 1, x_38); -lean_inc(x_21); -x_40 = l_Lean_indentExpr(x_21); -x_41 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_41, 0, x_39); -lean_ctor_set(x_41, 1, x_40); +lean_dec(x_21); +x_31 = lean_ctor_get(x_1, 0); +lean_inc(x_31); +lean_dec(x_1); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_32 = lean_infer_type(x_31, x_5, x_6, x_7, x_8, x_30); +if (lean_obj_tag(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; 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_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_instantiateMVars___at_Lean_Meta_SynthInstance_mkTableKeyFor___spec__1(x_33, x_3, x_4, x_5, x_6, x_7, x_8, 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); +lean_dec(x_35); +x_38 = l_Lean_MessageData_ofExpr(x_36); +x_39 = l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__5; +x_40 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_38); +x_41 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; x_42 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_32); -x_43 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_23, x_42, x_6, x_7, x_8, x_9, x_30); -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 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__1(x_3, x_21, x_44, x_6, x_7, x_8, x_9, x_45); +lean_ctor_set(x_42, 0, x_40); +lean_ctor_set(x_42, 1, x_41); +x_43 = l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__6; +x_44 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_44, 0, x_42); +lean_ctor_set(x_44, 1, x_43); +x_45 = l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__8; +x_46 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +x_47 = l_Nat_repr(x_11); +x_48 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_48, 0, x_47); +x_49 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_49, 0, x_48); +x_50 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_50, 0, x_46); +lean_ctor_set(x_50, 1, x_49); +x_51 = l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__10; +x_52 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_52, 0, x_50); +lean_ctor_set(x_52, 1, x_51); +x_53 = lean_ctor_get(x_3, 0); +lean_inc(x_53); +x_54 = l_Nat_repr(x_53); +x_55 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_55, 0, x_54); +x_56 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_56, 0, x_55); +x_57 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_57, 0, x_52); +lean_ctor_set(x_57, 1, x_56); +x_58 = l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__12; +x_59 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_59, 0, x_57); +lean_ctor_set(x_59, 1, x_58); +x_60 = l_Lean_addTrace___at_Lean_Meta_SynthInstance_wakeUp___spec__1(x_20, x_59, 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_44); -return x_46; -} +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_60; } else { -uint8_t x_47; +uint8_t x_61; +lean_dec(x_11); 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_47 = !lean_is_exclusive(x_20); -if (x_47 == 0) +x_61 = !lean_is_exclusive(x_32); +if (x_61 == 0) { -return x_20; +return x_32; } else { -lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_48 = lean_ctor_get(x_20, 0); -x_49 = lean_ctor_get(x_20, 1); -lean_inc(x_49); -lean_inc(x_48); -lean_dec(x_20); -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; +lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_62 = lean_ctor_get(x_32, 0); +x_63 = lean_ctor_get(x_32, 1); +lean_inc(x_63); +lean_inc(x_62); +lean_dec(x_32); +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 -{ -uint8_t x_51; -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -x_51 = !lean_is_exclusive(x_15); -if (x_51 == 0) -{ -return x_15; } -else +} +} +static lean_object* _init_l_Lean_Meta_SynthInstance_addAnswer___closed__1() { +_start: { -lean_object* x_52; lean_object* x_53; lean_object* x_54; -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_15); -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_object* x_1; +x_1 = lean_alloc_closure((void*)(l_ReaderT_read___at_Lean_Meta_SynthInstance_addAnswer___spec__1___boxed), 7, 0); +return x_1; } } -static lean_object* _init_l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__3___closed__1() { +LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_addAnswer(lean_object* x_1, lean_object* x_2, 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_1; -x_1 = lean_mk_string_from_bytes("redf", 4); -return x_1; +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_9 = lean_ctor_get(x_1, 2); +lean_inc(x_9); +x_10 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_addAnswer___lambda__4), 9, 1); +lean_closure_set(x_10, 0, x_1); +x_11 = l_Lean_Meta_SynthInstance_addAnswer___closed__1; +x_12 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_SynthInstance_newSubgoal___spec__7___rarg), 9, 2); +lean_closure_set(x_12, 0, x_11); +lean_closure_set(x_12, 1, x_10); +x_13 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__16___rarg(x_9, x_12, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +return x_13; } } -static lean_object* _init_l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__3___closed__2() { +LEAN_EXPORT lean_object* l_ReaderT_read___at_Lean_Meta_SynthInstance_addAnswer___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) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__3___closed__1; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; +lean_object* x_8; +x_8 = l_ReaderT_read___at_Lean_Meta_SynthInstance_addAnswer___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_8; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_addAnswer___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: { -lean_object* x_9; lean_object* x_10; uint8_t x_11; lean_object* x_12; lean_object* x_13; -x_9 = lean_box(0); -x_10 = lean_array_get_size(x_2); -x_11 = lean_nat_dec_le(x_10, x_10); -if (x_11 == 0) -{ -lean_object* x_33; uint8_t x_34; -x_33 = lean_unsigned_to_nat(0u); -x_34 = lean_nat_dec_lt(x_33, x_10); -if (x_34 == 0) -{ +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 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_addAnswer___spec__2(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_10); -x_12 = x_9; -x_13 = x_8; -goto block_32; +lean_dec(x_9); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +return x_15; } -else -{ -size_t x_35; size_t x_36; lean_object* x_37; -x_35 = lean_usize_of_nat(x_10); -lean_dec(x_10); -x_36 = 0; -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_37 = l_Array_foldrMUnsafe_fold___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___spec__3(x_3, x_2, x_35, x_36, x_9, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_37) == 0) -{ -lean_object* x_38; lean_object* x_39; -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_12 = x_38; -x_13 = x_39; -goto block_32; } -else +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_addAnswer___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) { +_start: { -uint8_t x_40; +lean_object* x_9; +x_9 = l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_addAnswer___spec__5(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); @@ -12680,8057 +12277,9530 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_40 = !lean_is_exclusive(x_37); -if (x_40 == 0) -{ -return x_37; +return x_9; } -else +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_addAnswer___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_41; lean_object* x_42; lean_object* x_43; -x_41 = lean_ctor_get(x_37, 0); -x_42 = lean_ctor_get(x_37, 1); -lean_inc(x_42); -lean_inc(x_41); -lean_dec(x_37); -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; +uint8_t x_15; lean_object* x_16; +x_15 = lean_unbox(x_4); +lean_dec(x_4); +x_16 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_addAnswer___spec__3___lambda__1(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); +lean_dec(x_13); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_7); +lean_dec(x_5); +return x_16; } } +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_addAnswer___spec__3___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) { +_start: +{ +uint8_t x_17; uint8_t x_18; double x_19; lean_object* x_20; +x_17 = lean_unbox(x_5); +lean_dec(x_5); +x_18 = lean_unbox(x_7); +lean_dec(x_7); +x_19 = lean_unbox_float(x_8); +lean_dec(x_8); +x_20 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_addAnswer___spec__3___lambda__2(x_1, x_2, x_3, x_4, x_17, x_6, x_18, x_19, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +return x_20; } } -else -{ -lean_object* x_44; uint8_t x_45; -x_44 = lean_unsigned_to_nat(0u); -x_45 = lean_nat_dec_lt(x_44, x_10); -if (x_45 == 0) +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_addAnswer___spec__3___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) { +_start: { -lean_dec(x_10); -x_12 = x_9; -x_13 = x_8; -goto block_32; +uint8_t x_15; uint8_t x_16; lean_object* x_17; +x_15 = lean_unbox(x_4); +lean_dec(x_4); +x_16 = lean_unbox(x_6); +lean_dec(x_6); +x_17 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_addAnswer___spec__3___lambda__3(x_1, x_2, x_3, x_15, x_5, x_16, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +return x_17; } -else -{ -size_t x_46; size_t x_47; lean_object* x_48; -x_46 = lean_usize_of_nat(x_10); -lean_dec(x_10); -x_47 = 0; -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_48 = l_Array_foldrMUnsafe_fold___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___spec__4(x_3, x_2, x_46, x_47, x_9, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_48) == 0) +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_addAnswer___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) { +_start: { -lean_object* x_49; lean_object* x_50; -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_12 = x_49; -x_13 = x_50; -goto block_32; +uint8_t x_12; lean_object* x_13; +x_12 = lean_unbox(x_4); +lean_dec(x_4); +x_13 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_addAnswer___spec__3(x_1, x_2, x_3, x_12, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_13; } -else +} +LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_addAnswer___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: { -uint8_t x_51; -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +lean_object* x_10; +x_10 = l_Lean_Meta_SynthInstance_addAnswer___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_51 = !lean_is_exclusive(x_48); -if (x_51 == 0) -{ -return x_48; -} -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_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; +return x_10; } } +LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_addAnswer___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) { +_start: +{ +lean_object* x_9; +x_9 = l_Lean_Meta_SynthInstance_addAnswer___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_3); +lean_dec(x_2); +return x_9; } } -block_32: +LEAN_EXPORT uint8_t l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_hasUnusedArguments(lean_object* x_1) { +_start: { -lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; uint8_t x_18; uint8_t x_19; lean_object* x_20; -x_14 = l_List_redLength___rarg(x_12); -x_15 = lean_mk_empty_array_with_capacity(x_14); -lean_dec(x_14); -x_16 = l_List_toArrayAux___rarg(x_12, x_15); -x_17 = 0; -x_18 = 1; -x_19 = 1; -lean_inc(x_16); -x_20 = l_Lean_Meta_mkForallFVars(x_16, x_3, x_17, x_18, x_19, x_4, x_5, x_6, x_7, x_13); -if (lean_obj_tag(x_20) == 0) +if (lean_obj_tag(x_1) == 7) { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; uint8_t x_26; lean_object* x_27; -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_21); -x_23 = lean_alloc_closure((void*)(l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__2___boxed), 10, 4); -lean_closure_set(x_23, 0, x_16); -lean_closure_set(x_23, 1, x_2); -lean_closure_set(x_23, 2, x_21); -lean_closure_set(x_23, 3, x_1); -x_24 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__3___closed__2; -x_25 = 0; -x_26 = 0; -x_27 = l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___spec__2___rarg(x_24, x_25, x_21, x_23, x_26, x_4, x_5, x_6, x_7, x_22); -return x_27; +lean_object* x_2; lean_object* x_3; uint8_t x_4; +x_2 = lean_ctor_get(x_1, 2); +x_3 = lean_unsigned_to_nat(0u); +x_4 = lean_expr_has_loose_bvar(x_2, x_3); +if (x_4 == 0) +{ +uint8_t x_5; +x_5 = 1; +return x_5; } else { -uint8_t x_28; -lean_dec(x_16); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -x_28 = !lean_is_exclusive(x_20); -if (x_28 == 0) -{ -return x_20; +x_1 = x_2; +goto _start; +} } else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = lean_ctor_get(x_20, 0); -x_30 = lean_ctor_get(x_20, 1); -lean_inc(x_30); -lean_inc(x_29); -lean_dec(x_20); -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; +uint8_t x_7; +x_7 = 0; +return x_7; } } } +LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_hasUnusedArguments___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_hasUnusedArguments(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___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_EXPORT lean_object* l_List_anyM___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___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) { _start: { -lean_object* x_7; +if (lean_obj_tag(x_2) == 0) +{ +uint8_t x_8; lean_object* x_9; lean_object* x_10; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_8 = 0; +x_9 = lean_box(x_8); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_7); +return x_10; +} +else +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = lean_ctor_get(x_2, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_2, 1); +lean_inc(x_12); +lean_dec(x_2); +lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -lean_inc(x_2); -x_7 = lean_infer_type(x_1, x_2, x_3, x_4, x_5, x_6); -if (lean_obj_tag(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_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_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_8, x_2, x_3, x_4, x_5, x_9); -x_11 = !lean_is_exclusive(x_10); -if (x_11 == 0) +x_13 = lean_infer_type(x_11, x_3, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_13) == 0) { -lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_12 = lean_ctor_get(x_10, 0); -x_13 = lean_ctor_get(x_10, 1); -x_14 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_hasUnusedArguments(x_12); +uint8_t x_14; +x_14 = !lean_is_exclusive(x_13); if (x_14 == 0) { -lean_object* x_15; +lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_15 = lean_ctor_get(x_13, 0); +x_16 = lean_ctor_get(x_13, 1); +x_17 = l_Lean_Expr_hasAnyFVar_visit___at_Lean_Expr_containsFVar___spec__1(x_1, x_15); +lean_dec(x_15); +if (x_17 == 0) +{ +lean_free_object(x_13); +x_2 = x_12; +x_7 = x_16; +goto _start; +} +else +{ +uint8_t x_19; lean_object* x_20; lean_dec(x_12); +lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -x_15 = lean_box(0); -lean_ctor_set(x_10, 0, x_15); -return x_10; +x_19 = 1; +x_20 = lean_box(x_19); +lean_ctor_set(x_13, 0, x_20); +return x_13; +} } else { -lean_object* x_16; lean_object* x_17; -lean_free_object(x_10); -lean_inc(x_12); -x_16 = lean_alloc_closure((void*)(l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__3), 8, 1); -lean_closure_set(x_16, 0, x_12); -x_17 = l_Lean_Meta_forallTelescope___at___private_Lean_Meta_InferType_0__Lean_Meta_inferForallType___spec__2___rarg(x_12, x_16, x_2, x_3, x_4, x_5, x_13); -return x_17; -} +lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_21 = lean_ctor_get(x_13, 0); +x_22 = lean_ctor_get(x_13, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_13); +x_23 = l_Lean_Expr_hasAnyFVar_visit___at_Lean_Expr_containsFVar___spec__1(x_1, x_21); +lean_dec(x_21); +if (x_23 == 0) +{ +x_2 = x_12; +x_7 = x_22; +goto _start; } else { -lean_object* x_18; lean_object* x_19; uint8_t x_20; -x_18 = lean_ctor_get(x_10, 0); -x_19 = lean_ctor_get(x_10, 1); -lean_inc(x_19); -lean_inc(x_18); -lean_dec(x_10); -x_20 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_hasUnusedArguments(x_18); -if (x_20 == 0) -{ -lean_object* x_21; lean_object* x_22; -lean_dec(x_18); +uint8_t x_25; lean_object* x_26; lean_object* x_27; +lean_dec(x_12); +lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -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_19); -return x_22; -} -else -{ -lean_object* x_23; lean_object* x_24; -lean_inc(x_18); -x_23 = lean_alloc_closure((void*)(l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__3), 8, 1); -lean_closure_set(x_23, 0, x_18); -x_24 = l_Lean_Meta_forallTelescope___at___private_Lean_Meta_InferType_0__Lean_Meta_inferForallType___spec__2___rarg(x_18, x_23, x_2, x_3, x_4, x_5, x_19); -return x_24; +x_25 = 1; +x_26 = lean_box(x_25); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_22); +return x_27; } } } else { -uint8_t x_25; +uint8_t x_28; +lean_dec(x_12); +lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -x_25 = !lean_is_exclusive(x_7); -if (x_25 == 0) +x_28 = !lean_is_exclusive(x_13); +if (x_28 == 0) { -return x_7; +return x_13; } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_26 = lean_ctor_get(x_7, 0); -x_27 = lean_ctor_get(x_7, 1); -lean_inc(x_27); -lean_inc(x_26); -lean_dec(x_7); -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_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; } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_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) { -_start: -{ -lean_object* x_8; lean_object* x_9; -x_8 = lean_alloc_closure((void*)(l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__4), 6, 1); -lean_closure_set(x_8, 0, x_2); -x_9 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_tryResolve___spec__1___rarg(x_1, x_8, x_3, x_4, x_5, x_6, x_7); -return x_9; } } -LEAN_EXPORT lean_object* l_List_anyM___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___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_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___spec__2___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) { _start: { -lean_object* x_8; -x_8 = l_List_anyM___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___spec__1(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_Meta_withLocalDecl___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___spec__2___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: +lean_object* x_11; +x_11 = l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDeclImp___rarg(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_11; uint8_t x_12; lean_object* x_13; -x_11 = lean_unbox(x_2); -lean_dec(x_2); -x_12 = lean_unbox(x_5); -lean_dec(x_5); -x_13 = l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___spec__2___rarg(x_1, x_11, x_3, x_4, x_12, x_6, x_7, x_8, x_9, x_10); -return x_13; -} +uint8_t x_12; +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) +{ +return x_11; } -LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___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: +else { -size_t x_11; size_t x_12; lean_object* x_13; -x_11 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_12 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_13 = l_Array_foldrMUnsafe_fold___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___spec__3(x_1, x_2, x_11, x_12, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_2); -lean_dec(x_1); -return x_13; +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_inc(x_13); +lean_dec(x_11); +x_15 = lean_alloc_ctor(0, 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_Array_foldrMUnsafe_fold___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___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: +else { -size_t x_11; size_t x_12; lean_object* x_13; -x_11 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_12 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_13 = l_Array_foldrMUnsafe_fold___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___spec__4(x_1, x_2, x_11, x_12, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_2); -lean_dec(x_1); -return x_13; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_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) { -_start: -{ -lean_object* x_9; -x_9 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__1(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___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_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) { -_start: +uint8_t x_16; +x_16 = !lean_is_exclusive(x_11); +if (x_16 == 0) { -lean_object* x_11; -x_11 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___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_9); -lean_dec(x_7); return x_11; } -} -LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at_Lean_Meta_SynthInstance_consume___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_9; uint8_t x_10; -x_9 = lean_st_ref_get(x_5, x_8); -x_10 = !lean_is_exclusive(x_9); -if (x_10 == 0) -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; -x_11 = lean_ctor_get(x_9, 0); -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -lean_dec(x_11); -x_13 = lean_ctor_get(x_12, 7); -lean_inc(x_13); -lean_dec(x_12); -x_14 = l_Lean_PersistentHashMap_contains___at_Lean_MVarId_isAssigned___spec__1(x_13, x_1); -x_15 = lean_box(x_14); -lean_ctor_set(x_9, 0, x_15); -return x_9; -} 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; lean_object* x_22; -x_16 = lean_ctor_get(x_9, 0); -x_17 = lean_ctor_get(x_9, 1); -lean_inc(x_17); -lean_inc(x_16); -lean_dec(x_9); -x_18 = lean_ctor_get(x_16, 0); +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_11, 0); +x_18 = lean_ctor_get(x_11, 1); lean_inc(x_18); -lean_dec(x_16); -x_19 = lean_ctor_get(x_18, 7); -lean_inc(x_19); -lean_dec(x_18); -x_20 = l_Lean_PersistentHashMap_contains___at_Lean_MVarId_isAssigned___spec__1(x_19, x_1); -x_21 = lean_box(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; +lean_inc(x_17); +lean_dec(x_11); +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_List_filterAuxM___at_Lean_Meta_SynthInstance_consume___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_Meta_withLocalDecl___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___spec__2(lean_object* x_1) { _start: { -if (lean_obj_tag(x_1) == 0) -{ -lean_object* x_10; -x_10 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_10, 0, x_2); -lean_ctor_set(x_10, 1, x_9); -return x_10; +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___spec__2___rarg___boxed), 10, 0); +return x_2; } -else +} +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_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, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: { uint8_t x_11; -x_11 = !lean_is_exclusive(x_1); +x_11 = lean_usize_dec_eq(x_3, 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; uint8_t x_17; -x_12 = lean_ctor_get(x_1, 0); -x_13 = lean_ctor_get(x_1, 1); -lean_inc(x_12); -x_14 = l_Lean_Expr_mvarId_x21(x_12); -x_15 = l_Lean_MVarId_isAssigned___at_Lean_Meta_SynthInstance_consume___spec__1(x_14, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_unbox(x_16); -lean_dec(x_16); -if (x_17 == 0) +size_t x_12; size_t x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_12 = 1; +x_13 = lean_usize_sub(x_3, x_12); +x_14 = lean_array_uget(x_2, x_13); +lean_inc(x_14); +x_15 = l_Lean_Expr_fvarId_x21(x_14); +x_16 = l_Lean_Expr_hasAnyFVar_visit___at_Lean_Expr_containsFVar___spec__1(x_15, x_1); +if (x_16 == 0) { -lean_object* x_18; -x_18 = lean_ctor_get(x_15, 1); -lean_inc(x_18); +lean_object* x_17; +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_17 = l_List_anyM___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___spec__1(x_15, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_15); -lean_ctor_set(x_1, 1, x_2); +if (lean_obj_tag(x_17) == 0) { -lean_object* _tmp_0 = x_13; -lean_object* _tmp_1 = x_1; -lean_object* _tmp_8 = x_18; -x_1 = _tmp_0; -x_2 = _tmp_1; -x_9 = _tmp_8; -} +lean_object* x_18; uint8_t x_19; +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_unbox(x_18); +lean_dec(x_18); +if (x_19 == 0) +{ +lean_object* x_20; +lean_dec(x_14); +x_20 = lean_ctor_get(x_17, 1); +lean_inc(x_20); +lean_dec(x_17); +x_3 = x_13; +x_10 = x_20; goto _start; } else { -lean_object* x_20; -lean_free_object(x_1); -lean_dec(x_12); -x_20 = lean_ctor_get(x_15, 1); -lean_inc(x_20); -lean_dec(x_15); -x_1 = x_13; -x_9 = x_20; +lean_object* x_22; lean_object* x_23; +x_22 = lean_ctor_get(x_17, 1); +lean_inc(x_22); +lean_dec(x_17); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_14); +lean_ctor_set(x_23, 1, x_5); +x_3 = x_13; +x_5 = x_23; +x_10 = x_22; goto _start; } } else { -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_22 = lean_ctor_get(x_1, 0); -x_23 = lean_ctor_get(x_1, 1); -lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_1); -lean_inc(x_22); -x_24 = l_Lean_Expr_mvarId_x21(x_22); -x_25 = l_Lean_MVarId_isAssigned___at_Lean_Meta_SynthInstance_consume___spec__1(x_24, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -x_26 = lean_ctor_get(x_25, 0); +uint8_t x_25; +lean_dec(x_14); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +x_25 = !lean_is_exclusive(x_17); +if (x_25 == 0) +{ +return x_17; +} +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_inc(x_26); -x_27 = lean_unbox(x_26); -lean_dec(x_26); -if (x_27 == 0) +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; +} +} +} +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); +lean_object* x_29; +lean_dec(x_15); x_29 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_29, 0, x_22); -lean_ctor_set(x_29, 1, x_2); -x_1 = x_23; -x_2 = x_29; -x_9 = x_28; +lean_ctor_set(x_29, 0, x_14); +lean_ctor_set(x_29, 1, x_5); +x_3 = x_13; +x_5 = x_29; goto _start; } +} else { lean_object* x_31; -lean_dec(x_22); -x_31 = lean_ctor_get(x_25, 1); -lean_inc(x_31); -lean_dec(x_25); -x_1 = x_23; -x_9 = x_31; -goto _start; -} -} +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_5); +lean_ctor_set(x_31, 1, x_10); +return x_31; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_SynthInstance_consume___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_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___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) { _start: { -lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; lean_object* x_16; -x_12 = l_Lean_Meta_SynthInstance_newSubgoal___lambda__2___closed__3; -x_13 = lean_array_push(x_12, x_4); -x_14 = 0; -x_15 = l_Lean_Expr_betaRev(x_1, x_13, x_14, x_14); -lean_dec(x_13); -lean_inc(x_15); -x_16 = lean_infer_type(x_15, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_16) == 0) +uint8_t x_11; +x_11 = lean_usize_dec_eq(x_3, x_4); +if (x_11 == 0) { -uint8_t x_17; -x_17 = !lean_is_exclusive(x_16); -if (x_17 == 0) +size_t x_12; size_t x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_12 = 1; +x_13 = lean_usize_sub(x_3, x_12); +x_14 = lean_array_uget(x_2, x_13); +lean_inc(x_14); +x_15 = l_Lean_Expr_fvarId_x21(x_14); +x_16 = l_Lean_Expr_hasAnyFVar_visit___at_Lean_Expr_containsFVar___spec__1(x_15, x_1); +if (x_16 == 0) { -uint8_t x_18; -x_18 = !lean_is_exclusive(x_2); -if (x_18 == 0) +lean_object* x_17; +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_17 = l_List_anyM___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___spec__1(x_15, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_15); +if (lean_obj_tag(x_17) == 0) { -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_19 = lean_ctor_get(x_16, 0); -x_20 = lean_ctor_get(x_2, 2); -lean_dec(x_20); -lean_ctor_set(x_2, 2, x_15); -x_21 = lean_ctor_get(x_3, 2); -lean_inc(x_21); -x_22 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_22, 0, x_2); -lean_ctor_set(x_22, 1, x_19); -lean_ctor_set(x_22, 2, x_21); -lean_ctor_set(x_16, 0, x_22); -return x_16; -} -else +lean_object* x_18; uint8_t x_19; +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_unbox(x_18); +lean_dec(x_18); +if (x_19 == 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_16, 0); -x_24 = lean_ctor_get(x_2, 0); -x_25 = lean_ctor_get(x_2, 1); -lean_inc(x_25); -lean_inc(x_24); -lean_dec(x_2); -x_26 = lean_alloc_ctor(0, 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_15); -x_27 = lean_ctor_get(x_3, 2); -lean_inc(x_27); -x_28 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_23); -lean_ctor_set(x_28, 2, x_27); -lean_ctor_set(x_16, 0, x_28); -return x_16; -} +lean_object* x_20; +lean_dec(x_14); +x_20 = lean_ctor_get(x_17, 1); +lean_inc(x_20); +lean_dec(x_17); +x_3 = x_13; +x_10 = x_20; +goto _start; } 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; lean_object* x_37; -x_29 = lean_ctor_get(x_16, 0); -x_30 = lean_ctor_get(x_16, 1); -lean_inc(x_30); -lean_inc(x_29); -lean_dec(x_16); -x_31 = lean_ctor_get(x_2, 0); -lean_inc(x_31); -x_32 = lean_ctor_get(x_2, 1); -lean_inc(x_32); -if (lean_is_exclusive(x_2)) { - lean_ctor_release(x_2, 0); - lean_ctor_release(x_2, 1); - lean_ctor_release(x_2, 2); - x_33 = x_2; -} else { - lean_dec_ref(x_2); - x_33 = lean_box(0); -} -if (lean_is_scalar(x_33)) { - x_34 = lean_alloc_ctor(0, 3, 0); -} else { - x_34 = x_33; -} -lean_ctor_set(x_34, 0, x_31); -lean_ctor_set(x_34, 1, x_32); -lean_ctor_set(x_34, 2, x_15); -x_35 = lean_ctor_get(x_3, 2); -lean_inc(x_35); -x_36 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_36, 0, x_34); -lean_ctor_set(x_36, 1, x_29); -lean_ctor_set(x_36, 2, x_35); -x_37 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_37, 0, x_36); -lean_ctor_set(x_37, 1, x_30); -return x_37; +lean_object* x_22; lean_object* x_23; +x_22 = lean_ctor_get(x_17, 1); +lean_inc(x_22); +lean_dec(x_17); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_14); +lean_ctor_set(x_23, 1, x_5); +x_3 = x_13; +x_5 = x_23; +x_10 = x_22; +goto _start; } } else { -uint8_t x_38; -lean_dec(x_15); -lean_dec(x_2); -x_38 = !lean_is_exclusive(x_16); -if (x_38 == 0) +uint8_t x_25; +lean_dec(x_14); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +x_25 = !lean_is_exclusive(x_17); +if (x_25 == 0) { -return x_16; +return x_17; } else { -lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_39 = lean_ctor_get(x_16, 0); -x_40 = lean_ctor_get(x_16, 1); -lean_inc(x_40); -lean_inc(x_39); -lean_dec(x_16); -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_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_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_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_SynthInstance_consume___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) { -_start: -{ -uint8_t x_13; -x_13 = lean_usize_dec_lt(x_4, x_3); -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_2); -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; } 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; -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); -x_18 = lean_ctor_get(x_15, 0); -lean_inc(x_18); -x_19 = lean_ctor_get(x_18, 2); -lean_inc(x_19); -x_20 = lean_alloc_closure((void*)(l_Lean_instantiateMVars___at_Lean_Meta_SynthInstance_mkTableKeyFor___spec__1___boxed), 8, 1); -lean_closure_set(x_20, 0, x_19); -lean_inc(x_2); -x_21 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_Meta_SynthInstance_consume___spec__3___lambda__1___boxed), 11, 3); -lean_closure_set(x_21, 0, x_2); -lean_closure_set(x_21, 1, x_18); -lean_closure_set(x_21, 2, x_15); -x_22 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_SynthInstance_newSubgoal___spec__7___rarg), 9, 2); -lean_closure_set(x_22, 0, x_20); -lean_closure_set(x_22, 1, x_21); -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_1); -x_23 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__18___rarg(x_1, x_22, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_23) == 0) -{ -lean_object* x_24; lean_object* x_25; size_t x_26; size_t x_27; lean_object* x_28; -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 = 1; -x_27 = lean_usize_add(x_4, x_26); -x_28 = lean_array_uset(x_17, x_4, x_24); -x_4 = x_27; -x_5 = x_28; -x_12 = x_25; +lean_object* x_29; +lean_dec(x_15); +x_29 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_29, 0, x_14); +lean_ctor_set(x_29, 1, x_5); +x_3 = x_13; +x_5 = x_29; goto _start; } +} else { -uint8_t x_30; -lean_dec(x_17); -lean_dec(x_11); -lean_dec(x_10); +lean_object* x_31; 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_30 = !lean_is_exclusive(x_23); -if (x_30 == 0) -{ -return x_23; -} -else -{ -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_23, 0); -x_32 = lean_ctor_get(x_23, 1); -lean_inc(x_32); -lean_inc(x_31); -lean_dec(x_23); -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; -} -} +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_5); +lean_ctor_set(x_31, 1, x_10); +return x_31; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_consume___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___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___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: { -uint8_t x_6; -x_6 = lean_usize_dec_eq(x_3, x_4); -if (x_6 == 0) -{ -lean_object* x_7; lean_object* x_8; lean_object* x_9; size_t x_10; size_t x_11; -x_7 = lean_array_uget(x_2, x_3); -lean_inc(x_1); -x_8 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_8, 0, x_1); -lean_ctor_set(x_8, 1, x_7); -x_9 = lean_array_push(x_5, x_8); -x_10 = 1; -x_11 = lean_usize_add(x_3, x_10); -x_3 = x_11; -x_5 = x_9; -goto _start; -} -else -{ -lean_dec(x_1); -return x_5; -} +lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_1); +lean_ctor_set(x_9, 1, x_2); +x_10 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_10, 0, x_9); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_8); +return x_11; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_consume___spec__5(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { +static lean_object* _init_l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__2___closed__1() { _start: { -uint8_t x_6; -x_6 = lean_usize_dec_eq(x_3, x_4); -if (x_6 == 0) -{ -lean_object* x_7; lean_object* x_8; lean_object* x_9; size_t x_10; size_t x_11; -x_7 = lean_array_uget(x_2, x_3); -lean_inc(x_1); -x_8 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_8, 0, x_1); -lean_ctor_set(x_8, 1, x_7); -x_9 = lean_array_push(x_5, x_8); -x_10 = 1; -x_11 = lean_usize_add(x_3, x_10); -x_3 = x_11; -x_5 = x_9; -goto _start; -} -else -{ -lean_dec(x_1); -return x_5; -} +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("unusedArgs", 10); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_consume___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) { +static lean_object* _init_l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__2___closed__2() { _start: { -lean_object* x_9; lean_object* x_10; -x_9 = l_List_reverse___rarg(x_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); -return x_10; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_6____closed__8; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_6____closed__1; +x_3 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__2___closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_consume___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) { +static lean_object* _init_l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__2___closed__3() { _start: { -uint8_t x_9; lean_object* x_10; lean_object* x_11; -x_9 = 0; -x_10 = lean_box(0); -x_11 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkFreshExprMVarImpl(x_1, x_9, x_10, x_4, x_5, x_6, x_7, x_8); -return x_11; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("\nhas unused arguments, reduced type", 35); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_consume___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) { +static lean_object* _init_l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__2___closed__4() { _start: { -lean_object* x_9; uint8_t x_10; -x_9 = lean_st_ref_get(x_5, x_8); -x_10 = !lean_is_exclusive(x_9); -if (x_10 == 0) -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_11 = lean_ctor_get(x_9, 0); -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -lean_dec(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_1); -lean_ctor_set(x_9, 0, x_13); -return x_9; -} -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_9, 0); -x_15 = lean_ctor_get(x_9, 1); -lean_inc(x_15); -lean_inc(x_14); -lean_dec(x_9); -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_1); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_17); -lean_ctor_set(x_18, 1, x_15); -return x_18; -} +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__2___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } -static lean_object* _init_l_Lean_Meta_SynthInstance_consume___closed__1() { +static lean_object* _init_l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__2___closed__5() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_consume___lambda__1___boxed), 8, 0); +x_1 = lean_mk_string_from_bytes("\nTransformer", 12); return x_1; } } -static lean_object* _init_l_Lean_Meta_SynthInstance_consume___closed__2() { +static lean_object* _init_l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__2___closed__6() { _start: { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_consume___lambda__3___boxed), 8, 0); -return x_1; +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__2___closed__5; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_consume(lean_object* x_1, lean_object* x_2, lean_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___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_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) { _start: { -uint8_t x_9; -x_9 = !lean_is_exclusive(x_1); -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; -x_10 = lean_ctor_get(x_1, 0); -x_11 = lean_ctor_get(x_1, 1); -x_12 = lean_ctor_get(x_1, 2); -x_13 = lean_ctor_get(x_1, 3); -x_14 = lean_ctor_get(x_1, 4); -x_15 = lean_box(0); -x_16 = lean_alloc_closure((void*)(l_List_filterAuxM___at_Lean_Meta_SynthInstance_consume___spec__2___boxed), 9, 2); -lean_closure_set(x_16, 0, x_13); -lean_closure_set(x_16, 1, x_15); -x_17 = l_Lean_Meta_SynthInstance_consume___closed__1; -x_18 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_SynthInstance_newSubgoal___spec__7___rarg), 9, 2); -lean_closure_set(x_18, 0, x_16); -lean_closure_set(x_18, 1, x_17); -lean_inc(x_7); -lean_inc(x_6); +lean_object* x_11; uint8_t x_12; uint8_t x_13; uint8_t x_14; lean_object* x_15; lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_12); -x_19 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__18___rarg(x_12, x_18, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_19) == 0) +x_11 = l_Lean_mkAppN(x_5, x_1); +x_12 = 0; +x_13 = 1; +x_14 = 1; +x_15 = l_Lean_Meta_mkLambdaFVars(x_2, x_11, x_12, x_13, x_14, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_15) == 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); -lean_dec(x_19); -lean_inc(x_14); -lean_inc(x_20); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_ctor_set(x_1, 3, x_20); +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, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_18 = l_Lean_Meta_SynthInstance_newSubgoal___lambda__2___closed__3; +x_19 = lean_array_push(x_18, x_5); +x_20 = l_Lean_Meta_mkLambdaFVars(x_19, x_16, x_12, x_13, x_14, x_6, x_7, x_8, x_9, x_17); if (lean_obj_tag(x_20) == 0) { -lean_object* x_22; -lean_dec(x_14); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -x_22 = l_Lean_Meta_SynthInstance_addAnswer(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_21); -return x_22; +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_20, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_20, 1); +lean_inc(x_22); +lean_dec(x_20); +x_23 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__2___closed__2; +x_24 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_23, x_6, x_7, x_8, x_9, x_22); +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_unbox(x_25); +lean_dec(x_25); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +lean_dec(x_4); +x_27 = lean_ctor_get(x_24, 1); +lean_inc(x_27); +lean_dec(x_24); +x_28 = lean_box(0); +x_29 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__1(x_3, x_21, x_28, x_6, x_7, x_8, x_9, x_27); +lean_dec(x_8); +lean_dec(x_6); +return x_29; } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_23 = lean_ctor_get(x_20, 0); -lean_inc(x_23); -lean_inc(x_1); -x_24 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_24, 0, x_1); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); +lean_object* x_30; lean_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_30 = lean_ctor_get(x_24, 1); +lean_inc(x_30); +lean_dec(x_24); +x_31 = l_Lean_MessageData_ofExpr(x_4); +x_32 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; +x_33 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_31); +x_34 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__2___closed__4; +x_35 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_23); -lean_inc(x_12); -x_25 = l_Lean_Meta_SynthInstance_mkTableKeyFor(x_12, x_23, x_2, x_3, x_4, x_5, x_6, x_7, x_21); -if (lean_obj_tag(x_25) == 0) -{ -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -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); -lean_inc(x_26); -x_28 = l_Lean_Meta_SynthInstance_findEntry_x3f(x_26, x_2, x_3, x_4, x_5, x_6, x_7, x_27); -x_29 = lean_ctor_get(x_28, 0); -lean_inc(x_29); -if (lean_obj_tag(x_29) == 0) +x_36 = l_Lean_indentExpr(x_3); +x_37 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_37, 0, x_35); +lean_ctor_set(x_37, 1, x_36); +x_38 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__2___closed__6; +x_39 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +lean_inc(x_21); +x_40 = l_Lean_indentExpr(x_21); +x_41 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_40); +x_42 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_32); +x_43 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_23, x_42, x_6, x_7, x_8, x_9, x_30); +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 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__1(x_3, x_21, x_44, x_6, x_7, x_8, x_9, x_45); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_44); +return x_46; +} +} +else { -lean_object* x_30; lean_object* x_31; -x_30 = lean_ctor_get(x_28, 1); -lean_inc(x_30); -lean_dec(x_28); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_23); -lean_inc(x_12); -x_31 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f(x_12, x_23, x_4, x_5, x_6, x_7, x_30); -if (lean_obj_tag(x_31) == 0) +uint8_t x_47; +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +x_47 = !lean_is_exclusive(x_20); +if (x_47 == 0) { -lean_object* x_32; -x_32 = lean_ctor_get(x_31, 0); -lean_inc(x_32); -if (lean_obj_tag(x_32) == 0) +return x_20; +} +else { -lean_object* x_33; lean_object* x_34; -lean_dec(x_1); +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_20, 0); +x_49 = lean_ctor_get(x_20, 1); +lean_inc(x_49); +lean_inc(x_48); lean_dec(x_20); -lean_dec(x_14); -lean_dec(x_11); -lean_dec(x_10); -x_33 = lean_ctor_get(x_31, 1); -lean_inc(x_33); -lean_dec(x_31); -x_34 = l_Lean_Meta_SynthInstance_newSubgoal(x_12, x_26, x_23, x_24, x_2, x_3, x_4, x_5, x_6, x_7, x_33); -return x_34; +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; +} +} } else { -uint8_t x_35; -lean_dec(x_26); -lean_dec(x_23); -x_35 = !lean_is_exclusive(x_32); -if (x_35 == 0) +uint8_t x_51; +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_51 = !lean_is_exclusive(x_15); +if (x_51 == 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; -x_36 = lean_ctor_get(x_32, 0); -x_37 = lean_ctor_get(x_31, 1); -lean_inc(x_37); -lean_dec(x_31); -x_38 = lean_ctor_get(x_36, 0); -lean_inc(x_38); -x_39 = lean_ctor_get(x_36, 1); -lean_inc(x_39); -lean_dec(x_36); -lean_inc(x_38); -x_40 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_mkTableKey___at_Lean_Meta_SynthInstance_mkTableKeyFor___spec__2___boxed), 8, 1); -lean_closure_set(x_40, 0, x_38); -lean_inc(x_7); -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_41 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__18___rarg(x_12, x_40, x_2, x_3, x_4, x_5, x_6, x_7, x_37); -if (lean_obj_tag(x_41) == 0) +return x_15; +} +else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; -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); -lean_inc(x_42); -x_44 = l_Lean_Meta_SynthInstance_findEntry_x3f(x_42, x_2, x_3, x_4, x_5, x_6, x_7, x_43); -x_45 = lean_ctor_get(x_44, 0); -lean_inc(x_45); -if (lean_obj_tag(x_45) == 0) +lean_object* x_52; lean_object* x_53; lean_object* x_54; +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_15); +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; +} +} +} +} +static lean_object* _init_l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__3___closed__1() { +_start: { -lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; -lean_dec(x_39); -lean_dec(x_24); -lean_dec(x_1); -x_46 = lean_ctor_get(x_44, 1); -lean_inc(x_46); -lean_dec(x_44); -lean_ctor_set(x_32, 0, x_38); -x_47 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_consume___lambda__2___boxed), 8, 1); -lean_closure_set(x_47, 0, x_32); -x_48 = l_Lean_Meta_SynthInstance_consume___closed__2; -x_49 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_SynthInstance_newSubgoal___spec__7___rarg), 9, 2); -lean_closure_set(x_49, 0, x_47); -lean_closure_set(x_49, 1, x_48); +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("redf", 4); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__3___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_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__3___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_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) { +_start: +{ +lean_object* x_9; lean_object* x_10; uint8_t x_11; lean_object* x_12; lean_object* x_13; +x_9 = lean_box(0); +x_10 = lean_array_get_size(x_2); +x_11 = lean_nat_dec_le(x_10, x_10); +if (x_11 == 0) +{ +lean_object* x_33; uint8_t x_34; +x_33 = lean_unsigned_to_nat(0u); +x_34 = lean_nat_dec_lt(x_33, x_10); +if (x_34 == 0) +{ +lean_dec(x_10); +x_12 = x_9; +x_13 = x_8; +goto block_32; +} +else +{ +size_t x_35; size_t x_36; lean_object* x_37; +x_35 = lean_usize_of_nat(x_10); +lean_dec(x_10); +x_36 = 0; 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_50 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__18___rarg(x_12, x_49, x_2, x_3, x_4, x_5, x_6, x_7, x_46); -if (lean_obj_tag(x_50) == 0) +x_37 = l_Array_foldrMUnsafe_fold___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___spec__3(x_3, x_2, x_35, x_36, x_9, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_37) == 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; -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_51, 0); -lean_inc(x_53); -x_54 = lean_ctor_get(x_51, 1); -lean_inc(x_54); -lean_dec(x_51); -lean_inc(x_54); -x_55 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_55, 0, x_54); -lean_ctor_set(x_55, 1, x_20); -lean_inc(x_53); -x_56 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_56, 0, x_10); -lean_ctor_set(x_56, 1, x_11); -lean_ctor_set(x_56, 2, x_53); -lean_ctor_set(x_56, 3, x_55); -lean_ctor_set(x_56, 4, x_14); -x_57 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_57, 0, x_56); -x_58 = l_Lean_Meta_SynthInstance_newSubgoal(x_53, x_42, x_54, x_57, x_2, x_3, x_4, x_5, x_6, x_7, x_52); -return x_58; +lean_object* x_38; lean_object* x_39; +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_12 = x_38; +x_13 = x_39; +goto block_32; } else { -uint8_t x_59; -lean_dec(x_42); -lean_dec(x_20); -lean_dec(x_14); -lean_dec(x_11); -lean_dec(x_10); +uint8_t x_40; 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_59 = !lean_is_exclusive(x_50); -if (x_59 == 0) +lean_dec(x_1); +x_40 = !lean_is_exclusive(x_37); +if (x_40 == 0) { -return x_50; +return x_37; } else { -lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_60 = lean_ctor_get(x_50, 0); -x_61 = lean_ctor_get(x_50, 1); -lean_inc(x_61); -lean_inc(x_60); -lean_dec(x_50); -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_37, 0); +x_42 = lean_ctor_get(x_37, 1); +lean_inc(x_42); +lean_inc(x_41); +lean_dec(x_37); +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_63; lean_object* x_64; uint8_t x_65; -lean_dec(x_38); -lean_free_object(x_32); -lean_dec(x_20); -lean_dec(x_14); -lean_dec(x_11); -lean_dec(x_10); -x_63 = lean_ctor_get(x_45, 0); -lean_inc(x_63); -lean_dec(x_45); -x_64 = lean_ctor_get(x_44, 1); -lean_inc(x_64); -lean_dec(x_44); -x_65 = !lean_is_exclusive(x_63); -if (x_65 == 0) -{ -lean_object* x_66; lean_object* x_67; lean_object* x_68; size_t x_69; size_t x_70; lean_object* x_71; -x_66 = lean_ctor_get(x_63, 0); -x_67 = lean_ctor_get(x_63, 1); -x_68 = lean_array_get_size(x_67); -x_69 = lean_usize_of_nat(x_68); -lean_dec(x_68); -x_70 = 0; -lean_inc(x_3); -lean_inc(x_67); -x_71 = l_Array_mapMUnsafe_map___at_Lean_Meta_SynthInstance_consume___spec__3(x_12, x_39, x_69, x_70, x_67, x_2, x_3, x_4, x_5, x_6, x_7, x_64); -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; uint8_t x_77; -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_st_ref_take(x_3, 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); -lean_dec(x_74); -x_77 = !lean_is_exclusive(x_75); -if (x_77 == 0) -{ -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; -x_78 = lean_ctor_get(x_75, 2); -x_79 = lean_ctor_get(x_75, 3); -x_80 = lean_array_get_size(x_72); -x_81 = lean_unsigned_to_nat(0u); -x_82 = lean_nat_dec_lt(x_81, x_80); -x_83 = lean_array_push(x_66, x_24); -lean_ctor_set(x_63, 0, x_83); -x_84 = l_Lean_HashMap_insert___at_Lean_Meta_SynthInstance_newSubgoal___spec__1(x_79, x_42, x_63); -if (x_82 == 0) -{ -lean_object* x_85; uint8_t x_86; -lean_dec(x_80); -lean_dec(x_72); -lean_dec(x_1); -lean_ctor_set(x_75, 3, x_84); -x_85 = lean_st_ref_set(x_3, x_75, x_76); -lean_dec(x_3); -x_86 = !lean_is_exclusive(x_85); -if (x_86 == 0) +lean_object* x_44; uint8_t x_45; +x_44 = lean_unsigned_to_nat(0u); +x_45 = lean_nat_dec_lt(x_44, x_10); +if (x_45 == 0) { -lean_object* x_87; lean_object* x_88; -x_87 = lean_ctor_get(x_85, 0); -lean_dec(x_87); -x_88 = lean_box(0); -lean_ctor_set(x_85, 0, x_88); -return x_85; +lean_dec(x_10); +x_12 = x_9; +x_13 = x_8; +goto block_32; } else { -lean_object* x_89; lean_object* x_90; lean_object* x_91; -x_89 = lean_ctor_get(x_85, 1); -lean_inc(x_89); -lean_dec(x_85); -x_90 = lean_box(0); -x_91 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_91, 0, x_90); -lean_ctor_set(x_91, 1, x_89); -return x_91; -} +size_t x_46; size_t x_47; lean_object* x_48; +x_46 = lean_usize_of_nat(x_10); +lean_dec(x_10); +x_47 = 0; +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_48 = l_Array_foldrMUnsafe_fold___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___spec__4(x_3, x_2, x_46, x_47, x_9, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_48) == 0) +{ +lean_object* x_49; lean_object* x_50; +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_12 = x_49; +x_13 = x_50; +goto block_32; } else { -uint8_t x_92; -x_92 = lean_nat_dec_le(x_80, x_80); -if (x_92 == 0) -{ -lean_object* x_93; uint8_t x_94; -lean_dec(x_80); -lean_dec(x_72); -lean_dec(x_1); -lean_ctor_set(x_75, 3, x_84); -x_93 = lean_st_ref_set(x_3, x_75, x_76); +uint8_t x_51; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); -x_94 = !lean_is_exclusive(x_93); -if (x_94 == 0) +lean_dec(x_2); +lean_dec(x_1); +x_51 = !lean_is_exclusive(x_48); +if (x_51 == 0) { -lean_object* x_95; lean_object* x_96; -x_95 = lean_ctor_get(x_93, 0); -lean_dec(x_95); -x_96 = lean_box(0); -lean_ctor_set(x_93, 0, x_96); -return x_93; +return x_48; } else { -lean_object* x_97; lean_object* x_98; lean_object* x_99; -x_97 = lean_ctor_get(x_93, 1); -lean_inc(x_97); -lean_dec(x_93); -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; +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_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; } } -else +} +} +block_32: { -size_t x_100; lean_object* x_101; lean_object* x_102; uint8_t x_103; -x_100 = lean_usize_of_nat(x_80); -lean_dec(x_80); -x_101 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_consume___spec__4(x_1, x_72, x_70, x_100, x_78); -lean_dec(x_72); -lean_ctor_set(x_75, 3, x_84); -lean_ctor_set(x_75, 2, x_101); -x_102 = lean_st_ref_set(x_3, x_75, x_76); -lean_dec(x_3); -x_103 = !lean_is_exclusive(x_102); -if (x_103 == 0) +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; uint8_t x_18; uint8_t x_19; lean_object* x_20; +x_14 = l_List_redLength___rarg(x_12); +x_15 = lean_mk_empty_array_with_capacity(x_14); +lean_dec(x_14); +x_16 = l_List_toArrayAux___rarg(x_12, x_15); +x_17 = 0; +x_18 = 1; +x_19 = 1; +lean_inc(x_16); +x_20 = l_Lean_Meta_mkForallFVars(x_16, x_3, x_17, x_18, x_19, x_4, x_5, x_6, x_7, x_13); +if (lean_obj_tag(x_20) == 0) { -lean_object* x_104; lean_object* x_105; -x_104 = lean_ctor_get(x_102, 0); -lean_dec(x_104); -x_105 = lean_box(0); -lean_ctor_set(x_102, 0, x_105); -return x_102; +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; uint8_t x_26; lean_object* x_27; +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_21); +x_23 = lean_alloc_closure((void*)(l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__2___boxed), 10, 4); +lean_closure_set(x_23, 0, x_16); +lean_closure_set(x_23, 1, x_2); +lean_closure_set(x_23, 2, x_21); +lean_closure_set(x_23, 3, x_1); +x_24 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__3___closed__2; +x_25 = 0; +x_26 = 0; +x_27 = l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___spec__2___rarg(x_24, x_25, x_21, x_23, x_26, x_4, x_5, x_6, x_7, x_22); +return x_27; } else { -lean_object* x_106; lean_object* x_107; lean_object* x_108; -x_106 = lean_ctor_get(x_102, 1); -lean_inc(x_106); -lean_dec(x_102); -x_107 = lean_box(0); -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; -} -} -} +uint8_t x_28; +lean_dec(x_16); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_28 = !lean_is_exclusive(x_20); +if (x_28 == 0) +{ +return x_20; } 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; uint8_t x_115; lean_object* x_116; lean_object* x_117; -x_109 = lean_ctor_get(x_75, 0); -x_110 = lean_ctor_get(x_75, 1); -x_111 = lean_ctor_get(x_75, 2); -x_112 = lean_ctor_get(x_75, 3); -lean_inc(x_112); -lean_inc(x_111); -lean_inc(x_110); -lean_inc(x_109); -lean_dec(x_75); -x_113 = lean_array_get_size(x_72); -x_114 = lean_unsigned_to_nat(0u); -x_115 = lean_nat_dec_lt(x_114, x_113); -x_116 = lean_array_push(x_66, x_24); -lean_ctor_set(x_63, 0, x_116); -x_117 = l_Lean_HashMap_insert___at_Lean_Meta_SynthInstance_newSubgoal___spec__1(x_112, x_42, x_63); -if (x_115 == 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_dec(x_113); -lean_dec(x_72); -lean_dec(x_1); -x_118 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_118, 0, x_109); -lean_ctor_set(x_118, 1, x_110); -lean_ctor_set(x_118, 2, x_111); -lean_ctor_set(x_118, 3, x_117); -x_119 = lean_st_ref_set(x_3, x_118, x_76); -lean_dec(x_3); -x_120 = lean_ctor_get(x_119, 1); -lean_inc(x_120); -if (lean_is_exclusive(x_119)) { - lean_ctor_release(x_119, 0); - lean_ctor_release(x_119, 1); - x_121 = x_119; -} else { - lean_dec_ref(x_119); - x_121 = lean_box(0); +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_20, 0); +x_30 = lean_ctor_get(x_20, 1); +lean_inc(x_30); +lean_inc(x_29); +lean_dec(x_20); +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; } -x_122 = lean_box(0); -if (lean_is_scalar(x_121)) { - x_123 = lean_alloc_ctor(0, 2, 0); -} else { - x_123 = x_121; } -lean_ctor_set(x_123, 0, x_122); -lean_ctor_set(x_123, 1, x_120); -return x_123; } -else +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___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) { +_start: { -uint8_t x_124; -x_124 = lean_nat_dec_le(x_113, x_113); -if (x_124 == 0) +lean_object* x_7; +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +x_7 = lean_infer_type(x_1, x_2, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_7) == 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_dec(x_113); -lean_dec(x_72); -lean_dec(x_1); -x_125 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_125, 0, x_109); -lean_ctor_set(x_125, 1, x_110); -lean_ctor_set(x_125, 2, x_111); -lean_ctor_set(x_125, 3, x_117); -x_126 = lean_st_ref_set(x_3, x_125, x_76); +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +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_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_8, x_2, x_3, x_4, x_5, x_9); +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_12 = lean_ctor_get(x_10, 0); +x_13 = lean_ctor_get(x_10, 1); +x_14 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_hasUnusedArguments(x_12); +if (x_14 == 0) +{ +lean_object* x_15; +lean_dec(x_12); +lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); -x_127 = lean_ctor_get(x_126, 1); -lean_inc(x_127); -if (lean_is_exclusive(x_126)) { - lean_ctor_release(x_126, 0); - lean_ctor_release(x_126, 1); - x_128 = x_126; -} else { - lean_dec_ref(x_126); - x_128 = lean_box(0); +lean_dec(x_2); +x_15 = lean_box(0); +lean_ctor_set(x_10, 0, x_15); +return x_10; } -x_129 = lean_box(0); -if (lean_is_scalar(x_128)) { - x_130 = lean_alloc_ctor(0, 2, 0); -} else { - x_130 = x_128; +else +{ +lean_object* x_16; lean_object* x_17; +lean_free_object(x_10); +lean_inc(x_12); +x_16 = lean_alloc_closure((void*)(l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__3), 8, 1); +lean_closure_set(x_16, 0, x_12); +x_17 = l_Lean_Meta_forallTelescope___at___private_Lean_Meta_InferType_0__Lean_Meta_inferForallType___spec__2___rarg(x_12, x_16, x_2, x_3, x_4, x_5, x_13); +return x_17; } -lean_ctor_set(x_130, 0, x_129); -lean_ctor_set(x_130, 1, x_127); -return x_130; } else { -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; -x_131 = lean_usize_of_nat(x_113); -lean_dec(x_113); -x_132 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_consume___spec__4(x_1, x_72, x_70, x_131, x_111); -lean_dec(x_72); -x_133 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_133, 0, x_109); -lean_ctor_set(x_133, 1, x_110); -lean_ctor_set(x_133, 2, x_132); -lean_ctor_set(x_133, 3, x_117); -x_134 = lean_st_ref_set(x_3, x_133, x_76); +lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_18 = lean_ctor_get(x_10, 0); +x_19 = lean_ctor_get(x_10, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_10); +x_20 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_hasUnusedArguments(x_18); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; +lean_dec(x_18); +lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); -x_135 = lean_ctor_get(x_134, 1); -lean_inc(x_135); -if (lean_is_exclusive(x_134)) { - lean_ctor_release(x_134, 0); - lean_ctor_release(x_134, 1); - x_136 = x_134; -} else { - lean_dec_ref(x_134); - x_136 = lean_box(0); -} -x_137 = lean_box(0); -if (lean_is_scalar(x_136)) { - x_138 = lean_alloc_ctor(0, 2, 0); -} else { - x_138 = x_136; -} -lean_ctor_set(x_138, 0, x_137); -lean_ctor_set(x_138, 1, x_135); -return x_138; +lean_dec(x_2); +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_19); +return x_22; } +else +{ +lean_object* x_23; lean_object* x_24; +lean_inc(x_18); +x_23 = lean_alloc_closure((void*)(l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__3), 8, 1); +lean_closure_set(x_23, 0, x_18); +x_24 = l_Lean_Meta_forallTelescope___at___private_Lean_Meta_InferType_0__Lean_Meta_inferForallType___spec__2___rarg(x_18, x_23, x_2, x_3, x_4, x_5, x_19); +return x_24; } } } else { -uint8_t x_139; -lean_free_object(x_63); -lean_dec(x_67); -lean_dec(x_66); -lean_dec(x_42); -lean_dec(x_24); -lean_dec(x_1); +uint8_t x_25; +lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); -x_139 = !lean_is_exclusive(x_71); -if (x_139 == 0) +lean_dec(x_2); +x_25 = !lean_is_exclusive(x_7); +if (x_25 == 0) { -return x_71; +return x_7; } else { -lean_object* x_140; lean_object* x_141; lean_object* x_142; -x_140 = lean_ctor_get(x_71, 0); -x_141 = lean_ctor_get(x_71, 1); -lean_inc(x_141); -lean_inc(x_140); -lean_dec(x_71); -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_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_7, 0); +x_27 = lean_ctor_get(x_7, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_7); +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 +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_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) { +_start: { -lean_object* x_143; lean_object* x_144; lean_object* x_145; size_t x_146; size_t x_147; lean_object* x_148; -x_143 = lean_ctor_get(x_63, 0); -x_144 = lean_ctor_get(x_63, 1); -lean_inc(x_144); -lean_inc(x_143); -lean_dec(x_63); -x_145 = lean_array_get_size(x_144); -x_146 = lean_usize_of_nat(x_145); -lean_dec(x_145); -x_147 = 0; -lean_inc(x_3); -lean_inc(x_144); -x_148 = l_Array_mapMUnsafe_map___at_Lean_Meta_SynthInstance_consume___spec__3(x_12, x_39, x_146, x_147, x_144, x_2, x_3, x_4, x_5, x_6, x_7, x_64); -if (lean_obj_tag(x_148) == 0) +lean_object* x_8; lean_object* x_9; +x_8 = lean_alloc_closure((void*)(l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__4), 6, 1); +lean_closure_set(x_8, 0, x_2); +x_9 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_tryResolve___spec__1___rarg(x_1, x_8, x_3, x_4, x_5, x_6, x_7); +return x_9; +} +} +LEAN_EXPORT lean_object* l_List_anyM___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___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) { +_start: { -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; uint8_t x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; -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_take(x_3, 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_152, 0); -lean_inc(x_154); -x_155 = lean_ctor_get(x_152, 1); -lean_inc(x_155); -x_156 = lean_ctor_get(x_152, 2); -lean_inc(x_156); -x_157 = lean_ctor_get(x_152, 3); -lean_inc(x_157); -if (lean_is_exclusive(x_152)) { - lean_ctor_release(x_152, 0); - lean_ctor_release(x_152, 1); - lean_ctor_release(x_152, 2); - lean_ctor_release(x_152, 3); - x_158 = x_152; -} else { - lean_dec_ref(x_152); - x_158 = lean_box(0); -} -x_159 = lean_array_get_size(x_149); -x_160 = lean_unsigned_to_nat(0u); -x_161 = lean_nat_dec_lt(x_160, x_159); -x_162 = lean_array_push(x_143, x_24); -x_163 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_163, 0, x_162); -lean_ctor_set(x_163, 1, x_144); -x_164 = l_Lean_HashMap_insert___at_Lean_Meta_SynthInstance_newSubgoal___spec__1(x_157, x_42, x_163); -if (x_161 == 0) -{ -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_dec(x_159); -lean_dec(x_149); +lean_object* x_8; +x_8 = l_List_anyM___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); lean_dec(x_1); -if (lean_is_scalar(x_158)) { - x_165 = lean_alloc_ctor(0, 4, 0); -} else { - x_165 = x_158; -} -lean_ctor_set(x_165, 0, x_154); -lean_ctor_set(x_165, 1, x_155); -lean_ctor_set(x_165, 2, x_156); -lean_ctor_set(x_165, 3, x_164); -x_166 = lean_st_ref_set(x_3, x_165, x_153); -lean_dec(x_3); -x_167 = lean_ctor_get(x_166, 1); -lean_inc(x_167); -if (lean_is_exclusive(x_166)) { - lean_ctor_release(x_166, 0); - lean_ctor_release(x_166, 1); - x_168 = x_166; -} else { - lean_dec_ref(x_166); - x_168 = lean_box(0); -} -x_169 = lean_box(0); -if (lean_is_scalar(x_168)) { - x_170 = lean_alloc_ctor(0, 2, 0); -} else { - x_170 = x_168; +return x_8; } -lean_ctor_set(x_170, 0, x_169); -lean_ctor_set(x_170, 1, x_167); -return x_170; } -else -{ -uint8_t x_171; -x_171 = lean_nat_dec_le(x_159, x_159); -if (x_171 == 0) +LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___spec__2___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: { -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_159); -lean_dec(x_149); -lean_dec(x_1); -if (lean_is_scalar(x_158)) { - x_172 = lean_alloc_ctor(0, 4, 0); -} else { - x_172 = x_158; -} -lean_ctor_set(x_172, 0, x_154); -lean_ctor_set(x_172, 1, x_155); -lean_ctor_set(x_172, 2, x_156); -lean_ctor_set(x_172, 3, x_164); -x_173 = lean_st_ref_set(x_3, x_172, x_153); -lean_dec(x_3); -x_174 = lean_ctor_get(x_173, 1); -lean_inc(x_174); -if (lean_is_exclusive(x_173)) { - lean_ctor_release(x_173, 0); - lean_ctor_release(x_173, 1); - x_175 = x_173; -} else { - lean_dec_ref(x_173); - x_175 = lean_box(0); -} -x_176 = lean_box(0); -if (lean_is_scalar(x_175)) { - x_177 = lean_alloc_ctor(0, 2, 0); -} else { - x_177 = x_175; +uint8_t x_11; uint8_t x_12; lean_object* x_13; +x_11 = lean_unbox(x_2); +lean_dec(x_2); +x_12 = lean_unbox(x_5); +lean_dec(x_5); +x_13 = l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___spec__2___rarg(x_1, x_11, x_3, x_4, x_12, x_6, x_7, x_8, x_9, x_10); +return x_13; } -lean_ctor_set(x_177, 0, x_176); -lean_ctor_set(x_177, 1, x_174); -return x_177; } -else +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___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_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_178 = lean_usize_of_nat(x_159); -lean_dec(x_159); -x_179 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_consume___spec__4(x_1, x_149, x_147, x_178, x_156); -lean_dec(x_149); -if (lean_is_scalar(x_158)) { - x_180 = lean_alloc_ctor(0, 4, 0); -} else { - x_180 = x_158; -} -lean_ctor_set(x_180, 0, x_154); -lean_ctor_set(x_180, 1, x_155); -lean_ctor_set(x_180, 2, x_179); -lean_ctor_set(x_180, 3, x_164); -x_181 = lean_st_ref_set(x_3, x_180, x_153); +size_t x_11; size_t x_12; lean_object* x_13; +x_11 = lean_unbox_usize(x_3); lean_dec(x_3); -x_182 = lean_ctor_get(x_181, 1); -lean_inc(x_182); -if (lean_is_exclusive(x_181)) { - lean_ctor_release(x_181, 0); - lean_ctor_release(x_181, 1); - x_183 = x_181; -} else { - lean_dec_ref(x_181); - x_183 = lean_box(0); -} -x_184 = lean_box(0); -if (lean_is_scalar(x_183)) { - x_185 = lean_alloc_ctor(0, 2, 0); -} else { - x_185 = x_183; -} -lean_ctor_set(x_185, 0, x_184); -lean_ctor_set(x_185, 1, x_182); -return x_185; -} +x_12 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_13 = l_Array_foldrMUnsafe_fold___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___spec__3(x_1, x_2, x_11, x_12, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_2); +lean_dec(x_1); +return x_13; } } -else +LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___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: { -lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; -lean_dec(x_144); -lean_dec(x_143); -lean_dec(x_42); -lean_dec(x_24); -lean_dec(x_1); +size_t x_11; size_t x_12; lean_object* x_13; +x_11 = lean_unbox_usize(x_3); lean_dec(x_3); -x_186 = lean_ctor_get(x_148, 0); -lean_inc(x_186); -x_187 = lean_ctor_get(x_148, 1); -lean_inc(x_187); -if (lean_is_exclusive(x_148)) { - lean_ctor_release(x_148, 0); - lean_ctor_release(x_148, 1); - x_188 = x_148; -} else { - lean_dec_ref(x_148); - x_188 = lean_box(0); -} -if (lean_is_scalar(x_188)) { - x_189 = lean_alloc_ctor(1, 2, 0); -} else { - x_189 = x_188; -} -lean_ctor_set(x_189, 0, x_186); -lean_ctor_set(x_189, 1, x_187); -return x_189; -} -} +x_12 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_13 = l_Array_foldrMUnsafe_fold___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___spec__4(x_1, x_2, x_11, x_12, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_2); +lean_dec(x_1); +return x_13; } } -else +LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_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) { +_start: { -uint8_t x_190; -lean_dec(x_39); -lean_dec(x_38); -lean_free_object(x_32); -lean_dec(x_24); -lean_dec(x_1); -lean_dec(x_20); -lean_dec(x_14); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); +lean_object* x_9; +x_9 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__1(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); -x_190 = !lean_is_exclusive(x_41); -if (x_190 == 0) -{ -return x_41; +return x_9; } -else +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_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) { +_start: { -lean_object* x_191; lean_object* x_192; lean_object* x_193; -x_191 = lean_ctor_get(x_41, 0); -x_192 = lean_ctor_get(x_41, 1); -lean_inc(x_192); -lean_inc(x_191); -lean_dec(x_41); -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_object* x_11; +x_11 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___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_9); +lean_dec(x_7); +return x_11; } } +LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at_Lean_Meta_SynthInstance_consume___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_9; uint8_t x_10; +x_9 = lean_st_ref_get(x_5, x_8); +x_10 = !lean_is_exclusive(x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; +x_11 = lean_ctor_get(x_9, 0); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +lean_dec(x_11); +x_13 = lean_ctor_get(x_12, 7); +lean_inc(x_13); +lean_dec(x_12); +x_14 = l_Lean_PersistentHashMap_contains___at_Lean_MVarId_isAssigned___spec__1(x_13, x_1); +x_15 = lean_box(x_14); +lean_ctor_set(x_9, 0, x_15); +return x_9; } 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; -x_194 = lean_ctor_get(x_32, 0); -lean_inc(x_194); -lean_dec(x_32); -x_195 = lean_ctor_get(x_31, 1); -lean_inc(x_195); -lean_dec(x_31); -x_196 = lean_ctor_get(x_194, 0); -lean_inc(x_196); -x_197 = lean_ctor_get(x_194, 1); -lean_inc(x_197); -lean_dec(x_194); -lean_inc(x_196); -x_198 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_mkTableKey___at_Lean_Meta_SynthInstance_mkTableKeyFor___spec__2___boxed), 8, 1); -lean_closure_set(x_198, 0, x_196); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -lean_inc(x_2); +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; lean_object* x_21; lean_object* x_22; +x_16 = lean_ctor_get(x_9, 0); +x_17 = lean_ctor_get(x_9, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_9); +x_18 = lean_ctor_get(x_16, 0); +lean_inc(x_18); +lean_dec(x_16); +x_19 = lean_ctor_get(x_18, 7); +lean_inc(x_19); +lean_dec(x_18); +x_20 = l_Lean_PersistentHashMap_contains___at_Lean_MVarId_isAssigned___spec__1(x_19, x_1); +x_21 = lean_box(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; +} +} +} +LEAN_EXPORT lean_object* l_List_filterAuxM___at_Lean_Meta_SynthInstance_consume___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: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_10; +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_2); +lean_ctor_set(x_10, 1, x_9); +return x_10; +} +else +{ +uint8_t x_11; +x_11 = !lean_is_exclusive(x_1); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_12 = lean_ctor_get(x_1, 0); +x_13 = lean_ctor_get(x_1, 1); lean_inc(x_12); -x_199 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__18___rarg(x_12, x_198, x_2, x_3, x_4, x_5, x_6, x_7, x_195); -if (lean_obj_tag(x_199) == 0) +x_14 = l_Lean_Expr_mvarId_x21(x_12); +x_15 = l_Lean_MVarId_isAssigned___at_Lean_Meta_SynthInstance_consume___spec__1(x_14, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_unbox(x_16); +lean_dec(x_16); +if (x_17 == 0) { -lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; -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); -lean_inc(x_200); -x_202 = l_Lean_Meta_SynthInstance_findEntry_x3f(x_200, x_2, x_3, x_4, x_5, x_6, x_7, x_201); -x_203 = lean_ctor_get(x_202, 0); -lean_inc(x_203); -if (lean_obj_tag(x_203) == 0) +lean_object* x_18; +x_18 = lean_ctor_get(x_15, 1); +lean_inc(x_18); +lean_dec(x_15); +lean_ctor_set(x_1, 1, x_2); { -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_dec(x_197); -lean_dec(x_24); +lean_object* _tmp_0 = x_13; +lean_object* _tmp_1 = x_1; +lean_object* _tmp_8 = x_18; +x_1 = _tmp_0; +x_2 = _tmp_1; +x_9 = _tmp_8; +} +goto _start; +} +else +{ +lean_object* x_20; +lean_free_object(x_1); +lean_dec(x_12); +x_20 = lean_ctor_get(x_15, 1); +lean_inc(x_20); +lean_dec(x_15); +x_1 = x_13; +x_9 = x_20; +goto _start; +} +} +else +{ +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_22 = lean_ctor_get(x_1, 0); +x_23 = lean_ctor_get(x_1, 1); +lean_inc(x_23); +lean_inc(x_22); lean_dec(x_1); -x_204 = lean_ctor_get(x_202, 1); -lean_inc(x_204); -lean_dec(x_202); -x_205 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_205, 0, x_196); -x_206 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_consume___lambda__2___boxed), 8, 1); -lean_closure_set(x_206, 0, x_205); -x_207 = l_Lean_Meta_SynthInstance_consume___closed__2; -x_208 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_SynthInstance_newSubgoal___spec__7___rarg), 9, 2); -lean_closure_set(x_208, 0, x_206); -lean_closure_set(x_208, 1, x_207); -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_209 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__18___rarg(x_12, x_208, x_2, x_3, x_4, x_5, x_6, x_7, x_204); -if (lean_obj_tag(x_209) == 0) +lean_inc(x_22); +x_24 = l_Lean_Expr_mvarId_x21(x_22); +x_25 = l_Lean_MVarId_isAssigned___at_Lean_Meta_SynthInstance_consume___spec__1(x_24, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +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_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_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_210, 0); -lean_inc(x_212); -x_213 = lean_ctor_get(x_210, 1); -lean_inc(x_213); -lean_dec(x_210); -lean_inc(x_213); -x_214 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_214, 0, x_213); -lean_ctor_set(x_214, 1, x_20); -lean_inc(x_212); -x_215 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_215, 0, x_10); -lean_ctor_set(x_215, 1, x_11); -lean_ctor_set(x_215, 2, x_212); -lean_ctor_set(x_215, 3, x_214); -lean_ctor_set(x_215, 4, x_14); -x_216 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_216, 0, x_215); -x_217 = l_Lean_Meta_SynthInstance_newSubgoal(x_212, x_200, x_213, x_216, x_2, x_3, x_4, x_5, x_6, x_7, x_211); -return x_217; +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(1, 2, 0); +lean_ctor_set(x_29, 0, x_22); +lean_ctor_set(x_29, 1, x_2); +x_1 = x_23; +x_2 = x_29; +x_9 = x_28; +goto _start; } else { -lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; -lean_dec(x_200); -lean_dec(x_20); -lean_dec(x_14); -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_218 = lean_ctor_get(x_209, 0); -lean_inc(x_218); -x_219 = lean_ctor_get(x_209, 1); -lean_inc(x_219); -if (lean_is_exclusive(x_209)) { - lean_ctor_release(x_209, 0); - lean_ctor_release(x_209, 1); - x_220 = x_209; -} else { - lean_dec_ref(x_209); - x_220 = lean_box(0); +lean_object* x_31; +lean_dec(x_22); +x_31 = lean_ctor_get(x_25, 1); +lean_inc(x_31); +lean_dec(x_25); +x_1 = x_23; +x_9 = x_31; +goto _start; } -if (lean_is_scalar(x_220)) { - x_221 = lean_alloc_ctor(1, 2, 0); -} else { - x_221 = x_220; } -lean_ctor_set(x_221, 0, x_218); -lean_ctor_set(x_221, 1, x_219); -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; size_t x_228; size_t x_229; lean_object* x_230; -lean_dec(x_196); -lean_dec(x_20); -lean_dec(x_14); -lean_dec(x_11); -lean_dec(x_10); -x_222 = lean_ctor_get(x_203, 0); -lean_inc(x_222); -lean_dec(x_203); -x_223 = lean_ctor_get(x_202, 1); -lean_inc(x_223); -lean_dec(x_202); -x_224 = lean_ctor_get(x_222, 0); -lean_inc(x_224); -x_225 = lean_ctor_get(x_222, 1); -lean_inc(x_225); -if (lean_is_exclusive(x_222)) { - lean_ctor_release(x_222, 0); - lean_ctor_release(x_222, 1); - x_226 = x_222; -} else { - lean_dec_ref(x_222); - x_226 = lean_box(0); } -x_227 = lean_array_get_size(x_225); -x_228 = lean_usize_of_nat(x_227); -lean_dec(x_227); -x_229 = 0; -lean_inc(x_3); -lean_inc(x_225); -x_230 = l_Array_mapMUnsafe_map___at_Lean_Meta_SynthInstance_consume___spec__3(x_12, x_197, x_228, x_229, x_225, x_2, x_3, x_4, x_5, x_6, x_7, x_223); -if (lean_obj_tag(x_230) == 0) +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_SynthInstance_consume___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) { +_start: { -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; uint8_t x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; -x_231 = lean_ctor_get(x_230, 0); -lean_inc(x_231); -x_232 = lean_ctor_get(x_230, 1); -lean_inc(x_232); -lean_dec(x_230); -x_233 = lean_st_ref_take(x_3, x_232); -x_234 = lean_ctor_get(x_233, 0); -lean_inc(x_234); -x_235 = lean_ctor_get(x_233, 1); -lean_inc(x_235); -lean_dec(x_233); -x_236 = lean_ctor_get(x_234, 0); -lean_inc(x_236); -x_237 = lean_ctor_get(x_234, 1); -lean_inc(x_237); -x_238 = lean_ctor_get(x_234, 2); -lean_inc(x_238); -x_239 = lean_ctor_get(x_234, 3); -lean_inc(x_239); -if (lean_is_exclusive(x_234)) { - lean_ctor_release(x_234, 0); - lean_ctor_release(x_234, 1); - lean_ctor_release(x_234, 2); - lean_ctor_release(x_234, 3); - x_240 = x_234; -} else { - lean_dec_ref(x_234); - x_240 = lean_box(0); -} -x_241 = lean_array_get_size(x_231); -x_242 = lean_unsigned_to_nat(0u); -x_243 = lean_nat_dec_lt(x_242, x_241); -x_244 = lean_array_push(x_224, x_24); -if (lean_is_scalar(x_226)) { - x_245 = lean_alloc_ctor(0, 2, 0); -} else { - x_245 = x_226; -} -lean_ctor_set(x_245, 0, x_244); -lean_ctor_set(x_245, 1, x_225); -x_246 = l_Lean_HashMap_insert___at_Lean_Meta_SynthInstance_newSubgoal___spec__1(x_239, x_200, x_245); -if (x_243 == 0) +lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; lean_object* x_16; +x_12 = l_Lean_Meta_SynthInstance_newSubgoal___lambda__2___closed__3; +x_13 = lean_array_push(x_12, x_4); +x_14 = 0; +x_15 = l_Lean_Expr_betaRev(x_1, x_13, x_14, x_14); +lean_dec(x_13); +lean_inc(x_15); +x_16 = lean_infer_type(x_15, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_16) == 0) { -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_dec(x_241); -lean_dec(x_231); -lean_dec(x_1); -if (lean_is_scalar(x_240)) { - x_247 = lean_alloc_ctor(0, 4, 0); -} else { - x_247 = x_240; -} -lean_ctor_set(x_247, 0, x_236); -lean_ctor_set(x_247, 1, x_237); -lean_ctor_set(x_247, 2, x_238); -lean_ctor_set(x_247, 3, x_246); -x_248 = lean_st_ref_set(x_3, x_247, x_235); -lean_dec(x_3); -x_249 = lean_ctor_get(x_248, 1); -lean_inc(x_249); -if (lean_is_exclusive(x_248)) { - lean_ctor_release(x_248, 0); - lean_ctor_release(x_248, 1); - x_250 = x_248; -} else { - lean_dec_ref(x_248); - x_250 = lean_box(0); +uint8_t x_17; +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) +{ +uint8_t x_18; +x_18 = !lean_is_exclusive(x_2); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_19 = lean_ctor_get(x_16, 0); +x_20 = lean_ctor_get(x_2, 2); +lean_dec(x_20); +lean_ctor_set(x_2, 2, x_15); +x_21 = lean_ctor_get(x_3, 2); +lean_inc(x_21); +x_22 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_22, 0, x_2); +lean_ctor_set(x_22, 1, x_19); +lean_ctor_set(x_22, 2, x_21); +lean_ctor_set(x_16, 0, x_22); +return x_16; } -x_251 = lean_box(0); -if (lean_is_scalar(x_250)) { - x_252 = lean_alloc_ctor(0, 2, 0); -} else { - x_252 = x_250; +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; +x_23 = lean_ctor_get(x_16, 0); +x_24 = lean_ctor_get(x_2, 0); +x_25 = lean_ctor_get(x_2, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_2); +x_26 = lean_alloc_ctor(0, 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_15); +x_27 = lean_ctor_get(x_3, 2); +lean_inc(x_27); +x_28 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 1, x_23); +lean_ctor_set(x_28, 2, x_27); +lean_ctor_set(x_16, 0, x_28); +return x_16; } -lean_ctor_set(x_252, 0, x_251); -lean_ctor_set(x_252, 1, x_249); -return x_252; } else { -uint8_t x_253; -x_253 = lean_nat_dec_le(x_241, x_241); -if (x_253 == 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_dec(x_241); -lean_dec(x_231); -lean_dec(x_1); -if (lean_is_scalar(x_240)) { - x_254 = lean_alloc_ctor(0, 4, 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; +x_29 = lean_ctor_get(x_16, 0); +x_30 = lean_ctor_get(x_16, 1); +lean_inc(x_30); +lean_inc(x_29); +lean_dec(x_16); +x_31 = lean_ctor_get(x_2, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_2, 1); +lean_inc(x_32); +if (lean_is_exclusive(x_2)) { + lean_ctor_release(x_2, 0); + lean_ctor_release(x_2, 1); + lean_ctor_release(x_2, 2); + x_33 = x_2; } else { - x_254 = x_240; + lean_dec_ref(x_2); + x_33 = lean_box(0); } -lean_ctor_set(x_254, 0, x_236); -lean_ctor_set(x_254, 1, x_237); -lean_ctor_set(x_254, 2, x_238); -lean_ctor_set(x_254, 3, x_246); -x_255 = lean_st_ref_set(x_3, x_254, x_235); -lean_dec(x_3); -x_256 = lean_ctor_get(x_255, 1); -lean_inc(x_256); -if (lean_is_exclusive(x_255)) { - lean_ctor_release(x_255, 0); - lean_ctor_release(x_255, 1); - x_257 = x_255; +if (lean_is_scalar(x_33)) { + x_34 = lean_alloc_ctor(0, 3, 0); } else { - lean_dec_ref(x_255); - x_257 = lean_box(0); + x_34 = x_33; } -x_258 = lean_box(0); -if (lean_is_scalar(x_257)) { - x_259 = lean_alloc_ctor(0, 2, 0); -} else { - x_259 = x_257; +lean_ctor_set(x_34, 0, x_31); +lean_ctor_set(x_34, 1, x_32); +lean_ctor_set(x_34, 2, x_15); +x_35 = lean_ctor_get(x_3, 2); +lean_inc(x_35); +x_36 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_29); +lean_ctor_set(x_36, 2, x_35); +x_37 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set(x_37, 1, x_30); +return x_37; } -lean_ctor_set(x_259, 0, x_258); -lean_ctor_set(x_259, 1, x_256); -return x_259; } else { -size_t 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_usize_of_nat(x_241); -lean_dec(x_241); -x_261 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_consume___spec__4(x_1, x_231, x_229, x_260, x_238); -lean_dec(x_231); -if (lean_is_scalar(x_240)) { - x_262 = lean_alloc_ctor(0, 4, 0); -} else { - x_262 = x_240; -} -lean_ctor_set(x_262, 0, x_236); -lean_ctor_set(x_262, 1, x_237); -lean_ctor_set(x_262, 2, x_261); -lean_ctor_set(x_262, 3, x_246); -x_263 = lean_st_ref_set(x_3, x_262, x_235); -lean_dec(x_3); -x_264 = lean_ctor_get(x_263, 1); -lean_inc(x_264); -if (lean_is_exclusive(x_263)) { - lean_ctor_release(x_263, 0); - lean_ctor_release(x_263, 1); - x_265 = x_263; -} else { - lean_dec_ref(x_263); - x_265 = lean_box(0); +uint8_t x_38; +lean_dec(x_15); +lean_dec(x_2); +x_38 = !lean_is_exclusive(x_16); +if (x_38 == 0) +{ +return x_16; } -x_266 = lean_box(0); -if (lean_is_scalar(x_265)) { - x_267 = lean_alloc_ctor(0, 2, 0); -} else { - x_267 = x_265; +else +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_16, 0); +x_40 = lean_ctor_get(x_16, 1); +lean_inc(x_40); +lean_inc(x_39); +lean_dec(x_16); +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_ctor_set(x_267, 0, x_266); -lean_ctor_set(x_267, 1, x_264); -return x_267; } } } -else +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_SynthInstance_consume___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) { +_start: { -lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; -lean_dec(x_226); -lean_dec(x_225); -lean_dec(x_224); -lean_dec(x_200); -lean_dec(x_24); -lean_dec(x_1); -lean_dec(x_3); -x_268 = lean_ctor_get(x_230, 0); -lean_inc(x_268); -x_269 = lean_ctor_get(x_230, 1); +uint8_t x_13; +x_13 = lean_usize_dec_lt(x_4, x_3); +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_2); +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; +} +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; +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); +x_18 = lean_ctor_get(x_15, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_18, 2); +lean_inc(x_19); +x_20 = lean_alloc_closure((void*)(l_Lean_instantiateMVars___at_Lean_Meta_SynthInstance_mkTableKeyFor___spec__1___boxed), 8, 1); +lean_closure_set(x_20, 0, x_19); +lean_inc(x_2); +x_21 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at_Lean_Meta_SynthInstance_consume___spec__3___lambda__1___boxed), 11, 3); +lean_closure_set(x_21, 0, x_2); +lean_closure_set(x_21, 1, x_18); +lean_closure_set(x_21, 2, x_15); +x_22 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_SynthInstance_newSubgoal___spec__7___rarg), 9, 2); +lean_closure_set(x_22, 0, x_20); +lean_closure_set(x_22, 1, x_21); +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_1); +x_23 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__16___rarg(x_1, x_22, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; lean_object* x_25; size_t x_26; size_t x_27; lean_object* x_28; +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 = 1; +x_27 = lean_usize_add(x_4, x_26); +x_28 = lean_array_uset(x_17, x_4, x_24); +x_4 = x_27; +x_5 = x_28; +x_12 = x_25; +goto _start; +} +else +{ +uint8_t x_30; +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_2); +lean_dec(x_1); +x_30 = !lean_is_exclusive(x_23); +if (x_30 == 0) +{ +return x_23; +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_23, 0); +x_32 = lean_ctor_get(x_23, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_23); +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_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_consume___spec__4(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; +x_6 = lean_usize_dec_eq(x_3, x_4); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; size_t x_10; size_t x_11; +x_7 = lean_array_uget(x_2, x_3); +lean_inc(x_1); +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_1); +lean_ctor_set(x_8, 1, x_7); +x_9 = lean_array_push(x_5, x_8); +x_10 = 1; +x_11 = lean_usize_add(x_3, x_10); +x_3 = x_11; +x_5 = x_9; +goto _start; +} +else +{ +lean_dec(x_1); +return x_5; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_consume___spec__5(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; +x_6 = lean_usize_dec_eq(x_3, x_4); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; size_t x_10; size_t x_11; +x_7 = lean_array_uget(x_2, x_3); +lean_inc(x_1); +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_1); +lean_ctor_set(x_8, 1, x_7); +x_9 = lean_array_push(x_5, x_8); +x_10 = 1; +x_11 = lean_usize_add(x_3, x_10); +x_3 = x_11; +x_5 = x_9; +goto _start; +} +else +{ +lean_dec(x_1); +return x_5; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_consume___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; +x_9 = l_List_reverse___rarg(x_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); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_consume___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: +{ +uint8_t x_9; lean_object* x_10; lean_object* x_11; +x_9 = 0; +x_10 = lean_box(0); +x_11 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkFreshExprMVarImpl(x_1, x_9, x_10, x_4, x_5, x_6, x_7, x_8); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_consume___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) { +_start: +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_st_ref_get(x_5, x_8); +x_10 = !lean_is_exclusive(x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = lean_ctor_get(x_9, 0); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +lean_dec(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_1); +lean_ctor_set(x_9, 0, x_13); +return x_9; +} +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_9, 0); +x_15 = lean_ctor_get(x_9, 1); +lean_inc(x_15); +lean_inc(x_14); +lean_dec(x_9); +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_1); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_15); +return x_18; +} +} +} +static lean_object* _init_l_Lean_Meta_SynthInstance_consume___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_consume___lambda__1___boxed), 8, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_SynthInstance_consume___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_consume___lambda__3___boxed), 8, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_consume(lean_object* x_1, lean_object* x_2, 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_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_10 = lean_ctor_get(x_1, 0); +x_11 = lean_ctor_get(x_1, 1); +x_12 = lean_ctor_get(x_1, 2); +x_13 = lean_ctor_get(x_1, 3); +x_14 = lean_ctor_get(x_1, 4); +x_15 = lean_box(0); +x_16 = lean_alloc_closure((void*)(l_List_filterAuxM___at_Lean_Meta_SynthInstance_consume___spec__2___boxed), 9, 2); +lean_closure_set(x_16, 0, x_13); +lean_closure_set(x_16, 1, x_15); +x_17 = l_Lean_Meta_SynthInstance_consume___closed__1; +x_18 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_SynthInstance_newSubgoal___spec__7___rarg), 9, 2); +lean_closure_set(x_18, 0, x_16); +lean_closure_set(x_18, 1, x_17); +lean_inc(x_7); +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_19 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__16___rarg(x_12, x_18, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_19) == 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); +lean_dec(x_19); +lean_inc(x_14); +lean_inc(x_20); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_ctor_set(x_1, 3, x_20); +if (lean_obj_tag(x_20) == 0) +{ +lean_object* x_22; +lean_dec(x_14); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +x_22 = l_Lean_Meta_SynthInstance_addAnswer(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_21); +return x_22; +} +else +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_20, 0); +lean_inc(x_23); +lean_inc(x_1); +x_24 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_24, 0, x_1); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_23); +lean_inc(x_12); +x_25 = l_Lean_Meta_SynthInstance_mkTableKeyFor(x_12, x_23, x_2, x_3, x_4, x_5, x_6, x_7, x_21); +if (lean_obj_tag(x_25) == 0) +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +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); +lean_inc(x_26); +x_28 = l_Lean_Meta_SynthInstance_findEntry_x3f(x_26, x_2, x_3, x_4, x_5, x_6, x_7, x_27); +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +if (lean_obj_tag(x_29) == 0) +{ +lean_object* x_30; lean_object* x_31; +x_30 = lean_ctor_get(x_28, 1); +lean_inc(x_30); +lean_dec(x_28); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_23); +lean_inc(x_12); +x_31 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f(x_12, x_23, x_4, x_5, x_6, x_7, x_30); +if (lean_obj_tag(x_31) == 0) +{ +lean_object* x_32; +x_32 = lean_ctor_get(x_31, 0); +lean_inc(x_32); +if (lean_obj_tag(x_32) == 0) +{ +lean_object* x_33; lean_object* x_34; +lean_dec(x_1); +lean_dec(x_20); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_10); +x_33 = lean_ctor_get(x_31, 1); +lean_inc(x_33); +lean_dec(x_31); +x_34 = l_Lean_Meta_SynthInstance_newSubgoal(x_12, x_26, x_23, x_24, x_2, x_3, x_4, x_5, x_6, x_7, x_33); +return x_34; +} +else +{ +uint8_t x_35; +lean_dec(x_26); +lean_dec(x_23); +x_35 = !lean_is_exclusive(x_32); +if (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; +x_36 = lean_ctor_get(x_32, 0); +x_37 = lean_ctor_get(x_31, 1); +lean_inc(x_37); +lean_dec(x_31); +x_38 = lean_ctor_get(x_36, 0); +lean_inc(x_38); +x_39 = lean_ctor_get(x_36, 1); +lean_inc(x_39); +lean_dec(x_36); +lean_inc(x_38); +x_40 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_mkTableKey___at_Lean_Meta_SynthInstance_mkTableKeyFor___spec__2___boxed), 8, 1); +lean_closure_set(x_40, 0, x_38); +lean_inc(x_7); +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_41 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__16___rarg(x_12, x_40, x_2, x_3, x_4, x_5, x_6, x_7, x_37); +if (lean_obj_tag(x_41) == 0) +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +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); +lean_inc(x_42); +x_44 = l_Lean_Meta_SynthInstance_findEntry_x3f(x_42, x_2, x_3, x_4, x_5, x_6, x_7, x_43); +x_45 = lean_ctor_get(x_44, 0); +lean_inc(x_45); +if (lean_obj_tag(x_45) == 0) +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +lean_dec(x_39); +lean_dec(x_24); +lean_dec(x_1); +x_46 = lean_ctor_get(x_44, 1); +lean_inc(x_46); +lean_dec(x_44); +lean_ctor_set(x_32, 0, x_38); +x_47 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_consume___lambda__2___boxed), 8, 1); +lean_closure_set(x_47, 0, x_32); +x_48 = l_Lean_Meta_SynthInstance_consume___closed__2; +x_49 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_SynthInstance_newSubgoal___spec__7___rarg), 9, 2); +lean_closure_set(x_49, 0, x_47); +lean_closure_set(x_49, 1, x_48); +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_50 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__16___rarg(x_12, x_49, x_2, x_3, x_4, x_5, x_6, x_7, x_46); +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; +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_51, 0); +lean_inc(x_53); +x_54 = lean_ctor_get(x_51, 1); +lean_inc(x_54); +lean_dec(x_51); +lean_inc(x_54); +x_55 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_55, 0, x_54); +lean_ctor_set(x_55, 1, x_20); +lean_inc(x_53); +x_56 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_56, 0, x_10); +lean_ctor_set(x_56, 1, x_11); +lean_ctor_set(x_56, 2, x_53); +lean_ctor_set(x_56, 3, x_55); +lean_ctor_set(x_56, 4, x_14); +x_57 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_57, 0, x_56); +x_58 = l_Lean_Meta_SynthInstance_newSubgoal(x_53, x_42, x_54, x_57, x_2, x_3, x_4, x_5, x_6, x_7, x_52); +return x_58; +} +else +{ +uint8_t x_59; +lean_dec(x_42); +lean_dec(x_20); +lean_dec(x_14); +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_59 = !lean_is_exclusive(x_50); +if (x_59 == 0) +{ +return x_50; +} +else +{ +lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_60 = lean_ctor_get(x_50, 0); +x_61 = lean_ctor_get(x_50, 1); +lean_inc(x_61); +lean_inc(x_60); +lean_dec(x_50); +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; +} +} +} +else +{ +lean_object* x_63; lean_object* x_64; uint8_t x_65; +lean_dec(x_38); +lean_free_object(x_32); +lean_dec(x_20); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_10); +x_63 = lean_ctor_get(x_45, 0); +lean_inc(x_63); +lean_dec(x_45); +x_64 = lean_ctor_get(x_44, 1); +lean_inc(x_64); +lean_dec(x_44); +x_65 = !lean_is_exclusive(x_63); +if (x_65 == 0) +{ +lean_object* x_66; lean_object* x_67; lean_object* x_68; size_t x_69; size_t x_70; lean_object* x_71; +x_66 = lean_ctor_get(x_63, 0); +x_67 = lean_ctor_get(x_63, 1); +x_68 = lean_array_get_size(x_67); +x_69 = lean_usize_of_nat(x_68); +lean_dec(x_68); +x_70 = 0; +lean_inc(x_3); +lean_inc(x_67); +x_71 = l_Array_mapMUnsafe_map___at_Lean_Meta_SynthInstance_consume___spec__3(x_12, x_39, x_69, x_70, x_67, x_2, x_3, x_4, x_5, x_6, x_7, x_64); +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; uint8_t x_77; +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_st_ref_take(x_3, 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); +lean_dec(x_74); +x_77 = !lean_is_exclusive(x_75); +if (x_77 == 0) +{ +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; +x_78 = lean_ctor_get(x_75, 2); +x_79 = lean_ctor_get(x_75, 3); +x_80 = lean_array_get_size(x_72); +x_81 = lean_unsigned_to_nat(0u); +x_82 = lean_nat_dec_lt(x_81, x_80); +x_83 = lean_array_push(x_66, x_24); +lean_ctor_set(x_63, 0, x_83); +x_84 = l_Lean_HashMap_insert___at_Lean_Meta_SynthInstance_newSubgoal___spec__1(x_79, x_42, x_63); +if (x_82 == 0) +{ +lean_object* x_85; uint8_t x_86; +lean_dec(x_80); +lean_dec(x_72); +lean_dec(x_1); +lean_ctor_set(x_75, 3, x_84); +x_85 = lean_st_ref_set(x_3, x_75, x_76); +lean_dec(x_3); +x_86 = !lean_is_exclusive(x_85); +if (x_86 == 0) +{ +lean_object* x_87; lean_object* x_88; +x_87 = lean_ctor_get(x_85, 0); +lean_dec(x_87); +x_88 = lean_box(0); +lean_ctor_set(x_85, 0, x_88); +return x_85; +} +else +{ +lean_object* x_89; lean_object* x_90; lean_object* x_91; +x_89 = lean_ctor_get(x_85, 1); +lean_inc(x_89); +lean_dec(x_85); +x_90 = lean_box(0); +x_91 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_91, 0, x_90); +lean_ctor_set(x_91, 1, x_89); +return x_91; +} +} +else +{ +uint8_t x_92; +x_92 = lean_nat_dec_le(x_80, x_80); +if (x_92 == 0) +{ +lean_object* x_93; uint8_t x_94; +lean_dec(x_80); +lean_dec(x_72); +lean_dec(x_1); +lean_ctor_set(x_75, 3, x_84); +x_93 = lean_st_ref_set(x_3, x_75, x_76); +lean_dec(x_3); +x_94 = !lean_is_exclusive(x_93); +if (x_94 == 0) +{ +lean_object* x_95; lean_object* x_96; +x_95 = lean_ctor_get(x_93, 0); +lean_dec(x_95); +x_96 = lean_box(0); +lean_ctor_set(x_93, 0, x_96); +return x_93; +} +else +{ +lean_object* x_97; lean_object* x_98; lean_object* x_99; +x_97 = lean_ctor_get(x_93, 1); +lean_inc(x_97); +lean_dec(x_93); +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; +} +} +else +{ +size_t x_100; lean_object* x_101; lean_object* x_102; uint8_t x_103; +x_100 = lean_usize_of_nat(x_80); +lean_dec(x_80); +x_101 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_consume___spec__4(x_1, x_72, x_70, x_100, x_78); +lean_dec(x_72); +lean_ctor_set(x_75, 3, x_84); +lean_ctor_set(x_75, 2, x_101); +x_102 = lean_st_ref_set(x_3, x_75, x_76); +lean_dec(x_3); +x_103 = !lean_is_exclusive(x_102); +if (x_103 == 0) +{ +lean_object* x_104; lean_object* x_105; +x_104 = lean_ctor_get(x_102, 0); +lean_dec(x_104); +x_105 = lean_box(0); +lean_ctor_set(x_102, 0, x_105); +return x_102; +} +else +{ +lean_object* x_106; lean_object* x_107; lean_object* x_108; +x_106 = lean_ctor_get(x_102, 1); +lean_inc(x_106); +lean_dec(x_102); +x_107 = lean_box(0); +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; +} +} +} +} +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; uint8_t x_115; lean_object* x_116; lean_object* x_117; +x_109 = lean_ctor_get(x_75, 0); +x_110 = lean_ctor_get(x_75, 1); +x_111 = lean_ctor_get(x_75, 2); +x_112 = lean_ctor_get(x_75, 3); +lean_inc(x_112); +lean_inc(x_111); +lean_inc(x_110); +lean_inc(x_109); +lean_dec(x_75); +x_113 = lean_array_get_size(x_72); +x_114 = lean_unsigned_to_nat(0u); +x_115 = lean_nat_dec_lt(x_114, x_113); +x_116 = lean_array_push(x_66, x_24); +lean_ctor_set(x_63, 0, x_116); +x_117 = l_Lean_HashMap_insert___at_Lean_Meta_SynthInstance_newSubgoal___spec__1(x_112, x_42, x_63); +if (x_115 == 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_dec(x_113); +lean_dec(x_72); +lean_dec(x_1); +x_118 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_118, 0, x_109); +lean_ctor_set(x_118, 1, x_110); +lean_ctor_set(x_118, 2, x_111); +lean_ctor_set(x_118, 3, x_117); +x_119 = lean_st_ref_set(x_3, x_118, x_76); +lean_dec(x_3); +x_120 = lean_ctor_get(x_119, 1); +lean_inc(x_120); +if (lean_is_exclusive(x_119)) { + lean_ctor_release(x_119, 0); + lean_ctor_release(x_119, 1); + x_121 = x_119; +} else { + lean_dec_ref(x_119); + x_121 = lean_box(0); +} +x_122 = lean_box(0); +if (lean_is_scalar(x_121)) { + x_123 = lean_alloc_ctor(0, 2, 0); +} else { + x_123 = x_121; +} +lean_ctor_set(x_123, 0, x_122); +lean_ctor_set(x_123, 1, x_120); +return x_123; +} +else +{ +uint8_t x_124; +x_124 = lean_nat_dec_le(x_113, x_113); +if (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_dec(x_113); +lean_dec(x_72); +lean_dec(x_1); +x_125 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_125, 0, x_109); +lean_ctor_set(x_125, 1, x_110); +lean_ctor_set(x_125, 2, x_111); +lean_ctor_set(x_125, 3, x_117); +x_126 = lean_st_ref_set(x_3, x_125, x_76); +lean_dec(x_3); +x_127 = lean_ctor_get(x_126, 1); +lean_inc(x_127); +if (lean_is_exclusive(x_126)) { + lean_ctor_release(x_126, 0); + lean_ctor_release(x_126, 1); + x_128 = x_126; +} else { + lean_dec_ref(x_126); + x_128 = lean_box(0); +} +x_129 = lean_box(0); +if (lean_is_scalar(x_128)) { + x_130 = lean_alloc_ctor(0, 2, 0); +} else { + x_130 = x_128; +} +lean_ctor_set(x_130, 0, x_129); +lean_ctor_set(x_130, 1, x_127); +return x_130; +} +else +{ +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; +x_131 = lean_usize_of_nat(x_113); +lean_dec(x_113); +x_132 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_consume___spec__4(x_1, x_72, x_70, x_131, x_111); +lean_dec(x_72); +x_133 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_133, 0, x_109); +lean_ctor_set(x_133, 1, x_110); +lean_ctor_set(x_133, 2, x_132); +lean_ctor_set(x_133, 3, x_117); +x_134 = lean_st_ref_set(x_3, x_133, x_76); +lean_dec(x_3); +x_135 = lean_ctor_get(x_134, 1); +lean_inc(x_135); +if (lean_is_exclusive(x_134)) { + lean_ctor_release(x_134, 0); + lean_ctor_release(x_134, 1); + x_136 = x_134; +} else { + lean_dec_ref(x_134); + x_136 = lean_box(0); +} +x_137 = lean_box(0); +if (lean_is_scalar(x_136)) { + x_138 = lean_alloc_ctor(0, 2, 0); +} else { + x_138 = x_136; +} +lean_ctor_set(x_138, 0, x_137); +lean_ctor_set(x_138, 1, x_135); +return x_138; +} +} +} +} +else +{ +uint8_t x_139; +lean_free_object(x_63); +lean_dec(x_67); +lean_dec(x_66); +lean_dec(x_42); +lean_dec(x_24); +lean_dec(x_1); +lean_dec(x_3); +x_139 = !lean_is_exclusive(x_71); +if (x_139 == 0) +{ +return x_71; +} +else +{ +lean_object* x_140; lean_object* x_141; lean_object* x_142; +x_140 = lean_ctor_get(x_71, 0); +x_141 = lean_ctor_get(x_71, 1); +lean_inc(x_141); +lean_inc(x_140); +lean_dec(x_71); +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 +{ +lean_object* x_143; lean_object* x_144; lean_object* x_145; size_t x_146; size_t x_147; lean_object* x_148; +x_143 = lean_ctor_get(x_63, 0); +x_144 = lean_ctor_get(x_63, 1); +lean_inc(x_144); +lean_inc(x_143); +lean_dec(x_63); +x_145 = lean_array_get_size(x_144); +x_146 = lean_usize_of_nat(x_145); +lean_dec(x_145); +x_147 = 0; +lean_inc(x_3); +lean_inc(x_144); +x_148 = l_Array_mapMUnsafe_map___at_Lean_Meta_SynthInstance_consume___spec__3(x_12, x_39, x_146, x_147, x_144, x_2, x_3, x_4, x_5, x_6, x_7, x_64); +if (lean_obj_tag(x_148) == 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; uint8_t x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; +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_take(x_3, 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_152, 0); +lean_inc(x_154); +x_155 = lean_ctor_get(x_152, 1); +lean_inc(x_155); +x_156 = lean_ctor_get(x_152, 2); +lean_inc(x_156); +x_157 = lean_ctor_get(x_152, 3); +lean_inc(x_157); +if (lean_is_exclusive(x_152)) { + lean_ctor_release(x_152, 0); + lean_ctor_release(x_152, 1); + lean_ctor_release(x_152, 2); + lean_ctor_release(x_152, 3); + x_158 = x_152; +} else { + lean_dec_ref(x_152); + x_158 = lean_box(0); +} +x_159 = lean_array_get_size(x_149); +x_160 = lean_unsigned_to_nat(0u); +x_161 = lean_nat_dec_lt(x_160, x_159); +x_162 = lean_array_push(x_143, x_24); +x_163 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_163, 0, x_162); +lean_ctor_set(x_163, 1, x_144); +x_164 = l_Lean_HashMap_insert___at_Lean_Meta_SynthInstance_newSubgoal___spec__1(x_157, x_42, x_163); +if (x_161 == 0) +{ +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_dec(x_159); +lean_dec(x_149); +lean_dec(x_1); +if (lean_is_scalar(x_158)) { + x_165 = lean_alloc_ctor(0, 4, 0); +} else { + x_165 = x_158; +} +lean_ctor_set(x_165, 0, x_154); +lean_ctor_set(x_165, 1, x_155); +lean_ctor_set(x_165, 2, x_156); +lean_ctor_set(x_165, 3, x_164); +x_166 = lean_st_ref_set(x_3, x_165, x_153); +lean_dec(x_3); +x_167 = lean_ctor_get(x_166, 1); +lean_inc(x_167); +if (lean_is_exclusive(x_166)) { + lean_ctor_release(x_166, 0); + lean_ctor_release(x_166, 1); + x_168 = x_166; +} else { + lean_dec_ref(x_166); + x_168 = lean_box(0); +} +x_169 = lean_box(0); +if (lean_is_scalar(x_168)) { + x_170 = lean_alloc_ctor(0, 2, 0); +} else { + x_170 = x_168; +} +lean_ctor_set(x_170, 0, x_169); +lean_ctor_set(x_170, 1, x_167); +return x_170; +} +else +{ +uint8_t x_171; +x_171 = lean_nat_dec_le(x_159, x_159); +if (x_171 == 0) +{ +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_159); +lean_dec(x_149); +lean_dec(x_1); +if (lean_is_scalar(x_158)) { + x_172 = lean_alloc_ctor(0, 4, 0); +} else { + x_172 = x_158; +} +lean_ctor_set(x_172, 0, x_154); +lean_ctor_set(x_172, 1, x_155); +lean_ctor_set(x_172, 2, x_156); +lean_ctor_set(x_172, 3, x_164); +x_173 = lean_st_ref_set(x_3, x_172, x_153); +lean_dec(x_3); +x_174 = lean_ctor_get(x_173, 1); +lean_inc(x_174); +if (lean_is_exclusive(x_173)) { + lean_ctor_release(x_173, 0); + lean_ctor_release(x_173, 1); + x_175 = x_173; +} else { + lean_dec_ref(x_173); + x_175 = lean_box(0); +} +x_176 = lean_box(0); +if (lean_is_scalar(x_175)) { + x_177 = lean_alloc_ctor(0, 2, 0); +} else { + x_177 = x_175; +} +lean_ctor_set(x_177, 0, x_176); +lean_ctor_set(x_177, 1, x_174); +return x_177; +} +else +{ +size_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; +x_178 = lean_usize_of_nat(x_159); +lean_dec(x_159); +x_179 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_consume___spec__4(x_1, x_149, x_147, x_178, x_156); +lean_dec(x_149); +if (lean_is_scalar(x_158)) { + x_180 = lean_alloc_ctor(0, 4, 0); +} else { + x_180 = x_158; +} +lean_ctor_set(x_180, 0, x_154); +lean_ctor_set(x_180, 1, x_155); +lean_ctor_set(x_180, 2, x_179); +lean_ctor_set(x_180, 3, x_164); +x_181 = lean_st_ref_set(x_3, x_180, x_153); +lean_dec(x_3); +x_182 = lean_ctor_get(x_181, 1); +lean_inc(x_182); +if (lean_is_exclusive(x_181)) { + lean_ctor_release(x_181, 0); + lean_ctor_release(x_181, 1); + x_183 = x_181; +} else { + lean_dec_ref(x_181); + x_183 = lean_box(0); +} +x_184 = lean_box(0); +if (lean_is_scalar(x_183)) { + x_185 = lean_alloc_ctor(0, 2, 0); +} else { + x_185 = x_183; +} +lean_ctor_set(x_185, 0, x_184); +lean_ctor_set(x_185, 1, x_182); +return x_185; +} +} +} +else +{ +lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; +lean_dec(x_144); +lean_dec(x_143); +lean_dec(x_42); +lean_dec(x_24); +lean_dec(x_1); +lean_dec(x_3); +x_186 = lean_ctor_get(x_148, 0); +lean_inc(x_186); +x_187 = lean_ctor_get(x_148, 1); +lean_inc(x_187); +if (lean_is_exclusive(x_148)) { + lean_ctor_release(x_148, 0); + lean_ctor_release(x_148, 1); + x_188 = x_148; +} else { + lean_dec_ref(x_148); + x_188 = lean_box(0); +} +if (lean_is_scalar(x_188)) { + x_189 = lean_alloc_ctor(1, 2, 0); +} else { + x_189 = x_188; +} +lean_ctor_set(x_189, 0, x_186); +lean_ctor_set(x_189, 1, x_187); +return x_189; +} +} +} +} +else +{ +uint8_t x_190; +lean_dec(x_39); +lean_dec(x_38); +lean_free_object(x_32); +lean_dec(x_24); +lean_dec(x_1); +lean_dec(x_20); +lean_dec(x_14); +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_190 = !lean_is_exclusive(x_41); +if (x_190 == 0) +{ +return x_41; +} +else +{ +lean_object* x_191; lean_object* x_192; lean_object* x_193; +x_191 = lean_ctor_get(x_41, 0); +x_192 = lean_ctor_get(x_41, 1); +lean_inc(x_192); +lean_inc(x_191); +lean_dec(x_41); +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; +} +} +} +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; +x_194 = lean_ctor_get(x_32, 0); +lean_inc(x_194); +lean_dec(x_32); +x_195 = lean_ctor_get(x_31, 1); +lean_inc(x_195); +lean_dec(x_31); +x_196 = lean_ctor_get(x_194, 0); +lean_inc(x_196); +x_197 = lean_ctor_get(x_194, 1); +lean_inc(x_197); +lean_dec(x_194); +lean_inc(x_196); +x_198 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_mkTableKey___at_Lean_Meta_SynthInstance_mkTableKeyFor___spec__2___boxed), 8, 1); +lean_closure_set(x_198, 0, x_196); +lean_inc(x_7); +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_199 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__16___rarg(x_12, x_198, x_2, x_3, x_4, x_5, x_6, x_7, x_195); +if (lean_obj_tag(x_199) == 0) +{ +lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; +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); +lean_inc(x_200); +x_202 = l_Lean_Meta_SynthInstance_findEntry_x3f(x_200, x_2, x_3, x_4, x_5, x_6, x_7, x_201); +x_203 = lean_ctor_get(x_202, 0); +lean_inc(x_203); +if (lean_obj_tag(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_dec(x_197); +lean_dec(x_24); +lean_dec(x_1); +x_204 = lean_ctor_get(x_202, 1); +lean_inc(x_204); +lean_dec(x_202); +x_205 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_205, 0, x_196); +x_206 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_consume___lambda__2___boxed), 8, 1); +lean_closure_set(x_206, 0, x_205); +x_207 = l_Lean_Meta_SynthInstance_consume___closed__2; +x_208 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_SynthInstance_newSubgoal___spec__7___rarg), 9, 2); +lean_closure_set(x_208, 0, x_206); +lean_closure_set(x_208, 1, x_207); +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_209 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__16___rarg(x_12, x_208, x_2, x_3, x_4, x_5, x_6, x_7, x_204); +if (lean_obj_tag(x_209) == 0) +{ +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_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_210, 0); +lean_inc(x_212); +x_213 = lean_ctor_get(x_210, 1); +lean_inc(x_213); +lean_dec(x_210); +lean_inc(x_213); +x_214 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_214, 0, x_213); +lean_ctor_set(x_214, 1, x_20); +lean_inc(x_212); +x_215 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_215, 0, x_10); +lean_ctor_set(x_215, 1, x_11); +lean_ctor_set(x_215, 2, x_212); +lean_ctor_set(x_215, 3, x_214); +lean_ctor_set(x_215, 4, x_14); +x_216 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_216, 0, x_215); +x_217 = l_Lean_Meta_SynthInstance_newSubgoal(x_212, x_200, x_213, x_216, x_2, x_3, x_4, x_5, x_6, x_7, x_211); +return x_217; +} +else +{ +lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; +lean_dec(x_200); +lean_dec(x_20); +lean_dec(x_14); +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_218 = lean_ctor_get(x_209, 0); +lean_inc(x_218); +x_219 = lean_ctor_get(x_209, 1); +lean_inc(x_219); +if (lean_is_exclusive(x_209)) { + lean_ctor_release(x_209, 0); + lean_ctor_release(x_209, 1); + x_220 = x_209; +} else { + lean_dec_ref(x_209); + x_220 = lean_box(0); +} +if (lean_is_scalar(x_220)) { + x_221 = lean_alloc_ctor(1, 2, 0); +} else { + x_221 = x_220; +} +lean_ctor_set(x_221, 0, x_218); +lean_ctor_set(x_221, 1, x_219); +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; size_t x_228; size_t x_229; lean_object* x_230; +lean_dec(x_196); +lean_dec(x_20); +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_10); +x_222 = lean_ctor_get(x_203, 0); +lean_inc(x_222); +lean_dec(x_203); +x_223 = lean_ctor_get(x_202, 1); +lean_inc(x_223); +lean_dec(x_202); +x_224 = lean_ctor_get(x_222, 0); +lean_inc(x_224); +x_225 = lean_ctor_get(x_222, 1); +lean_inc(x_225); +if (lean_is_exclusive(x_222)) { + lean_ctor_release(x_222, 0); + lean_ctor_release(x_222, 1); + x_226 = x_222; +} else { + lean_dec_ref(x_222); + x_226 = lean_box(0); +} +x_227 = lean_array_get_size(x_225); +x_228 = lean_usize_of_nat(x_227); +lean_dec(x_227); +x_229 = 0; +lean_inc(x_3); +lean_inc(x_225); +x_230 = l_Array_mapMUnsafe_map___at_Lean_Meta_SynthInstance_consume___spec__3(x_12, x_197, x_228, x_229, x_225, x_2, x_3, x_4, x_5, x_6, x_7, x_223); +if (lean_obj_tag(x_230) == 0) +{ +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; uint8_t x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; +x_231 = lean_ctor_get(x_230, 0); +lean_inc(x_231); +x_232 = lean_ctor_get(x_230, 1); +lean_inc(x_232); +lean_dec(x_230); +x_233 = lean_st_ref_take(x_3, x_232); +x_234 = lean_ctor_get(x_233, 0); +lean_inc(x_234); +x_235 = lean_ctor_get(x_233, 1); +lean_inc(x_235); +lean_dec(x_233); +x_236 = lean_ctor_get(x_234, 0); +lean_inc(x_236); +x_237 = lean_ctor_get(x_234, 1); +lean_inc(x_237); +x_238 = lean_ctor_get(x_234, 2); +lean_inc(x_238); +x_239 = lean_ctor_get(x_234, 3); +lean_inc(x_239); +if (lean_is_exclusive(x_234)) { + lean_ctor_release(x_234, 0); + lean_ctor_release(x_234, 1); + lean_ctor_release(x_234, 2); + lean_ctor_release(x_234, 3); + x_240 = x_234; +} else { + lean_dec_ref(x_234); + x_240 = lean_box(0); +} +x_241 = lean_array_get_size(x_231); +x_242 = lean_unsigned_to_nat(0u); +x_243 = lean_nat_dec_lt(x_242, x_241); +x_244 = lean_array_push(x_224, x_24); +if (lean_is_scalar(x_226)) { + x_245 = lean_alloc_ctor(0, 2, 0); +} else { + x_245 = x_226; +} +lean_ctor_set(x_245, 0, x_244); +lean_ctor_set(x_245, 1, x_225); +x_246 = l_Lean_HashMap_insert___at_Lean_Meta_SynthInstance_newSubgoal___spec__1(x_239, x_200, x_245); +if (x_243 == 0) +{ +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_dec(x_241); +lean_dec(x_231); +lean_dec(x_1); +if (lean_is_scalar(x_240)) { + x_247 = lean_alloc_ctor(0, 4, 0); +} else { + x_247 = x_240; +} +lean_ctor_set(x_247, 0, x_236); +lean_ctor_set(x_247, 1, x_237); +lean_ctor_set(x_247, 2, x_238); +lean_ctor_set(x_247, 3, x_246); +x_248 = lean_st_ref_set(x_3, x_247, x_235); +lean_dec(x_3); +x_249 = lean_ctor_get(x_248, 1); +lean_inc(x_249); +if (lean_is_exclusive(x_248)) { + lean_ctor_release(x_248, 0); + lean_ctor_release(x_248, 1); + x_250 = x_248; +} else { + lean_dec_ref(x_248); + x_250 = lean_box(0); +} +x_251 = lean_box(0); +if (lean_is_scalar(x_250)) { + x_252 = lean_alloc_ctor(0, 2, 0); +} else { + x_252 = x_250; +} +lean_ctor_set(x_252, 0, x_251); +lean_ctor_set(x_252, 1, x_249); +return x_252; +} +else +{ +uint8_t x_253; +x_253 = lean_nat_dec_le(x_241, x_241); +if (x_253 == 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_dec(x_241); +lean_dec(x_231); +lean_dec(x_1); +if (lean_is_scalar(x_240)) { + x_254 = lean_alloc_ctor(0, 4, 0); +} else { + x_254 = x_240; +} +lean_ctor_set(x_254, 0, x_236); +lean_ctor_set(x_254, 1, x_237); +lean_ctor_set(x_254, 2, x_238); +lean_ctor_set(x_254, 3, x_246); +x_255 = lean_st_ref_set(x_3, x_254, x_235); +lean_dec(x_3); +x_256 = lean_ctor_get(x_255, 1); +lean_inc(x_256); +if (lean_is_exclusive(x_255)) { + lean_ctor_release(x_255, 0); + lean_ctor_release(x_255, 1); + x_257 = x_255; +} else { + lean_dec_ref(x_255); + x_257 = lean_box(0); +} +x_258 = lean_box(0); +if (lean_is_scalar(x_257)) { + x_259 = lean_alloc_ctor(0, 2, 0); +} else { + x_259 = x_257; +} +lean_ctor_set(x_259, 0, x_258); +lean_ctor_set(x_259, 1, x_256); +return x_259; +} +else +{ +size_t 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_usize_of_nat(x_241); +lean_dec(x_241); +x_261 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_consume___spec__4(x_1, x_231, x_229, x_260, x_238); +lean_dec(x_231); +if (lean_is_scalar(x_240)) { + x_262 = lean_alloc_ctor(0, 4, 0); +} else { + x_262 = x_240; +} +lean_ctor_set(x_262, 0, x_236); +lean_ctor_set(x_262, 1, x_237); +lean_ctor_set(x_262, 2, x_261); +lean_ctor_set(x_262, 3, x_246); +x_263 = lean_st_ref_set(x_3, x_262, x_235); +lean_dec(x_3); +x_264 = lean_ctor_get(x_263, 1); +lean_inc(x_264); +if (lean_is_exclusive(x_263)) { + lean_ctor_release(x_263, 0); + lean_ctor_release(x_263, 1); + x_265 = x_263; +} else { + lean_dec_ref(x_263); + x_265 = lean_box(0); +} +x_266 = lean_box(0); +if (lean_is_scalar(x_265)) { + x_267 = lean_alloc_ctor(0, 2, 0); +} else { + x_267 = x_265; +} +lean_ctor_set(x_267, 0, x_266); +lean_ctor_set(x_267, 1, x_264); +return x_267; +} +} +} +else +{ +lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; +lean_dec(x_226); +lean_dec(x_225); +lean_dec(x_224); +lean_dec(x_200); +lean_dec(x_24); +lean_dec(x_1); +lean_dec(x_3); +x_268 = lean_ctor_get(x_230, 0); +lean_inc(x_268); +x_269 = lean_ctor_get(x_230, 1); lean_inc(x_269); if (lean_is_exclusive(x_230)) { lean_ctor_release(x_230, 0); lean_ctor_release(x_230, 1); x_270 = x_230; } else { - lean_dec_ref(x_230); - x_270 = lean_box(0); + lean_dec_ref(x_230); + x_270 = lean_box(0); +} +if (lean_is_scalar(x_270)) { + x_271 = lean_alloc_ctor(1, 2, 0); +} else { + x_271 = x_270; +} +lean_ctor_set(x_271, 0, x_268); +lean_ctor_set(x_271, 1, x_269); +return x_271; +} +} +} +else +{ +lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; +lean_dec(x_197); +lean_dec(x_196); +lean_dec(x_24); +lean_dec(x_1); +lean_dec(x_20); +lean_dec(x_14); +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_272 = lean_ctor_get(x_199, 0); +lean_inc(x_272); +x_273 = lean_ctor_get(x_199, 1); +lean_inc(x_273); +if (lean_is_exclusive(x_199)) { + lean_ctor_release(x_199, 0); + lean_ctor_release(x_199, 1); + x_274 = x_199; +} else { + lean_dec_ref(x_199); + x_274 = lean_box(0); +} +if (lean_is_scalar(x_274)) { + x_275 = lean_alloc_ctor(1, 2, 0); +} else { + x_275 = x_274; +} +lean_ctor_set(x_275, 0, x_272); +lean_ctor_set(x_275, 1, x_273); +return x_275; +} +} +} +} +else +{ +uint8_t x_276; +lean_dec(x_26); +lean_dec(x_24); +lean_dec(x_23); +lean_dec(x_1); +lean_dec(x_20); +lean_dec(x_14); +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_276 = !lean_is_exclusive(x_31); +if (x_276 == 0) +{ +return x_31; +} +else +{ +lean_object* x_277; lean_object* x_278; lean_object* x_279; +x_277 = lean_ctor_get(x_31, 0); +x_278 = lean_ctor_get(x_31, 1); +lean_inc(x_278); +lean_inc(x_277); +lean_dec(x_31); +x_279 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_279, 0, x_277); +lean_ctor_set(x_279, 1, x_278); +return x_279; +} +} +} +else +{ +lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; uint8_t x_285; +lean_dec(x_23); +lean_dec(x_20); +lean_dec(x_14); +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_2); +x_280 = lean_ctor_get(x_28, 1); +lean_inc(x_280); +lean_dec(x_28); +x_281 = lean_ctor_get(x_29, 0); +lean_inc(x_281); +lean_dec(x_29); +x_282 = lean_st_ref_take(x_3, x_280); +x_283 = lean_ctor_get(x_282, 0); +lean_inc(x_283); +x_284 = lean_ctor_get(x_282, 1); +lean_inc(x_284); +lean_dec(x_282); +x_285 = !lean_is_exclusive(x_283); +if (x_285 == 0) +{ +uint8_t x_286; +x_286 = !lean_is_exclusive(x_281); +if (x_286 == 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; uint8_t x_293; lean_object* x_294; lean_object* x_295; +x_287 = lean_ctor_get(x_283, 2); +x_288 = lean_ctor_get(x_283, 3); +x_289 = lean_ctor_get(x_281, 0); +x_290 = lean_ctor_get(x_281, 1); +x_291 = lean_array_get_size(x_290); +x_292 = lean_unsigned_to_nat(0u); +x_293 = lean_nat_dec_lt(x_292, x_291); +x_294 = lean_array_push(x_289, x_24); +lean_inc(x_290); +lean_ctor_set(x_281, 0, x_294); +x_295 = l_Lean_HashMap_insert___at_Lean_Meta_SynthInstance_newSubgoal___spec__1(x_288, x_26, x_281); +if (x_293 == 0) +{ +lean_object* x_296; uint8_t x_297; +lean_dec(x_291); +lean_dec(x_290); +lean_dec(x_1); +lean_ctor_set(x_283, 3, x_295); +x_296 = lean_st_ref_set(x_3, x_283, x_284); +lean_dec(x_3); +x_297 = !lean_is_exclusive(x_296); +if (x_297 == 0) +{ +lean_object* x_298; lean_object* x_299; +x_298 = lean_ctor_get(x_296, 0); +lean_dec(x_298); +x_299 = lean_box(0); +lean_ctor_set(x_296, 0, x_299); +return x_296; +} +else +{ +lean_object* x_300; lean_object* x_301; lean_object* x_302; +x_300 = lean_ctor_get(x_296, 1); +lean_inc(x_300); +lean_dec(x_296); +x_301 = lean_box(0); +x_302 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_302, 0, x_301); +lean_ctor_set(x_302, 1, x_300); +return x_302; +} +} +else +{ +uint8_t x_303; +x_303 = lean_nat_dec_le(x_291, x_291); +if (x_303 == 0) +{ +lean_object* x_304; uint8_t x_305; +lean_dec(x_291); +lean_dec(x_290); +lean_dec(x_1); +lean_ctor_set(x_283, 3, x_295); +x_304 = lean_st_ref_set(x_3, x_283, x_284); +lean_dec(x_3); +x_305 = !lean_is_exclusive(x_304); +if (x_305 == 0) +{ +lean_object* x_306; lean_object* x_307; +x_306 = lean_ctor_get(x_304, 0); +lean_dec(x_306); +x_307 = lean_box(0); +lean_ctor_set(x_304, 0, x_307); +return x_304; +} +else +{ +lean_object* x_308; lean_object* x_309; lean_object* x_310; +x_308 = lean_ctor_get(x_304, 1); +lean_inc(x_308); +lean_dec(x_304); +x_309 = lean_box(0); +x_310 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_310, 0, x_309); +lean_ctor_set(x_310, 1, x_308); +return x_310; +} +} +else +{ +size_t x_311; size_t x_312; lean_object* x_313; lean_object* x_314; uint8_t x_315; +x_311 = 0; +x_312 = lean_usize_of_nat(x_291); +lean_dec(x_291); +x_313 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_consume___spec__5(x_1, x_290, x_311, x_312, x_287); +lean_dec(x_290); +lean_ctor_set(x_283, 3, x_295); +lean_ctor_set(x_283, 2, x_313); +x_314 = lean_st_ref_set(x_3, x_283, x_284); +lean_dec(x_3); +x_315 = !lean_is_exclusive(x_314); +if (x_315 == 0) +{ +lean_object* x_316; lean_object* x_317; +x_316 = lean_ctor_get(x_314, 0); +lean_dec(x_316); +x_317 = lean_box(0); +lean_ctor_set(x_314, 0, x_317); +return x_314; +} +else +{ +lean_object* x_318; lean_object* x_319; lean_object* x_320; +x_318 = lean_ctor_get(x_314, 1); +lean_inc(x_318); +lean_dec(x_314); +x_319 = lean_box(0); +x_320 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_320, 0, x_319); +lean_ctor_set(x_320, 1, x_318); +return x_320; +} +} +} +} +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; uint8_t x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; +x_321 = lean_ctor_get(x_283, 2); +x_322 = lean_ctor_get(x_283, 3); +x_323 = lean_ctor_get(x_281, 0); +x_324 = lean_ctor_get(x_281, 1); +lean_inc(x_324); +lean_inc(x_323); +lean_dec(x_281); +x_325 = lean_array_get_size(x_324); +x_326 = lean_unsigned_to_nat(0u); +x_327 = lean_nat_dec_lt(x_326, x_325); +x_328 = lean_array_push(x_323, x_24); +lean_inc(x_324); +x_329 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_329, 0, x_328); +lean_ctor_set(x_329, 1, x_324); +x_330 = l_Lean_HashMap_insert___at_Lean_Meta_SynthInstance_newSubgoal___spec__1(x_322, x_26, x_329); +if (x_327 == 0) +{ +lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; +lean_dec(x_325); +lean_dec(x_324); +lean_dec(x_1); +lean_ctor_set(x_283, 3, x_330); +x_331 = lean_st_ref_set(x_3, x_283, x_284); +lean_dec(x_3); +x_332 = lean_ctor_get(x_331, 1); +lean_inc(x_332); +if (lean_is_exclusive(x_331)) { + lean_ctor_release(x_331, 0); + lean_ctor_release(x_331, 1); + x_333 = x_331; +} else { + lean_dec_ref(x_331); + x_333 = lean_box(0); +} +x_334 = lean_box(0); +if (lean_is_scalar(x_333)) { + x_335 = lean_alloc_ctor(0, 2, 0); +} else { + x_335 = x_333; +} +lean_ctor_set(x_335, 0, x_334); +lean_ctor_set(x_335, 1, x_332); +return x_335; +} +else +{ +uint8_t x_336; +x_336 = lean_nat_dec_le(x_325, x_325); +if (x_336 == 0) +{ +lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; +lean_dec(x_325); +lean_dec(x_324); +lean_dec(x_1); +lean_ctor_set(x_283, 3, x_330); +x_337 = lean_st_ref_set(x_3, x_283, x_284); +lean_dec(x_3); +x_338 = lean_ctor_get(x_337, 1); +lean_inc(x_338); +if (lean_is_exclusive(x_337)) { + lean_ctor_release(x_337, 0); + lean_ctor_release(x_337, 1); + x_339 = x_337; +} else { + lean_dec_ref(x_337); + x_339 = lean_box(0); +} +x_340 = lean_box(0); +if (lean_is_scalar(x_339)) { + x_341 = lean_alloc_ctor(0, 2, 0); +} else { + x_341 = x_339; +} +lean_ctor_set(x_341, 0, x_340); +lean_ctor_set(x_341, 1, x_338); +return x_341; +} +else +{ +size_t x_342; size_t 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 = 0; +x_343 = lean_usize_of_nat(x_325); +lean_dec(x_325); +x_344 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_consume___spec__5(x_1, x_324, x_342, x_343, x_321); +lean_dec(x_324); +lean_ctor_set(x_283, 3, x_330); +lean_ctor_set(x_283, 2, x_344); +x_345 = lean_st_ref_set(x_3, x_283, x_284); +lean_dec(x_3); +x_346 = lean_ctor_get(x_345, 1); +lean_inc(x_346); +if (lean_is_exclusive(x_345)) { + lean_ctor_release(x_345, 0); + lean_ctor_release(x_345, 1); + x_347 = x_345; +} else { + lean_dec_ref(x_345); + x_347 = lean_box(0); +} +x_348 = lean_box(0); +if (lean_is_scalar(x_347)) { + x_349 = lean_alloc_ctor(0, 2, 0); +} else { + x_349 = x_347; +} +lean_ctor_set(x_349, 0, x_348); +lean_ctor_set(x_349, 1, x_346); +return x_349; +} +} +} +} +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; uint8_t x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; +x_350 = lean_ctor_get(x_283, 0); +x_351 = lean_ctor_get(x_283, 1); +x_352 = lean_ctor_get(x_283, 2); +x_353 = lean_ctor_get(x_283, 3); +lean_inc(x_353); +lean_inc(x_352); +lean_inc(x_351); +lean_inc(x_350); +lean_dec(x_283); +x_354 = lean_ctor_get(x_281, 0); +lean_inc(x_354); +x_355 = lean_ctor_get(x_281, 1); +lean_inc(x_355); +if (lean_is_exclusive(x_281)) { + lean_ctor_release(x_281, 0); + lean_ctor_release(x_281, 1); + x_356 = x_281; +} else { + lean_dec_ref(x_281); + x_356 = lean_box(0); +} +x_357 = lean_array_get_size(x_355); +x_358 = lean_unsigned_to_nat(0u); +x_359 = lean_nat_dec_lt(x_358, x_357); +x_360 = lean_array_push(x_354, x_24); +lean_inc(x_355); +if (lean_is_scalar(x_356)) { + x_361 = lean_alloc_ctor(0, 2, 0); +} else { + x_361 = x_356; +} +lean_ctor_set(x_361, 0, x_360); +lean_ctor_set(x_361, 1, x_355); +x_362 = l_Lean_HashMap_insert___at_Lean_Meta_SynthInstance_newSubgoal___spec__1(x_353, x_26, x_361); +if (x_359 == 0) +{ +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_357); +lean_dec(x_355); +lean_dec(x_1); +x_363 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_363, 0, x_350); +lean_ctor_set(x_363, 1, x_351); +lean_ctor_set(x_363, 2, x_352); +lean_ctor_set(x_363, 3, x_362); +x_364 = lean_st_ref_set(x_3, x_363, x_284); +lean_dec(x_3); +x_365 = lean_ctor_get(x_364, 1); +lean_inc(x_365); +if (lean_is_exclusive(x_364)) { + lean_ctor_release(x_364, 0); + lean_ctor_release(x_364, 1); + x_366 = x_364; +} else { + lean_dec_ref(x_364); + x_366 = lean_box(0); +} +x_367 = lean_box(0); +if (lean_is_scalar(x_366)) { + x_368 = lean_alloc_ctor(0, 2, 0); +} else { + x_368 = x_366; +} +lean_ctor_set(x_368, 0, x_367); +lean_ctor_set(x_368, 1, x_365); +return x_368; +} +else +{ +uint8_t x_369; +x_369 = lean_nat_dec_le(x_357, x_357); +if (x_369 == 0) +{ +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_dec(x_357); +lean_dec(x_355); +lean_dec(x_1); +x_370 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_370, 0, x_350); +lean_ctor_set(x_370, 1, x_351); +lean_ctor_set(x_370, 2, x_352); +lean_ctor_set(x_370, 3, x_362); +x_371 = lean_st_ref_set(x_3, x_370, x_284); +lean_dec(x_3); +x_372 = lean_ctor_get(x_371, 1); +lean_inc(x_372); +if (lean_is_exclusive(x_371)) { + lean_ctor_release(x_371, 0); + lean_ctor_release(x_371, 1); + x_373 = x_371; +} else { + lean_dec_ref(x_371); + x_373 = lean_box(0); +} +x_374 = lean_box(0); +if (lean_is_scalar(x_373)) { + x_375 = lean_alloc_ctor(0, 2, 0); +} else { + x_375 = x_373; +} +lean_ctor_set(x_375, 0, x_374); +lean_ctor_set(x_375, 1, x_372); +return x_375; +} +else +{ +size_t x_376; size_t 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; +x_376 = 0; +x_377 = lean_usize_of_nat(x_357); +lean_dec(x_357); +x_378 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_consume___spec__5(x_1, x_355, x_376, x_377, x_352); +lean_dec(x_355); +x_379 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_379, 0, x_350); +lean_ctor_set(x_379, 1, x_351); +lean_ctor_set(x_379, 2, x_378); +lean_ctor_set(x_379, 3, x_362); +x_380 = lean_st_ref_set(x_3, x_379, x_284); +lean_dec(x_3); +x_381 = lean_ctor_get(x_380, 1); +lean_inc(x_381); +if (lean_is_exclusive(x_380)) { + lean_ctor_release(x_380, 0); + lean_ctor_release(x_380, 1); + x_382 = x_380; +} else { + lean_dec_ref(x_380); + x_382 = lean_box(0); +} +x_383 = lean_box(0); +if (lean_is_scalar(x_382)) { + x_384 = lean_alloc_ctor(0, 2, 0); +} else { + x_384 = x_382; +} +lean_ctor_set(x_384, 0, x_383); +lean_ctor_set(x_384, 1, x_381); +return x_384; +} +} +} +} +} +else +{ +uint8_t x_385; +lean_dec(x_24); +lean_dec(x_23); +lean_dec(x_1); +lean_dec(x_20); +lean_dec(x_14); +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_385 = !lean_is_exclusive(x_25); +if (x_385 == 0) +{ +return x_25; +} +else +{ +lean_object* x_386; lean_object* x_387; lean_object* x_388; +x_386 = lean_ctor_get(x_25, 0); +x_387 = lean_ctor_get(x_25, 1); +lean_inc(x_387); +lean_inc(x_386); +lean_dec(x_25); +x_388 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_388, 0, x_386); +lean_ctor_set(x_388, 1, x_387); +return x_388; } -if (lean_is_scalar(x_270)) { - x_271 = lean_alloc_ctor(1, 2, 0); -} else { - x_271 = x_270; } -lean_ctor_set(x_271, 0, x_268); -lean_ctor_set(x_271, 1, x_269); -return x_271; } } +else +{ +uint8_t x_389; +lean_free_object(x_1); +lean_dec(x_14); +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_389 = !lean_is_exclusive(x_19); +if (x_389 == 0) +{ +return x_19; +} +else +{ +lean_object* x_390; lean_object* x_391; lean_object* x_392; +x_390 = lean_ctor_get(x_19, 0); +x_391 = lean_ctor_get(x_19, 1); +lean_inc(x_391); +lean_inc(x_390); +lean_dec(x_19); +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; +} +} +} +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; +x_393 = lean_ctor_get(x_1, 0); +x_394 = lean_ctor_get(x_1, 1); +x_395 = lean_ctor_get(x_1, 2); +x_396 = lean_ctor_get(x_1, 3); +x_397 = lean_ctor_get(x_1, 4); +lean_inc(x_397); +lean_inc(x_396); +lean_inc(x_395); +lean_inc(x_394); +lean_inc(x_393); +lean_dec(x_1); +x_398 = lean_box(0); +x_399 = lean_alloc_closure((void*)(l_List_filterAuxM___at_Lean_Meta_SynthInstance_consume___spec__2___boxed), 9, 2); +lean_closure_set(x_399, 0, x_396); +lean_closure_set(x_399, 1, x_398); +x_400 = l_Lean_Meta_SynthInstance_consume___closed__1; +x_401 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_SynthInstance_newSubgoal___spec__7___rarg), 9, 2); +lean_closure_set(x_401, 0, x_399); +lean_closure_set(x_401, 1, x_400); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_395); +x_402 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__16___rarg(x_395, x_401, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_402) == 0) +{ +lean_object* x_403; lean_object* x_404; lean_object* x_405; +x_403 = lean_ctor_get(x_402, 0); +lean_inc(x_403); +x_404 = lean_ctor_get(x_402, 1); +lean_inc(x_404); +lean_dec(x_402); +lean_inc(x_397); +lean_inc(x_403); +lean_inc(x_395); +lean_inc(x_394); +lean_inc(x_393); +x_405 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_405, 0, x_393); +lean_ctor_set(x_405, 1, x_394); +lean_ctor_set(x_405, 2, x_395); +lean_ctor_set(x_405, 3, x_403); +lean_ctor_set(x_405, 4, x_397); +if (lean_obj_tag(x_403) == 0) +{ +lean_object* x_406; +lean_dec(x_397); +lean_dec(x_395); +lean_dec(x_394); +lean_dec(x_393); +x_406 = l_Lean_Meta_SynthInstance_addAnswer(x_405, x_2, x_3, x_4, x_5, x_6, x_7, x_404); +return x_406; +} +else +{ +lean_object* x_407; lean_object* x_408; lean_object* x_409; +x_407 = lean_ctor_get(x_403, 0); +lean_inc(x_407); +lean_inc(x_405); +x_408 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_408, 0, x_405); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_407); +lean_inc(x_395); +x_409 = l_Lean_Meta_SynthInstance_mkTableKeyFor(x_395, x_407, x_2, x_3, x_4, x_5, x_6, x_7, x_404); +if (lean_obj_tag(x_409) == 0) +{ +lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; +x_410 = lean_ctor_get(x_409, 0); +lean_inc(x_410); +x_411 = lean_ctor_get(x_409, 1); +lean_inc(x_411); +lean_dec(x_409); +lean_inc(x_410); +x_412 = l_Lean_Meta_SynthInstance_findEntry_x3f(x_410, x_2, x_3, x_4, x_5, x_6, x_7, x_411); +x_413 = lean_ctor_get(x_412, 0); +lean_inc(x_413); +if (lean_obj_tag(x_413) == 0) +{ +lean_object* x_414; lean_object* x_415; +x_414 = lean_ctor_get(x_412, 1); +lean_inc(x_414); +lean_dec(x_412); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_407); +lean_inc(x_395); +x_415 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f(x_395, x_407, x_4, x_5, x_6, x_7, x_414); +if (lean_obj_tag(x_415) == 0) +{ +lean_object* x_416; +x_416 = lean_ctor_get(x_415, 0); +lean_inc(x_416); +if (lean_obj_tag(x_416) == 0) +{ +lean_object* x_417; lean_object* x_418; +lean_dec(x_405); +lean_dec(x_403); +lean_dec(x_397); +lean_dec(x_394); +lean_dec(x_393); +x_417 = lean_ctor_get(x_415, 1); +lean_inc(x_417); +lean_dec(x_415); +x_418 = l_Lean_Meta_SynthInstance_newSubgoal(x_395, x_410, x_407, x_408, x_2, x_3, x_4, x_5, x_6, x_7, x_417); +return x_418; +} +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_dec(x_410); +lean_dec(x_407); +x_419 = lean_ctor_get(x_416, 0); +lean_inc(x_419); +if (lean_is_exclusive(x_416)) { + lean_ctor_release(x_416, 0); + x_420 = x_416; +} else { + lean_dec_ref(x_416); + x_420 = lean_box(0); +} +x_421 = lean_ctor_get(x_415, 1); +lean_inc(x_421); +lean_dec(x_415); +x_422 = lean_ctor_get(x_419, 0); +lean_inc(x_422); +x_423 = lean_ctor_get(x_419, 1); +lean_inc(x_423); +lean_dec(x_419); +lean_inc(x_422); +x_424 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_mkTableKey___at_Lean_Meta_SynthInstance_mkTableKeyFor___spec__2___boxed), 8, 1); +lean_closure_set(x_424, 0, x_422); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_395); +x_425 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__16___rarg(x_395, x_424, x_2, x_3, x_4, x_5, x_6, x_7, x_421); +if (lean_obj_tag(x_425) == 0) +{ +lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; +x_426 = lean_ctor_get(x_425, 0); +lean_inc(x_426); +x_427 = lean_ctor_get(x_425, 1); +lean_inc(x_427); +lean_dec(x_425); +lean_inc(x_426); +x_428 = l_Lean_Meta_SynthInstance_findEntry_x3f(x_426, x_2, x_3, x_4, x_5, x_6, x_7, x_427); +x_429 = lean_ctor_get(x_428, 0); +lean_inc(x_429); +if (lean_obj_tag(x_429) == 0) +{ +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_dec(x_423); +lean_dec(x_408); +lean_dec(x_405); +x_430 = lean_ctor_get(x_428, 1); +lean_inc(x_430); +lean_dec(x_428); +if (lean_is_scalar(x_420)) { + x_431 = lean_alloc_ctor(1, 1, 0); +} else { + x_431 = x_420; +} +lean_ctor_set(x_431, 0, x_422); +x_432 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_consume___lambda__2___boxed), 8, 1); +lean_closure_set(x_432, 0, x_431); +x_433 = l_Lean_Meta_SynthInstance_consume___closed__2; +x_434 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_SynthInstance_newSubgoal___spec__7___rarg), 9, 2); +lean_closure_set(x_434, 0, x_432); +lean_closure_set(x_434, 1, x_433); +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_435 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__16___rarg(x_395, x_434, x_2, x_3, x_4, x_5, x_6, x_7, x_430); +if (lean_obj_tag(x_435) == 0) +{ +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; +x_436 = lean_ctor_get(x_435, 0); +lean_inc(x_436); +x_437 = lean_ctor_get(x_435, 1); +lean_inc(x_437); +lean_dec(x_435); +x_438 = lean_ctor_get(x_436, 0); +lean_inc(x_438); +x_439 = lean_ctor_get(x_436, 1); +lean_inc(x_439); +lean_dec(x_436); +lean_inc(x_439); +x_440 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_440, 0, x_439); +lean_ctor_set(x_440, 1, x_403); +lean_inc(x_438); +x_441 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_441, 0, x_393); +lean_ctor_set(x_441, 1, x_394); +lean_ctor_set(x_441, 2, x_438); +lean_ctor_set(x_441, 3, x_440); +lean_ctor_set(x_441, 4, x_397); +x_442 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_442, 0, x_441); +x_443 = l_Lean_Meta_SynthInstance_newSubgoal(x_438, x_426, x_439, x_442, x_2, x_3, x_4, x_5, x_6, x_7, x_437); +return x_443; } else { -lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; -lean_dec(x_197); -lean_dec(x_196); -lean_dec(x_24); -lean_dec(x_1); -lean_dec(x_20); -lean_dec(x_14); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); +lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; +lean_dec(x_426); +lean_dec(x_403); +lean_dec(x_397); +lean_dec(x_394); +lean_dec(x_393); 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_272 = lean_ctor_get(x_199, 0); -lean_inc(x_272); -x_273 = lean_ctor_get(x_199, 1); -lean_inc(x_273); -if (lean_is_exclusive(x_199)) { - lean_ctor_release(x_199, 0); - lean_ctor_release(x_199, 1); - x_274 = x_199; +x_444 = lean_ctor_get(x_435, 0); +lean_inc(x_444); +x_445 = lean_ctor_get(x_435, 1); +lean_inc(x_445); +if (lean_is_exclusive(x_435)) { + lean_ctor_release(x_435, 0); + lean_ctor_release(x_435, 1); + x_446 = x_435; } else { - lean_dec_ref(x_199); - x_274 = lean_box(0); + lean_dec_ref(x_435); + x_446 = lean_box(0); } -if (lean_is_scalar(x_274)) { - x_275 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_446)) { + x_447 = lean_alloc_ctor(1, 2, 0); } else { - x_275 = x_274; -} -lean_ctor_set(x_275, 0, x_272); -lean_ctor_set(x_275, 1, x_273); -return x_275; -} -} -} -} -else -{ -uint8_t x_276; -lean_dec(x_26); -lean_dec(x_24); -lean_dec(x_23); -lean_dec(x_1); -lean_dec(x_20); -lean_dec(x_14); -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_276 = !lean_is_exclusive(x_31); -if (x_276 == 0) -{ -return x_31; -} -else -{ -lean_object* x_277; lean_object* x_278; lean_object* x_279; -x_277 = lean_ctor_get(x_31, 0); -x_278 = lean_ctor_get(x_31, 1); -lean_inc(x_278); -lean_inc(x_277); -lean_dec(x_31); -x_279 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_279, 0, x_277); -lean_ctor_set(x_279, 1, x_278); -return x_279; -} -} -} -else -{ -lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; uint8_t x_285; -lean_dec(x_23); -lean_dec(x_20); -lean_dec(x_14); -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_2); -x_280 = lean_ctor_get(x_28, 1); -lean_inc(x_280); -lean_dec(x_28); -x_281 = lean_ctor_get(x_29, 0); -lean_inc(x_281); -lean_dec(x_29); -x_282 = lean_st_ref_take(x_3, x_280); -x_283 = lean_ctor_get(x_282, 0); -lean_inc(x_283); -x_284 = lean_ctor_get(x_282, 1); -lean_inc(x_284); -lean_dec(x_282); -x_285 = !lean_is_exclusive(x_283); -if (x_285 == 0) -{ -uint8_t x_286; -x_286 = !lean_is_exclusive(x_281); -if (x_286 == 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; uint8_t x_293; lean_object* x_294; lean_object* x_295; -x_287 = lean_ctor_get(x_283, 2); -x_288 = lean_ctor_get(x_283, 3); -x_289 = lean_ctor_get(x_281, 0); -x_290 = lean_ctor_get(x_281, 1); -x_291 = lean_array_get_size(x_290); -x_292 = lean_unsigned_to_nat(0u); -x_293 = lean_nat_dec_lt(x_292, x_291); -x_294 = lean_array_push(x_289, x_24); -lean_inc(x_290); -lean_ctor_set(x_281, 0, x_294); -x_295 = l_Lean_HashMap_insert___at_Lean_Meta_SynthInstance_newSubgoal___spec__1(x_288, x_26, x_281); -if (x_293 == 0) -{ -lean_object* x_296; uint8_t x_297; -lean_dec(x_291); -lean_dec(x_290); -lean_dec(x_1); -lean_ctor_set(x_283, 3, x_295); -x_296 = lean_st_ref_set(x_3, x_283, x_284); -lean_dec(x_3); -x_297 = !lean_is_exclusive(x_296); -if (x_297 == 0) -{ -lean_object* x_298; lean_object* x_299; -x_298 = lean_ctor_get(x_296, 0); -lean_dec(x_298); -x_299 = lean_box(0); -lean_ctor_set(x_296, 0, x_299); -return x_296; + x_447 = x_446; } -else -{ -lean_object* x_300; lean_object* x_301; lean_object* x_302; -x_300 = lean_ctor_get(x_296, 1); -lean_inc(x_300); -lean_dec(x_296); -x_301 = lean_box(0); -x_302 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_302, 0, x_301); -lean_ctor_set(x_302, 1, x_300); -return x_302; +lean_ctor_set(x_447, 0, x_444); +lean_ctor_set(x_447, 1, x_445); +return x_447; } } else { -uint8_t x_303; -x_303 = lean_nat_dec_le(x_291, x_291); -if (x_303 == 0) -{ -lean_object* x_304; uint8_t x_305; -lean_dec(x_291); -lean_dec(x_290); -lean_dec(x_1); -lean_ctor_set(x_283, 3, x_295); -x_304 = lean_st_ref_set(x_3, x_283, x_284); -lean_dec(x_3); -x_305 = !lean_is_exclusive(x_304); -if (x_305 == 0) -{ -lean_object* x_306; lean_object* x_307; -x_306 = lean_ctor_get(x_304, 0); -lean_dec(x_306); -x_307 = lean_box(0); -lean_ctor_set(x_304, 0, x_307); -return x_304; +lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; size_t x_454; size_t x_455; lean_object* x_456; +lean_dec(x_422); +lean_dec(x_420); +lean_dec(x_403); +lean_dec(x_397); +lean_dec(x_394); +lean_dec(x_393); +x_448 = lean_ctor_get(x_429, 0); +lean_inc(x_448); +lean_dec(x_429); +x_449 = lean_ctor_get(x_428, 1); +lean_inc(x_449); +lean_dec(x_428); +x_450 = lean_ctor_get(x_448, 0); +lean_inc(x_450); +x_451 = lean_ctor_get(x_448, 1); +lean_inc(x_451); +if (lean_is_exclusive(x_448)) { + lean_ctor_release(x_448, 0); + lean_ctor_release(x_448, 1); + x_452 = x_448; +} else { + lean_dec_ref(x_448); + x_452 = lean_box(0); } -else +x_453 = lean_array_get_size(x_451); +x_454 = lean_usize_of_nat(x_453); +lean_dec(x_453); +x_455 = 0; +lean_inc(x_3); +lean_inc(x_451); +x_456 = l_Array_mapMUnsafe_map___at_Lean_Meta_SynthInstance_consume___spec__3(x_395, x_423, x_454, x_455, x_451, x_2, x_3, x_4, x_5, x_6, x_7, x_449); +if (lean_obj_tag(x_456) == 0) { -lean_object* x_308; lean_object* x_309; lean_object* x_310; -x_308 = lean_ctor_get(x_304, 1); -lean_inc(x_308); -lean_dec(x_304); -x_309 = lean_box(0); -x_310 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_310, 0, x_309); -lean_ctor_set(x_310, 1, x_308); -return x_310; -} +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; uint8_t x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; +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_st_ref_take(x_3, x_458); +x_460 = lean_ctor_get(x_459, 0); +lean_inc(x_460); +x_461 = lean_ctor_get(x_459, 1); +lean_inc(x_461); +lean_dec(x_459); +x_462 = lean_ctor_get(x_460, 0); +lean_inc(x_462); +x_463 = lean_ctor_get(x_460, 1); +lean_inc(x_463); +x_464 = lean_ctor_get(x_460, 2); +lean_inc(x_464); +x_465 = lean_ctor_get(x_460, 3); +lean_inc(x_465); +if (lean_is_exclusive(x_460)) { + lean_ctor_release(x_460, 0); + lean_ctor_release(x_460, 1); + lean_ctor_release(x_460, 2); + lean_ctor_release(x_460, 3); + x_466 = x_460; +} else { + lean_dec_ref(x_460); + x_466 = lean_box(0); } -else -{ -size_t x_311; size_t x_312; lean_object* x_313; lean_object* x_314; uint8_t x_315; -x_311 = 0; -x_312 = lean_usize_of_nat(x_291); -lean_dec(x_291); -x_313 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_consume___spec__5(x_1, x_290, x_311, x_312, x_287); -lean_dec(x_290); -lean_ctor_set(x_283, 3, x_295); -lean_ctor_set(x_283, 2, x_313); -x_314 = lean_st_ref_set(x_3, x_283, x_284); -lean_dec(x_3); -x_315 = !lean_is_exclusive(x_314); -if (x_315 == 0) -{ -lean_object* x_316; lean_object* x_317; -x_316 = lean_ctor_get(x_314, 0); -lean_dec(x_316); -x_317 = lean_box(0); -lean_ctor_set(x_314, 0, x_317); -return x_314; +x_467 = lean_array_get_size(x_457); +x_468 = lean_unsigned_to_nat(0u); +x_469 = lean_nat_dec_lt(x_468, x_467); +x_470 = lean_array_push(x_450, x_408); +if (lean_is_scalar(x_452)) { + x_471 = lean_alloc_ctor(0, 2, 0); +} else { + x_471 = x_452; } -else +lean_ctor_set(x_471, 0, x_470); +lean_ctor_set(x_471, 1, x_451); +x_472 = l_Lean_HashMap_insert___at_Lean_Meta_SynthInstance_newSubgoal___spec__1(x_465, x_426, x_471); +if (x_469 == 0) { -lean_object* x_318; lean_object* x_319; lean_object* x_320; -x_318 = lean_ctor_get(x_314, 1); -lean_inc(x_318); -lean_dec(x_314); -x_319 = lean_box(0); -x_320 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_320, 0, x_319); -lean_ctor_set(x_320, 1, x_318); -return x_320; +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_dec(x_467); +lean_dec(x_457); +lean_dec(x_405); +if (lean_is_scalar(x_466)) { + x_473 = lean_alloc_ctor(0, 4, 0); +} else { + x_473 = x_466; } +lean_ctor_set(x_473, 0, x_462); +lean_ctor_set(x_473, 1, x_463); +lean_ctor_set(x_473, 2, x_464); +lean_ctor_set(x_473, 3, x_472); +x_474 = lean_st_ref_set(x_3, x_473, x_461); +lean_dec(x_3); +x_475 = lean_ctor_get(x_474, 1); +lean_inc(x_475); +if (lean_is_exclusive(x_474)) { + lean_ctor_release(x_474, 0); + lean_ctor_release(x_474, 1); + x_476 = x_474; +} else { + lean_dec_ref(x_474); + x_476 = lean_box(0); } +x_477 = lean_box(0); +if (lean_is_scalar(x_476)) { + x_478 = lean_alloc_ctor(0, 2, 0); +} else { + x_478 = x_476; } +lean_ctor_set(x_478, 0, x_477); +lean_ctor_set(x_478, 1, x_475); +return x_478; } 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; uint8_t x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; -x_321 = lean_ctor_get(x_283, 2); -x_322 = lean_ctor_get(x_283, 3); -x_323 = lean_ctor_get(x_281, 0); -x_324 = lean_ctor_get(x_281, 1); -lean_inc(x_324); -lean_inc(x_323); -lean_dec(x_281); -x_325 = lean_array_get_size(x_324); -x_326 = lean_unsigned_to_nat(0u); -x_327 = lean_nat_dec_lt(x_326, x_325); -x_328 = lean_array_push(x_323, x_24); -lean_inc(x_324); -x_329 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_329, 0, x_328); -lean_ctor_set(x_329, 1, x_324); -x_330 = l_Lean_HashMap_insert___at_Lean_Meta_SynthInstance_newSubgoal___spec__1(x_322, x_26, x_329); -if (x_327 == 0) +uint8_t x_479; +x_479 = lean_nat_dec_le(x_467, x_467); +if (x_479 == 0) { -lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; -lean_dec(x_325); -lean_dec(x_324); -lean_dec(x_1); -lean_ctor_set(x_283, 3, x_330); -x_331 = lean_st_ref_set(x_3, x_283, x_284); +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_dec(x_467); +lean_dec(x_457); +lean_dec(x_405); +if (lean_is_scalar(x_466)) { + x_480 = lean_alloc_ctor(0, 4, 0); +} else { + x_480 = x_466; +} +lean_ctor_set(x_480, 0, x_462); +lean_ctor_set(x_480, 1, x_463); +lean_ctor_set(x_480, 2, x_464); +lean_ctor_set(x_480, 3, x_472); +x_481 = lean_st_ref_set(x_3, x_480, x_461); lean_dec(x_3); -x_332 = lean_ctor_get(x_331, 1); -lean_inc(x_332); -if (lean_is_exclusive(x_331)) { - lean_ctor_release(x_331, 0); - lean_ctor_release(x_331, 1); - x_333 = x_331; +x_482 = lean_ctor_get(x_481, 1); +lean_inc(x_482); +if (lean_is_exclusive(x_481)) { + lean_ctor_release(x_481, 0); + lean_ctor_release(x_481, 1); + x_483 = x_481; } else { - lean_dec_ref(x_331); - x_333 = lean_box(0); + lean_dec_ref(x_481); + x_483 = lean_box(0); } -x_334 = lean_box(0); -if (lean_is_scalar(x_333)) { - x_335 = lean_alloc_ctor(0, 2, 0); +x_484 = lean_box(0); +if (lean_is_scalar(x_483)) { + x_485 = lean_alloc_ctor(0, 2, 0); } else { - x_335 = x_333; + x_485 = x_483; } -lean_ctor_set(x_335, 0, x_334); -lean_ctor_set(x_335, 1, x_332); -return x_335; +lean_ctor_set(x_485, 0, x_484); +lean_ctor_set(x_485, 1, x_482); +return x_485; } else { -uint8_t x_336; -x_336 = lean_nat_dec_le(x_325, x_325); -if (x_336 == 0) -{ -lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; -lean_dec(x_325); -lean_dec(x_324); -lean_dec(x_1); -lean_ctor_set(x_283, 3, x_330); -x_337 = lean_st_ref_set(x_3, x_283, x_284); +size_t 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; +x_486 = lean_usize_of_nat(x_467); +lean_dec(x_467); +x_487 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_consume___spec__4(x_405, x_457, x_455, x_486, x_464); +lean_dec(x_457); +if (lean_is_scalar(x_466)) { + x_488 = lean_alloc_ctor(0, 4, 0); +} else { + x_488 = x_466; +} +lean_ctor_set(x_488, 0, x_462); +lean_ctor_set(x_488, 1, x_463); +lean_ctor_set(x_488, 2, x_487); +lean_ctor_set(x_488, 3, x_472); +x_489 = lean_st_ref_set(x_3, x_488, x_461); lean_dec(x_3); -x_338 = lean_ctor_get(x_337, 1); -lean_inc(x_338); -if (lean_is_exclusive(x_337)) { - lean_ctor_release(x_337, 0); - lean_ctor_release(x_337, 1); - x_339 = x_337; +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_337); - x_339 = lean_box(0); + lean_dec_ref(x_489); + x_491 = lean_box(0); } -x_340 = lean_box(0); -if (lean_is_scalar(x_339)) { - x_341 = lean_alloc_ctor(0, 2, 0); +x_492 = lean_box(0); +if (lean_is_scalar(x_491)) { + x_493 = lean_alloc_ctor(0, 2, 0); } else { - x_341 = x_339; + x_493 = x_491; +} +lean_ctor_set(x_493, 0, x_492); +lean_ctor_set(x_493, 1, x_490); +return x_493; +} } -lean_ctor_set(x_341, 0, x_340); -lean_ctor_set(x_341, 1, x_338); -return x_341; } else { -size_t x_342; size_t 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 = 0; -x_343 = lean_usize_of_nat(x_325); -lean_dec(x_325); -x_344 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_consume___spec__5(x_1, x_324, x_342, x_343, x_321); -lean_dec(x_324); -lean_ctor_set(x_283, 3, x_330); -lean_ctor_set(x_283, 2, x_344); -x_345 = lean_st_ref_set(x_3, x_283, x_284); +lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; +lean_dec(x_452); +lean_dec(x_451); +lean_dec(x_450); +lean_dec(x_426); +lean_dec(x_408); +lean_dec(x_405); lean_dec(x_3); -x_346 = lean_ctor_get(x_345, 1); -lean_inc(x_346); -if (lean_is_exclusive(x_345)) { - lean_ctor_release(x_345, 0); - lean_ctor_release(x_345, 1); - x_347 = x_345; +x_494 = lean_ctor_get(x_456, 0); +lean_inc(x_494); +x_495 = lean_ctor_get(x_456, 1); +lean_inc(x_495); +if (lean_is_exclusive(x_456)) { + lean_ctor_release(x_456, 0); + lean_ctor_release(x_456, 1); + x_496 = x_456; } else { - lean_dec_ref(x_345); - x_347 = lean_box(0); + lean_dec_ref(x_456); + x_496 = lean_box(0); } -x_348 = lean_box(0); -if (lean_is_scalar(x_347)) { - x_349 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_496)) { + x_497 = lean_alloc_ctor(1, 2, 0); } else { - x_349 = x_347; -} -lean_ctor_set(x_349, 0, x_348); -lean_ctor_set(x_349, 1, x_346); -return x_349; + x_497 = x_496; } +lean_ctor_set(x_497, 0, x_494); +lean_ctor_set(x_497, 1, x_495); +return x_497; } } } 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; uint8_t x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; -x_350 = lean_ctor_get(x_283, 0); -x_351 = lean_ctor_get(x_283, 1); -x_352 = lean_ctor_get(x_283, 2); -x_353 = lean_ctor_get(x_283, 3); -lean_inc(x_353); -lean_inc(x_352); -lean_inc(x_351); -lean_inc(x_350); -lean_dec(x_283); -x_354 = lean_ctor_get(x_281, 0); -lean_inc(x_354); -x_355 = lean_ctor_get(x_281, 1); -lean_inc(x_355); -if (lean_is_exclusive(x_281)) { - lean_ctor_release(x_281, 0); - lean_ctor_release(x_281, 1); - x_356 = x_281; +lean_object* x_498; lean_object* x_499; lean_object* x_500; lean_object* x_501; +lean_dec(x_423); +lean_dec(x_422); +lean_dec(x_420); +lean_dec(x_408); +lean_dec(x_405); +lean_dec(x_403); +lean_dec(x_397); +lean_dec(x_395); +lean_dec(x_394); +lean_dec(x_393); +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_498 = lean_ctor_get(x_425, 0); +lean_inc(x_498); +x_499 = lean_ctor_get(x_425, 1); +lean_inc(x_499); +if (lean_is_exclusive(x_425)) { + lean_ctor_release(x_425, 0); + lean_ctor_release(x_425, 1); + x_500 = x_425; } else { - lean_dec_ref(x_281); - x_356 = lean_box(0); + lean_dec_ref(x_425); + x_500 = lean_box(0); } -x_357 = lean_array_get_size(x_355); -x_358 = lean_unsigned_to_nat(0u); -x_359 = lean_nat_dec_lt(x_358, x_357); -x_360 = lean_array_push(x_354, x_24); -lean_inc(x_355); -if (lean_is_scalar(x_356)) { - x_361 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_500)) { + x_501 = lean_alloc_ctor(1, 2, 0); } else { - x_361 = x_356; + x_501 = x_500; } -lean_ctor_set(x_361, 0, x_360); -lean_ctor_set(x_361, 1, x_355); -x_362 = l_Lean_HashMap_insert___at_Lean_Meta_SynthInstance_newSubgoal___spec__1(x_353, x_26, x_361); -if (x_359 == 0) +lean_ctor_set(x_501, 0, x_498); +lean_ctor_set(x_501, 1, x_499); +return x_501; +} +} +} +else { -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_357); -lean_dec(x_355); -lean_dec(x_1); -x_363 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_363, 0, x_350); -lean_ctor_set(x_363, 1, x_351); -lean_ctor_set(x_363, 2, x_352); -lean_ctor_set(x_363, 3, x_362); -x_364 = lean_st_ref_set(x_3, x_363, x_284); +lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; +lean_dec(x_410); +lean_dec(x_408); +lean_dec(x_407); +lean_dec(x_405); +lean_dec(x_403); +lean_dec(x_397); +lean_dec(x_395); +lean_dec(x_394); +lean_dec(x_393); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); -x_365 = lean_ctor_get(x_364, 1); -lean_inc(x_365); -if (lean_is_exclusive(x_364)) { - lean_ctor_release(x_364, 0); - lean_ctor_release(x_364, 1); - x_366 = x_364; +lean_dec(x_2); +x_502 = lean_ctor_get(x_415, 0); +lean_inc(x_502); +x_503 = lean_ctor_get(x_415, 1); +lean_inc(x_503); +if (lean_is_exclusive(x_415)) { + lean_ctor_release(x_415, 0); + lean_ctor_release(x_415, 1); + x_504 = x_415; } else { - lean_dec_ref(x_364); - x_366 = lean_box(0); + lean_dec_ref(x_415); + x_504 = lean_box(0); } -x_367 = lean_box(0); -if (lean_is_scalar(x_366)) { - x_368 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_504)) { + x_505 = lean_alloc_ctor(1, 2, 0); } else { - x_368 = x_366; + x_505 = x_504; +} +lean_ctor_set(x_505, 0, x_502); +lean_ctor_set(x_505, 1, x_503); +return x_505; } -lean_ctor_set(x_368, 0, x_367); -lean_ctor_set(x_368, 1, x_365); -return x_368; } else { -uint8_t x_369; -x_369 = lean_nat_dec_le(x_357, x_357); -if (x_369 == 0) -{ -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_dec(x_357); -lean_dec(x_355); -lean_dec(x_1); -x_370 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_370, 0, x_350); -lean_ctor_set(x_370, 1, x_351); -lean_ctor_set(x_370, 2, x_352); -lean_ctor_set(x_370, 3, x_362); -x_371 = lean_st_ref_set(x_3, x_370, x_284); -lean_dec(x_3); -x_372 = lean_ctor_get(x_371, 1); -lean_inc(x_372); -if (lean_is_exclusive(x_371)) { - lean_ctor_release(x_371, 0); - lean_ctor_release(x_371, 1); - x_373 = x_371; +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; uint8_t x_521; lean_object* x_522; lean_object* x_523; lean_object* x_524; +lean_dec(x_407); +lean_dec(x_403); +lean_dec(x_397); +lean_dec(x_395); +lean_dec(x_394); +lean_dec(x_393); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +x_506 = lean_ctor_get(x_412, 1); +lean_inc(x_506); +lean_dec(x_412); +x_507 = lean_ctor_get(x_413, 0); +lean_inc(x_507); +lean_dec(x_413); +x_508 = lean_st_ref_take(x_3, x_506); +x_509 = lean_ctor_get(x_508, 0); +lean_inc(x_509); +x_510 = lean_ctor_get(x_508, 1); +lean_inc(x_510); +lean_dec(x_508); +x_511 = lean_ctor_get(x_509, 0); +lean_inc(x_511); +x_512 = lean_ctor_get(x_509, 1); +lean_inc(x_512); +x_513 = lean_ctor_get(x_509, 2); +lean_inc(x_513); +x_514 = lean_ctor_get(x_509, 3); +lean_inc(x_514); +if (lean_is_exclusive(x_509)) { + lean_ctor_release(x_509, 0); + lean_ctor_release(x_509, 1); + lean_ctor_release(x_509, 2); + lean_ctor_release(x_509, 3); + x_515 = x_509; } else { - lean_dec_ref(x_371); - x_373 = lean_box(0); + lean_dec_ref(x_509); + x_515 = lean_box(0); } -x_374 = lean_box(0); -if (lean_is_scalar(x_373)) { - x_375 = lean_alloc_ctor(0, 2, 0); +x_516 = lean_ctor_get(x_507, 0); +lean_inc(x_516); +x_517 = lean_ctor_get(x_507, 1); +lean_inc(x_517); +if (lean_is_exclusive(x_507)) { + lean_ctor_release(x_507, 0); + lean_ctor_release(x_507, 1); + x_518 = x_507; } else { - x_375 = x_373; + lean_dec_ref(x_507); + x_518 = lean_box(0); } -lean_ctor_set(x_375, 0, x_374); -lean_ctor_set(x_375, 1, x_372); -return x_375; +x_519 = lean_array_get_size(x_517); +x_520 = lean_unsigned_to_nat(0u); +x_521 = lean_nat_dec_lt(x_520, x_519); +x_522 = lean_array_push(x_516, x_408); +lean_inc(x_517); +if (lean_is_scalar(x_518)) { + x_523 = lean_alloc_ctor(0, 2, 0); +} else { + x_523 = x_518; } -else +lean_ctor_set(x_523, 0, x_522); +lean_ctor_set(x_523, 1, x_517); +x_524 = l_Lean_HashMap_insert___at_Lean_Meta_SynthInstance_newSubgoal___spec__1(x_514, x_410, x_523); +if (x_521 == 0) { -size_t x_376; size_t 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; -x_376 = 0; -x_377 = lean_usize_of_nat(x_357); -lean_dec(x_357); -x_378 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_consume___spec__5(x_1, x_355, x_376, x_377, x_352); -lean_dec(x_355); -x_379 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_379, 0, x_350); -lean_ctor_set(x_379, 1, x_351); -lean_ctor_set(x_379, 2, x_378); -lean_ctor_set(x_379, 3, x_362); -x_380 = lean_st_ref_set(x_3, x_379, x_284); +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_dec(x_519); +lean_dec(x_517); +lean_dec(x_405); +if (lean_is_scalar(x_515)) { + x_525 = lean_alloc_ctor(0, 4, 0); +} else { + x_525 = x_515; +} +lean_ctor_set(x_525, 0, x_511); +lean_ctor_set(x_525, 1, x_512); +lean_ctor_set(x_525, 2, x_513); +lean_ctor_set(x_525, 3, x_524); +x_526 = lean_st_ref_set(x_3, x_525, x_510); lean_dec(x_3); -x_381 = lean_ctor_get(x_380, 1); -lean_inc(x_381); -if (lean_is_exclusive(x_380)) { - lean_ctor_release(x_380, 0); - lean_ctor_release(x_380, 1); - x_382 = x_380; +x_527 = lean_ctor_get(x_526, 1); +lean_inc(x_527); +if (lean_is_exclusive(x_526)) { + lean_ctor_release(x_526, 0); + lean_ctor_release(x_526, 1); + x_528 = x_526; } else { - lean_dec_ref(x_380); - x_382 = lean_box(0); + lean_dec_ref(x_526); + x_528 = lean_box(0); } -x_383 = lean_box(0); -if (lean_is_scalar(x_382)) { - x_384 = lean_alloc_ctor(0, 2, 0); +x_529 = lean_box(0); +if (lean_is_scalar(x_528)) { + x_530 = lean_alloc_ctor(0, 2, 0); } else { - x_384 = x_382; -} -lean_ctor_set(x_384, 0, x_383); -lean_ctor_set(x_384, 1, x_381); -return x_384; -} -} -} + x_530 = x_528; } +lean_ctor_set(x_530, 0, x_529); +lean_ctor_set(x_530, 1, x_527); +return x_530; } else { -uint8_t x_385; -lean_dec(x_24); -lean_dec(x_23); -lean_dec(x_1); -lean_dec(x_20); -lean_dec(x_14); -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_385 = !lean_is_exclusive(x_25); -if (x_385 == 0) -{ -return x_25; -} -else +uint8_t x_531; +x_531 = lean_nat_dec_le(x_519, x_519); +if (x_531 == 0) { -lean_object* x_386; lean_object* x_387; lean_object* x_388; -x_386 = lean_ctor_get(x_25, 0); -x_387 = lean_ctor_get(x_25, 1); -lean_inc(x_387); -lean_inc(x_386); -lean_dec(x_25); -x_388 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_388, 0, x_386); -lean_ctor_set(x_388, 1, x_387); -return x_388; -} +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_dec(x_519); +lean_dec(x_517); +lean_dec(x_405); +if (lean_is_scalar(x_515)) { + x_532 = lean_alloc_ctor(0, 4, 0); +} else { + x_532 = x_515; } +lean_ctor_set(x_532, 0, x_511); +lean_ctor_set(x_532, 1, x_512); +lean_ctor_set(x_532, 2, x_513); +lean_ctor_set(x_532, 3, x_524); +x_533 = lean_st_ref_set(x_3, x_532, x_510); +lean_dec(x_3); +x_534 = lean_ctor_get(x_533, 1); +lean_inc(x_534); +if (lean_is_exclusive(x_533)) { + lean_ctor_release(x_533, 0); + lean_ctor_release(x_533, 1); + x_535 = x_533; +} else { + lean_dec_ref(x_533); + x_535 = lean_box(0); } +x_536 = lean_box(0); +if (lean_is_scalar(x_535)) { + x_537 = lean_alloc_ctor(0, 2, 0); +} else { + x_537 = x_535; } -else -{ -uint8_t x_389; -lean_free_object(x_1); -lean_dec(x_14); -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_389 = !lean_is_exclusive(x_19); -if (x_389 == 0) -{ -return x_19; +lean_ctor_set(x_537, 0, x_536); +lean_ctor_set(x_537, 1, x_534); +return x_537; } else { -lean_object* x_390; lean_object* x_391; lean_object* x_392; -x_390 = lean_ctor_get(x_19, 0); -x_391 = lean_ctor_get(x_19, 1); -lean_inc(x_391); -lean_inc(x_390); -lean_dec(x_19); -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; +size_t x_538; size_t 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; +x_538 = 0; +x_539 = lean_usize_of_nat(x_519); +lean_dec(x_519); +x_540 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_consume___spec__5(x_405, x_517, x_538, x_539, x_513); +lean_dec(x_517); +if (lean_is_scalar(x_515)) { + x_541 = lean_alloc_ctor(0, 4, 0); +} else { + x_541 = x_515; } +lean_ctor_set(x_541, 0, x_511); +lean_ctor_set(x_541, 1, x_512); +lean_ctor_set(x_541, 2, x_540); +lean_ctor_set(x_541, 3, x_524); +x_542 = lean_st_ref_set(x_3, x_541, x_510); +lean_dec(x_3); +x_543 = lean_ctor_get(x_542, 1); +lean_inc(x_543); +if (lean_is_exclusive(x_542)) { + lean_ctor_release(x_542, 0); + lean_ctor_release(x_542, 1); + x_544 = x_542; +} else { + lean_dec_ref(x_542); + x_544 = lean_box(0); } +x_545 = lean_box(0); +if (lean_is_scalar(x_544)) { + x_546 = lean_alloc_ctor(0, 2, 0); +} else { + x_546 = x_544; } -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; -x_393 = lean_ctor_get(x_1, 0); -x_394 = lean_ctor_get(x_1, 1); -x_395 = lean_ctor_get(x_1, 2); -x_396 = lean_ctor_get(x_1, 3); -x_397 = lean_ctor_get(x_1, 4); -lean_inc(x_397); -lean_inc(x_396); -lean_inc(x_395); -lean_inc(x_394); -lean_inc(x_393); -lean_dec(x_1); -x_398 = lean_box(0); -x_399 = lean_alloc_closure((void*)(l_List_filterAuxM___at_Lean_Meta_SynthInstance_consume___spec__2___boxed), 9, 2); -lean_closure_set(x_399, 0, x_396); -lean_closure_set(x_399, 1, x_398); -x_400 = l_Lean_Meta_SynthInstance_consume___closed__1; -x_401 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_SynthInstance_newSubgoal___spec__7___rarg), 9, 2); -lean_closure_set(x_401, 0, x_399); -lean_closure_set(x_401, 1, x_400); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_395); -x_402 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__18___rarg(x_395, x_401, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_402) == 0) -{ -lean_object* x_403; lean_object* x_404; lean_object* x_405; -x_403 = lean_ctor_get(x_402, 0); -lean_inc(x_403); -x_404 = lean_ctor_get(x_402, 1); -lean_inc(x_404); -lean_dec(x_402); -lean_inc(x_397); -lean_inc(x_403); -lean_inc(x_395); -lean_inc(x_394); -lean_inc(x_393); -x_405 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_405, 0, x_393); -lean_ctor_set(x_405, 1, x_394); -lean_ctor_set(x_405, 2, x_395); -lean_ctor_set(x_405, 3, x_403); -lean_ctor_set(x_405, 4, x_397); -if (lean_obj_tag(x_403) == 0) -{ -lean_object* x_406; -lean_dec(x_397); -lean_dec(x_395); -lean_dec(x_394); -lean_dec(x_393); -x_406 = l_Lean_Meta_SynthInstance_addAnswer(x_405, x_2, x_3, x_4, x_5, x_6, x_7, x_404); -return x_406; +lean_ctor_set(x_546, 0, x_545); +lean_ctor_set(x_546, 1, x_543); +return x_546; } -else -{ -lean_object* x_407; lean_object* x_408; lean_object* x_409; -x_407 = lean_ctor_get(x_403, 0); -lean_inc(x_407); -lean_inc(x_405); -x_408 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_408, 0, x_405); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_407); -lean_inc(x_395); -x_409 = l_Lean_Meta_SynthInstance_mkTableKeyFor(x_395, x_407, x_2, x_3, x_4, x_5, x_6, x_7, x_404); -if (lean_obj_tag(x_409) == 0) -{ -lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; -x_410 = lean_ctor_get(x_409, 0); -lean_inc(x_410); -x_411 = lean_ctor_get(x_409, 1); -lean_inc(x_411); -lean_dec(x_409); -lean_inc(x_410); -x_412 = l_Lean_Meta_SynthInstance_findEntry_x3f(x_410, x_2, x_3, x_4, x_5, x_6, x_7, x_411); -x_413 = lean_ctor_get(x_412, 0); -lean_inc(x_413); -if (lean_obj_tag(x_413) == 0) -{ -lean_object* x_414; lean_object* x_415; -x_414 = lean_ctor_get(x_412, 1); -lean_inc(x_414); -lean_dec(x_412); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_407); -lean_inc(x_395); -x_415 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f(x_395, x_407, x_4, x_5, x_6, x_7, x_414); -if (lean_obj_tag(x_415) == 0) -{ -lean_object* x_416; -x_416 = lean_ctor_get(x_415, 0); -lean_inc(x_416); -if (lean_obj_tag(x_416) == 0) +} +} +} +else { -lean_object* x_417; lean_object* x_418; +lean_object* x_547; lean_object* x_548; lean_object* x_549; lean_object* x_550; +lean_dec(x_408); +lean_dec(x_407); lean_dec(x_405); lean_dec(x_403); lean_dec(x_397); +lean_dec(x_395); lean_dec(x_394); lean_dec(x_393); -x_417 = lean_ctor_get(x_415, 1); -lean_inc(x_417); -lean_dec(x_415); -x_418 = l_Lean_Meta_SynthInstance_newSubgoal(x_395, x_410, x_407, x_408, x_2, x_3, x_4, x_5, x_6, x_7, x_417); -return x_418; +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_547 = lean_ctor_get(x_409, 0); +lean_inc(x_547); +x_548 = lean_ctor_get(x_409, 1); +lean_inc(x_548); +if (lean_is_exclusive(x_409)) { + lean_ctor_release(x_409, 0); + lean_ctor_release(x_409, 1); + x_549 = x_409; +} else { + lean_dec_ref(x_409); + x_549 = lean_box(0); +} +if (lean_is_scalar(x_549)) { + x_550 = lean_alloc_ctor(1, 2, 0); +} else { + x_550 = x_549; +} +lean_ctor_set(x_550, 0, x_547); +lean_ctor_set(x_550, 1, x_548); +return x_550; +} +} } 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_dec(x_410); -lean_dec(x_407); -x_419 = lean_ctor_get(x_416, 0); -lean_inc(x_419); -if (lean_is_exclusive(x_416)) { - lean_ctor_release(x_416, 0); - x_420 = x_416; +lean_object* x_551; lean_object* x_552; lean_object* x_553; lean_object* x_554; +lean_dec(x_397); +lean_dec(x_395); +lean_dec(x_394); +lean_dec(x_393); +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_551 = lean_ctor_get(x_402, 0); +lean_inc(x_551); +x_552 = lean_ctor_get(x_402, 1); +lean_inc(x_552); +if (lean_is_exclusive(x_402)) { + lean_ctor_release(x_402, 0); + lean_ctor_release(x_402, 1); + x_553 = x_402; } else { - lean_dec_ref(x_416); - x_420 = lean_box(0); + lean_dec_ref(x_402); + x_553 = lean_box(0); } -x_421 = lean_ctor_get(x_415, 1); -lean_inc(x_421); -lean_dec(x_415); -x_422 = lean_ctor_get(x_419, 0); -lean_inc(x_422); -x_423 = lean_ctor_get(x_419, 1); -lean_inc(x_423); -lean_dec(x_419); -lean_inc(x_422); -x_424 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_mkTableKey___at_Lean_Meta_SynthInstance_mkTableKeyFor___spec__2___boxed), 8, 1); -lean_closure_set(x_424, 0, x_422); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_395); -x_425 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__18___rarg(x_395, x_424, x_2, x_3, x_4, x_5, x_6, x_7, x_421); -if (lean_obj_tag(x_425) == 0) +if (lean_is_scalar(x_553)) { + x_554 = lean_alloc_ctor(1, 2, 0); +} else { + x_554 = x_553; +} +lean_ctor_set(x_554, 0, x_551); +lean_ctor_set(x_554, 1, x_552); +return x_554; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at_Lean_Meta_SynthInstance_consume___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: { -lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; -x_426 = lean_ctor_get(x_425, 0); -lean_inc(x_426); -x_427 = lean_ctor_get(x_425, 1); -lean_inc(x_427); -lean_dec(x_425); -lean_inc(x_426); -x_428 = l_Lean_Meta_SynthInstance_findEntry_x3f(x_426, x_2, x_3, x_4, x_5, x_6, x_7, x_427); -x_429 = lean_ctor_get(x_428, 0); -lean_inc(x_429); -if (lean_obj_tag(x_429) == 0) +lean_object* x_9; +x_9 = l_Lean_MVarId_isAssigned___at_Lean_Meta_SynthInstance_consume___spec__1(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_List_filterAuxM___at_Lean_Meta_SynthInstance_consume___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: { -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_dec(x_423); -lean_dec(x_408); -lean_dec(x_405); -x_430 = lean_ctor_get(x_428, 1); -lean_inc(x_430); -lean_dec(x_428); -if (lean_is_scalar(x_420)) { - x_431 = lean_alloc_ctor(1, 1, 0); -} else { - x_431 = x_420; +lean_object* x_10; +x_10 = l_List_filterAuxM___at_Lean_Meta_SynthInstance_consume___spec__2(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); +return x_10; } -lean_ctor_set(x_431, 0, x_422); -x_432 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_consume___lambda__2___boxed), 8, 1); -lean_closure_set(x_432, 0, x_431); -x_433 = l_Lean_Meta_SynthInstance_consume___closed__2; -x_434 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_SynthInstance_newSubgoal___spec__7___rarg), 9, 2); -lean_closure_set(x_434, 0, x_432); -lean_closure_set(x_434, 1, x_433); -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_435 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__18___rarg(x_395, x_434, x_2, x_3, x_4, x_5, x_6, x_7, x_430); -if (lean_obj_tag(x_435) == 0) +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_SynthInstance_consume___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) { +_start: { -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; -x_436 = lean_ctor_get(x_435, 0); -lean_inc(x_436); -x_437 = lean_ctor_get(x_435, 1); -lean_inc(x_437); -lean_dec(x_435); -x_438 = lean_ctor_get(x_436, 0); -lean_inc(x_438); -x_439 = lean_ctor_get(x_436, 1); -lean_inc(x_439); -lean_dec(x_436); -lean_inc(x_439); -x_440 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_440, 0, x_439); -lean_ctor_set(x_440, 1, x_403); -lean_inc(x_438); -x_441 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_441, 0, x_393); -lean_ctor_set(x_441, 1, x_394); -lean_ctor_set(x_441, 2, x_438); -lean_ctor_set(x_441, 3, x_440); -lean_ctor_set(x_441, 4, x_397); -x_442 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_442, 0, x_441); -x_443 = l_Lean_Meta_SynthInstance_newSubgoal(x_438, x_426, x_439, x_442, x_2, x_3, x_4, x_5, x_6, x_7, x_437); -return x_443; +lean_object* x_12; +x_12 = l_Array_mapMUnsafe_map___at_Lean_Meta_SynthInstance_consume___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); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_SynthInstance_consume___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) { +_start: +{ +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 = l_Array_mapMUnsafe_map___at_Lean_Meta_SynthInstance_consume___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); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_consume___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_3); +lean_dec(x_3); +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_consume___spec__4(x_1, x_2, x_6, x_7, x_5); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_consume___spec__5___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_3); +lean_dec(x_3); +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_consume___spec__5(x_1, x_2, x_6, x_7, x_5); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_consume___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_Meta_SynthInstance_consume___lambda__1(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_Meta_SynthInstance_consume___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) { +_start: +{ +lean_object* x_9; +x_9 = l_Lean_Meta_SynthInstance_consume___lambda__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_3); +lean_dec(x_2); +return x_9; } -else +} +LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_consume___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) { +_start: { -lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; -lean_dec(x_426); -lean_dec(x_403); -lean_dec(x_397); -lean_dec(x_394); -lean_dec(x_393); +lean_object* x_9; +x_9 = l_Lean_Meta_SynthInstance_consume___lambda__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); lean_dec(x_2); -x_444 = lean_ctor_get(x_435, 0); -lean_inc(x_444); -x_445 = lean_ctor_get(x_435, 1); -lean_inc(x_445); -if (lean_is_exclusive(x_435)) { - lean_ctor_release(x_435, 0); - lean_ctor_release(x_435, 1); - x_446 = x_435; -} else { - lean_dec_ref(x_435); - x_446 = lean_box(0); -} -if (lean_is_scalar(x_446)) { - x_447 = lean_alloc_ctor(1, 2, 0); -} else { - x_447 = x_446; +return x_9; } -lean_ctor_set(x_447, 0, x_444); -lean_ctor_set(x_447, 1, x_445); -return x_447; } +LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_getTop___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; uint8_t x_8; +x_7 = lean_st_ref_get(x_1, x_6); +x_8 = !lean_is_exclusive(x_7); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_9 = lean_ctor_get(x_7, 0); +x_10 = lean_ctor_get(x_9, 1); +lean_inc(x_10); +lean_dec(x_9); +x_11 = l_Lean_Meta_SynthInstance_instInhabitedGeneratorNode; +x_12 = l_Array_back___rarg(x_11, x_10); +lean_dec(x_10); +lean_ctor_set(x_7, 0, x_12); +return x_7; } else { -lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_object* x_453; size_t x_454; size_t x_455; lean_object* x_456; -lean_dec(x_422); -lean_dec(x_420); -lean_dec(x_403); -lean_dec(x_397); -lean_dec(x_394); -lean_dec(x_393); -x_448 = lean_ctor_get(x_429, 0); -lean_inc(x_448); -lean_dec(x_429); -x_449 = lean_ctor_get(x_428, 1); -lean_inc(x_449); -lean_dec(x_428); -x_450 = lean_ctor_get(x_448, 0); -lean_inc(x_450); -x_451 = lean_ctor_get(x_448, 1); -lean_inc(x_451); -if (lean_is_exclusive(x_448)) { - lean_ctor_release(x_448, 0); - lean_ctor_release(x_448, 1); - x_452 = x_448; -} else { - lean_dec_ref(x_448); - x_452 = lean_box(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; +x_13 = lean_ctor_get(x_7, 0); +x_14 = lean_ctor_get(x_7, 1); +lean_inc(x_14); +lean_inc(x_13); +lean_dec(x_7); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Lean_Meta_SynthInstance_instInhabitedGeneratorNode; +x_17 = l_Array_back___rarg(x_16, x_15); +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_14); +return x_18; } -x_453 = lean_array_get_size(x_451); -x_454 = lean_usize_of_nat(x_453); -lean_dec(x_453); -x_455 = 0; -lean_inc(x_3); -lean_inc(x_451); -x_456 = l_Array_mapMUnsafe_map___at_Lean_Meta_SynthInstance_consume___spec__3(x_395, x_423, x_454, x_455, x_451, x_2, x_3, x_4, x_5, x_6, x_7, x_449); -if (lean_obj_tag(x_456) == 0) -{ -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; uint8_t x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; -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_st_ref_take(x_3, x_458); -x_460 = lean_ctor_get(x_459, 0); -lean_inc(x_460); -x_461 = lean_ctor_get(x_459, 1); -lean_inc(x_461); -lean_dec(x_459); -x_462 = lean_ctor_get(x_460, 0); -lean_inc(x_462); -x_463 = lean_ctor_get(x_460, 1); -lean_inc(x_463); -x_464 = lean_ctor_get(x_460, 2); -lean_inc(x_464); -x_465 = lean_ctor_get(x_460, 3); -lean_inc(x_465); -if (lean_is_exclusive(x_460)) { - lean_ctor_release(x_460, 0); - lean_ctor_release(x_460, 1); - lean_ctor_release(x_460, 2); - lean_ctor_release(x_460, 3); - x_466 = x_460; -} else { - lean_dec_ref(x_460); - x_466 = lean_box(0); } -x_467 = lean_array_get_size(x_457); -x_468 = lean_unsigned_to_nat(0u); -x_469 = lean_nat_dec_lt(x_468, x_467); -x_470 = lean_array_push(x_450, x_408); -if (lean_is_scalar(x_452)) { - x_471 = lean_alloc_ctor(0, 2, 0); -} else { - x_471 = x_452; } -lean_ctor_set(x_471, 0, x_470); -lean_ctor_set(x_471, 1, x_451); -x_472 = l_Lean_HashMap_insert___at_Lean_Meta_SynthInstance_newSubgoal___spec__1(x_465, x_426, x_471); -if (x_469 == 0) +LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_getTop(lean_object* x_1) { +_start: { -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_dec(x_467); -lean_dec(x_457); -lean_dec(x_405); -if (lean_is_scalar(x_466)) { - x_473 = lean_alloc_ctor(0, 4, 0); -} else { - x_473 = x_466; +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_getTop___rarg___boxed), 6, 0); +return x_2; } -lean_ctor_set(x_473, 0, x_462); -lean_ctor_set(x_473, 1, x_463); -lean_ctor_set(x_473, 2, x_464); -lean_ctor_set(x_473, 3, x_472); -x_474 = lean_st_ref_set(x_3, x_473, x_461); +} +LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_getTop___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_Meta_SynthInstance_getTop___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); -x_475 = lean_ctor_get(x_474, 1); -lean_inc(x_475); -if (lean_is_exclusive(x_474)) { - lean_ctor_release(x_474, 0); - lean_ctor_release(x_474, 1); - x_476 = x_474; -} else { - lean_dec_ref(x_474); - x_476 = lean_box(0); +lean_dec(x_2); +lean_dec(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_getTop___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_Meta_SynthInstance_getTop(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_modifyTop(lean_object* x_1, lean_object* x_2, 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; +x_9 = lean_st_ref_take(x_3, 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_is_exclusive(x_10); +if (x_12 == 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_10, 1); +x_14 = lean_array_get_size(x_13); +x_15 = lean_unsigned_to_nat(1u); +x_16 = lean_nat_sub(x_14, x_15); +x_17 = lean_nat_dec_lt(x_16, x_14); +lean_dec(x_14); +if (x_17 == 0) +{ +lean_object* x_18; uint8_t x_19; +lean_dec(x_16); +lean_dec(x_1); +x_18 = lean_st_ref_set(x_3, x_10, x_11); +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_18, 0); +lean_dec(x_20); +x_21 = lean_box(0); +lean_ctor_set(x_18, 0, x_21); +return x_18; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_18, 1); +lean_inc(x_22); +lean_dec(x_18); +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; +} +} +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; uint8_t x_31; +x_25 = lean_array_fget(x_13, x_16); +x_26 = lean_box(0); +x_27 = lean_array_fset(x_13, x_16, x_26); +x_28 = lean_apply_1(x_1, x_25); +x_29 = lean_array_fset(x_27, x_16, x_28); +lean_dec(x_16); +lean_ctor_set(x_10, 1, x_29); +x_30 = lean_st_ref_set(x_3, x_10, x_11); +x_31 = !lean_is_exclusive(x_30); +if (x_31 == 0) +{ +lean_object* x_32; +x_32 = lean_ctor_get(x_30, 0); +lean_dec(x_32); +lean_ctor_set(x_30, 0, x_26); +return x_30; +} +else +{ +lean_object* x_33; lean_object* x_34; +x_33 = lean_ctor_get(x_30, 1); +lean_inc(x_33); +lean_dec(x_30); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_26); +lean_ctor_set(x_34, 1, x_33); +return x_34; } -x_477 = lean_box(0); -if (lean_is_scalar(x_476)) { - x_478 = lean_alloc_ctor(0, 2, 0); -} else { - x_478 = x_476; } -lean_ctor_set(x_478, 0, x_477); -lean_ctor_set(x_478, 1, x_475); -return x_478; } else { -uint8_t x_479; -x_479 = lean_nat_dec_le(x_467, x_467); -if (x_479 == 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; uint8_t x_42; +x_35 = lean_ctor_get(x_10, 0); +x_36 = lean_ctor_get(x_10, 1); +x_37 = lean_ctor_get(x_10, 2); +x_38 = lean_ctor_get(x_10, 3); +lean_inc(x_38); +lean_inc(x_37); +lean_inc(x_36); +lean_inc(x_35); +lean_dec(x_10); +x_39 = lean_array_get_size(x_36); +x_40 = lean_unsigned_to_nat(1u); +x_41 = lean_nat_sub(x_39, x_40); +x_42 = lean_nat_dec_lt(x_41, x_39); +lean_dec(x_39); +if (x_42 == 0) { -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_dec(x_467); -lean_dec(x_457); -lean_dec(x_405); -if (lean_is_scalar(x_466)) { - x_480 = lean_alloc_ctor(0, 4, 0); -} else { - x_480 = x_466; -} -lean_ctor_set(x_480, 0, x_462); -lean_ctor_set(x_480, 1, x_463); -lean_ctor_set(x_480, 2, x_464); -lean_ctor_set(x_480, 3, x_472); -x_481 = lean_st_ref_set(x_3, x_480, x_461); -lean_dec(x_3); -x_482 = lean_ctor_get(x_481, 1); -lean_inc(x_482); -if (lean_is_exclusive(x_481)) { - lean_ctor_release(x_481, 0); - lean_ctor_release(x_481, 1); - x_483 = x_481; +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_41); +lean_dec(x_1); +x_43 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_43, 0, x_35); +lean_ctor_set(x_43, 1, x_36); +lean_ctor_set(x_43, 2, x_37); +lean_ctor_set(x_43, 3, x_38); +x_44 = lean_st_ref_set(x_3, x_43, x_11); +x_45 = lean_ctor_get(x_44, 1); +lean_inc(x_45); +if (lean_is_exclusive(x_44)) { + lean_ctor_release(x_44, 0); + lean_ctor_release(x_44, 1); + x_46 = x_44; } else { - lean_dec_ref(x_481); - x_483 = lean_box(0); + lean_dec_ref(x_44); + x_46 = lean_box(0); } -x_484 = lean_box(0); -if (lean_is_scalar(x_483)) { - x_485 = lean_alloc_ctor(0, 2, 0); +x_47 = lean_box(0); +if (lean_is_scalar(x_46)) { + x_48 = lean_alloc_ctor(0, 2, 0); } else { - x_485 = x_483; + x_48 = x_46; } -lean_ctor_set(x_485, 0, x_484); -lean_ctor_set(x_485, 1, x_482); -return x_485; +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set(x_48, 1, x_45); +return x_48; } else { -size_t 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; -x_486 = lean_usize_of_nat(x_467); -lean_dec(x_467); -x_487 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_consume___spec__4(x_405, x_457, x_455, x_486, x_464); -lean_dec(x_457); -if (lean_is_scalar(x_466)) { - x_488 = lean_alloc_ctor(0, 4, 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; +x_49 = lean_array_fget(x_36, x_41); +x_50 = lean_box(0); +x_51 = lean_array_fset(x_36, x_41, x_50); +x_52 = lean_apply_1(x_1, x_49); +x_53 = lean_array_fset(x_51, x_41, x_52); +lean_dec(x_41); +x_54 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_54, 0, x_35); +lean_ctor_set(x_54, 1, x_53); +lean_ctor_set(x_54, 2, x_37); +lean_ctor_set(x_54, 3, x_38); +x_55 = lean_st_ref_set(x_3, x_54, x_11); +x_56 = lean_ctor_get(x_55, 1); +lean_inc(x_56); +if (lean_is_exclusive(x_55)) { + lean_ctor_release(x_55, 0); + lean_ctor_release(x_55, 1); + x_57 = x_55; } else { - x_488 = x_466; + lean_dec_ref(x_55); + x_57 = lean_box(0); } -lean_ctor_set(x_488, 0, x_462); -lean_ctor_set(x_488, 1, x_463); -lean_ctor_set(x_488, 2, x_487); -lean_ctor_set(x_488, 3, x_472); -x_489 = lean_st_ref_set(x_3, x_488, x_461); -lean_dec(x_3); -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; +if (lean_is_scalar(x_57)) { + x_58 = lean_alloc_ctor(0, 2, 0); } else { - lean_dec_ref(x_489); - x_491 = lean_box(0); + x_58 = x_57; } -x_492 = lean_box(0); -if (lean_is_scalar(x_491)) { - x_493 = lean_alloc_ctor(0, 2, 0); -} else { - x_493 = x_491; +lean_ctor_set(x_58, 0, x_50); +lean_ctor_set(x_58, 1, x_56); +return x_58; } -lean_ctor_set(x_493, 0, x_492); -lean_ctor_set(x_493, 1, x_490); -return x_493; } } } -else +LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_modifyTop___boxed(lean_object* x_1, lean_object* x_2, 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_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; -lean_dec(x_452); -lean_dec(x_451); -lean_dec(x_450); -lean_dec(x_426); -lean_dec(x_408); -lean_dec(x_405); +lean_object* x_9; +x_9 = l_Lean_Meta_SynthInstance_modifyTop(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); -x_494 = lean_ctor_get(x_456, 0); -lean_inc(x_494); -x_495 = lean_ctor_get(x_456, 1); -lean_inc(x_495); -if (lean_is_exclusive(x_456)) { - lean_ctor_release(x_456, 0); - lean_ctor_release(x_456, 1); - x_496 = x_456; -} else { - lean_dec_ref(x_456); - x_496 = lean_box(0); +lean_dec(x_2); +return x_9; } -if (lean_is_scalar(x_496)) { - x_497 = lean_alloc_ctor(1, 2, 0); -} else { - x_497 = x_496; } -lean_ctor_set(x_497, 0, x_494); -lean_ctor_set(x_497, 1, x_495); -return x_497; +LEAN_EXPORT lean_object* l_Lean_withSeconds___at_Lean_Meta_SynthInstance_generate___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; +x_9 = lean_io_mono_nanos_now(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_apply_7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_11); +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_io_mono_nanos_now(x_14); +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; uint8_t x_19; lean_object* x_20; double x_21; double x_22; double x_23; lean_object* x_24; lean_object* x_25; +x_17 = lean_ctor_get(x_15, 0); +x_18 = lean_nat_sub(x_17, x_10); +lean_dec(x_10); +lean_dec(x_17); +x_19 = 0; +x_20 = lean_unsigned_to_nat(0u); +x_21 = l_Float_ofScientific(x_18, x_19, x_20); +lean_dec(x_18); +x_22 = l_Lean_withSeconds___at_Lean_Meta_SynthInstance_newSubgoal___spec__12___closed__1; +x_23 = lean_float_div(x_21, x_22); +x_24 = lean_box_float(x_23); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_13); +lean_ctor_set(x_25, 1, x_24); +lean_ctor_set(x_15, 0, x_25); +return x_15; +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; lean_object* x_30; double x_31; double x_32; double x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_26 = lean_ctor_get(x_15, 0); +x_27 = lean_ctor_get(x_15, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_15); +x_28 = lean_nat_sub(x_26, x_10); +lean_dec(x_10); +lean_dec(x_26); +x_29 = 0; +x_30 = lean_unsigned_to_nat(0u); +x_31 = l_Float_ofScientific(x_28, x_29, x_30); +lean_dec(x_28); +x_32 = l_Lean_withSeconds___at_Lean_Meta_SynthInstance_newSubgoal___spec__12___closed__1; +x_33 = lean_float_div(x_31, x_32); +x_34 = lean_box_float(x_33); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_13); +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_27); +return x_36; } } +else +{ +uint8_t x_37; +lean_dec(x_10); +x_37 = !lean_is_exclusive(x_12); +if (x_37 == 0) +{ +return x_12; } else { -lean_object* x_498; lean_object* x_499; lean_object* x_500; lean_object* x_501; -lean_dec(x_423); -lean_dec(x_422); -lean_dec(x_420); -lean_dec(x_408); -lean_dec(x_405); -lean_dec(x_403); -lean_dec(x_397); -lean_dec(x_395); -lean_dec(x_394); -lean_dec(x_393); -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_498 = lean_ctor_get(x_425, 0); -lean_inc(x_498); -x_499 = lean_ctor_get(x_425, 1); -lean_inc(x_499); -if (lean_is_exclusive(x_425)) { - lean_ctor_release(x_425, 0); - lean_ctor_release(x_425, 1); - x_500 = x_425; -} else { - lean_dec_ref(x_425); - x_500 = lean_box(0); +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_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; } -if (lean_is_scalar(x_500)) { - x_501 = lean_alloc_ctor(1, 2, 0); -} else { - x_501 = x_500; } -lean_ctor_set(x_501, 0, x_498); -lean_ctor_set(x_501, 1, x_499); -return x_501; } } +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_generate___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: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_9; lean_object* x_10; +x_9 = lean_ctor_get(x_1, 0); +lean_inc(x_9); +x_10 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_8); +return x_10; } else { -lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; -lean_dec(x_410); -lean_dec(x_408); -lean_dec(x_407); -lean_dec(x_405); -lean_dec(x_403); -lean_dec(x_397); -lean_dec(x_395); -lean_dec(x_394); -lean_dec(x_393); -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_502 = lean_ctor_get(x_415, 0); -lean_inc(x_502); -x_503 = lean_ctor_get(x_415, 1); -lean_inc(x_503); -if (lean_is_exclusive(x_415)) { - lean_ctor_release(x_415, 0); - lean_ctor_release(x_415, 1); - x_504 = x_415; -} else { - lean_dec_ref(x_415); - x_504 = lean_box(0); +lean_object* x_11; lean_object* x_12; +x_11 = lean_ctor_get(x_1, 0); +lean_inc(x_11); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_8); +return x_12; } -if (lean_is_scalar(x_504)) { - x_505 = lean_alloc_ctor(1, 2, 0); -} else { - x_505 = x_504; } -lean_ctor_set(x_505, 0, x_502); -lean_ctor_set(x_505, 1, x_503); -return x_505; +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_generate___spec__1___lambda__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) { +_start: +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +lean_inc(x_12); +x_15 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__13(x_1, x_2, x_3, x_6, x_4, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +x_16 = lean_ctor_get(x_15, 1); +lean_inc(x_16); +lean_dec(x_15); +x_17 = l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_generate___spec__3(x_5, x_8, x_9, x_10, x_11, x_12, x_13, x_16); +lean_dec(x_12); +lean_dec(x_8); +return x_17; } } -else +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_generate___spec__1___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, uint8_t x_7, double 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_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; uint8_t x_521; lean_object* x_522; lean_object* x_523; lean_object* x_524; -lean_dec(x_407); -lean_dec(x_403); -lean_dec(x_397); -lean_dec(x_395); -lean_dec(x_394); -lean_dec(x_393); -lean_dec(x_7); +lean_object* x_17; lean_object* x_18; +lean_dec(x_9); +x_17 = lean_ctor_get(x_14, 5); +lean_inc(x_17); +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_2); +x_18 = lean_apply_8(x_1, x_2, x_10, x_11, x_12, x_13, x_14, x_15, 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; +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_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__1; +x_22 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_6, x_21); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); +if (x_22 == 0) +{ +if (x_7 == 0) +{ +lean_object* x_23; lean_object* x_24; +x_23 = lean_box(0); +x_24 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_generate___spec__1___lambda__1(x_3, x_4, x_17, x_5, x_2, x_19, x_23, x_10, x_11, x_12, x_13, x_14, x_15, x_20); +lean_dec(x_15); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_2); -x_506 = lean_ctor_get(x_412, 1); -lean_inc(x_506); -lean_dec(x_412); -x_507 = lean_ctor_get(x_413, 0); -lean_inc(x_507); -lean_dec(x_413); -x_508 = lean_st_ref_take(x_3, x_506); -x_509 = lean_ctor_get(x_508, 0); -lean_inc(x_509); -x_510 = lean_ctor_get(x_508, 1); -lean_inc(x_510); -lean_dec(x_508); -x_511 = lean_ctor_get(x_509, 0); -lean_inc(x_511); -x_512 = lean_ctor_get(x_509, 1); -lean_inc(x_512); -x_513 = lean_ctor_get(x_509, 2); -lean_inc(x_513); -x_514 = lean_ctor_get(x_509, 3); -lean_inc(x_514); -if (lean_is_exclusive(x_509)) { - lean_ctor_release(x_509, 0); - lean_ctor_release(x_509, 1); - lean_ctor_release(x_509, 2); - lean_ctor_release(x_509, 3); - x_515 = x_509; -} else { - lean_dec_ref(x_509); - x_515 = lean_box(0); -} -x_516 = lean_ctor_get(x_507, 0); -lean_inc(x_516); -x_517 = lean_ctor_get(x_507, 1); -lean_inc(x_517); -if (lean_is_exclusive(x_507)) { - lean_ctor_release(x_507, 0); - lean_ctor_release(x_507, 1); - x_518 = x_507; -} else { - lean_dec_ref(x_507); - x_518 = lean_box(0); -} -x_519 = lean_array_get_size(x_517); -x_520 = lean_unsigned_to_nat(0u); -x_521 = lean_nat_dec_lt(x_520, x_519); -x_522 = lean_array_push(x_516, x_408); -lean_inc(x_517); -if (lean_is_scalar(x_518)) { - x_523 = lean_alloc_ctor(0, 2, 0); -} else { - x_523 = x_518; +return x_24; } -lean_ctor_set(x_523, 0, x_522); -lean_ctor_set(x_523, 1, x_517); -x_524 = l_Lean_HashMap_insert___at_Lean_Meta_SynthInstance_newSubgoal___spec__1(x_514, x_410, x_523); -if (x_521 == 0) +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_dec(x_519); -lean_dec(x_517); -lean_dec(x_405); -if (lean_is_scalar(x_515)) { - x_525 = lean_alloc_ctor(0, 4, 0); -} else { - x_525 = x_515; -} -lean_ctor_set(x_525, 0, x_511); -lean_ctor_set(x_525, 1, x_512); -lean_ctor_set(x_525, 2, x_513); -lean_ctor_set(x_525, 3, x_524); -x_526 = lean_st_ref_set(x_3, x_525, x_510); -lean_dec(x_3); -x_527 = lean_ctor_get(x_526, 1); -lean_inc(x_527); -if (lean_is_exclusive(x_526)) { - lean_ctor_release(x_526, 0); - lean_ctor_release(x_526, 1); - x_528 = x_526; -} else { - lean_dec_ref(x_526); - x_528 = lean_box(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; +x_25 = lean_float_to_string(x_8); +x_26 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_26, 0, x_25); +x_27 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_27, 0, x_26); +x_28 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__3; +x_29 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_27); +x_30 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__5; +x_31 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_30); +x_32 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_19); +x_33 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; +x_34 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +x_35 = lean_box(0); +x_36 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_generate___spec__1___lambda__1(x_3, x_4, x_17, x_5, x_2, x_34, x_35, x_10, x_11, x_12, x_13, x_14, x_15, x_20); +lean_dec(x_15); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_2); +return x_36; } -x_529 = lean_box(0); -if (lean_is_scalar(x_528)) { - x_530 = lean_alloc_ctor(0, 2, 0); -} else { - x_530 = x_528; } -lean_ctor_set(x_530, 0, x_529); -lean_ctor_set(x_530, 1, x_527); -return x_530; +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; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_37 = lean_float_to_string(x_8); +x_38 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_38, 0, x_37); +x_39 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_39, 0, x_38); +x_40 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__3; +x_41 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_39); +x_42 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__5; +x_43 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_43, 0, x_41); +lean_ctor_set(x_43, 1, x_42); +x_44 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_44, 0, x_43); +lean_ctor_set(x_44, 1, x_19); +x_45 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; +x_46 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +x_47 = lean_box(0); +x_48 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_generate___spec__1___lambda__1(x_3, x_4, x_17, x_5, x_2, x_46, x_47, x_10, x_11, x_12, x_13, x_14, x_15, x_20); +lean_dec(x_15); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_2); +return x_48; +} } else { -uint8_t x_531; -x_531 = lean_nat_dec_le(x_519, x_519); -if (x_531 == 0) +uint8_t x_49; +lean_dec(x_17); +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_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_49 = !lean_is_exclusive(x_18); +if (x_49 == 0) { -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_dec(x_519); -lean_dec(x_517); -lean_dec(x_405); -if (lean_is_scalar(x_515)) { - x_532 = lean_alloc_ctor(0, 4, 0); -} else { - x_532 = x_515; +return x_18; } -lean_ctor_set(x_532, 0, x_511); -lean_ctor_set(x_532, 1, x_512); -lean_ctor_set(x_532, 2, x_513); -lean_ctor_set(x_532, 3, x_524); -x_533 = lean_st_ref_set(x_3, x_532, x_510); -lean_dec(x_3); -x_534 = lean_ctor_get(x_533, 1); -lean_inc(x_534); -if (lean_is_exclusive(x_533)) { - lean_ctor_release(x_533, 0); - lean_ctor_release(x_533, 1); - x_535 = x_533; -} else { - lean_dec_ref(x_533); - x_535 = lean_box(0); +else +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; +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_alloc_ctor(1, 2, 0); +lean_ctor_set(x_52, 0, x_50); +lean_ctor_set(x_52, 1, x_51); +return x_52; } -x_536 = lean_box(0); -if (lean_is_scalar(x_535)) { - x_537 = lean_alloc_ctor(0, 2, 0); -} else { - x_537 = x_535; } -lean_ctor_set(x_537, 0, x_536); -lean_ctor_set(x_537, 1, x_534); -return x_537; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_generate___spec__1___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t 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, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +lean_dec(x_7); +x_15 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_SynthInstance_newSubgoal___spec__11___rarg(x_13, x_14); +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_closure((void*)(l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__1), 8, 1); +lean_closure_set(x_18, 0, x_1); +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_19 = l_Lean_withSeconds___at_Lean_Meta_SynthInstance_generate___spec__2(x_18, x_8, x_9, x_10, x_11, x_12, x_13, x_17); +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_69; uint8_t x_70; +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_20, 0); +lean_inc(x_22); +x_23 = lean_ctor_get(x_20, 1); +lean_inc(x_23); +lean_dec(x_20); +x_69 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__4___closed__1; +x_70 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_5, x_69); +if (x_70 == 0) +{ +uint8_t x_71; +x_71 = 0; +x_24 = x_71; +goto block_68; } else { -size_t x_538; size_t 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; -x_538 = 0; -x_539 = lean_usize_of_nat(x_519); -lean_dec(x_519); -x_540 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_consume___spec__5(x_405, x_517, x_538, x_539, x_513); -lean_dec(x_517); -if (lean_is_scalar(x_515)) { - x_541 = lean_alloc_ctor(0, 4, 0); -} else { - x_541 = x_515; +double x_72; double x_73; uint8_t x_74; +x_72 = l_Lean_trace_profiler_threshold_getSecs(x_5); +x_73 = lean_unbox_float(x_23); +x_74 = lean_float_decLt(x_72, x_73); +x_24 = x_74; +goto block_68; } -lean_ctor_set(x_541, 0, x_511); -lean_ctor_set(x_541, 1, x_512); -lean_ctor_set(x_541, 2, x_540); -lean_ctor_set(x_541, 3, x_524); -x_542 = lean_st_ref_set(x_3, x_541, x_510); +block_68: +{ +if (x_6 == 0) +{ +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; +lean_dec(x_23); +lean_dec(x_5); lean_dec(x_3); -x_543 = lean_ctor_get(x_542, 1); -lean_inc(x_543); -if (lean_is_exclusive(x_542)) { - lean_ctor_release(x_542, 0); - lean_ctor_release(x_542, 1); - x_544 = x_542; -} else { - lean_dec_ref(x_542); - x_544 = lean_box(0); +lean_dec(x_2); +x_25 = lean_st_ref_take(x_13, x_21); +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_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; +x_29 = lean_ctor_get(x_26, 3); +x_30 = l_Lean_PersistentArray_append___rarg(x_16, x_29); +lean_ctor_set(x_26, 3, x_30); +x_31 = lean_st_ref_set(x_13, x_26, x_27); +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +lean_dec(x_31); +x_33 = l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_generate___spec__3(x_22, x_8, x_9, x_10, x_11, x_12, x_13, x_32); +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_22); +if (lean_obj_tag(x_33) == 0) +{ +uint8_t x_34; +x_34 = !lean_is_exclusive(x_33); +if (x_34 == 0) +{ +return x_33; } -x_545 = lean_box(0); -if (lean_is_scalar(x_544)) { - x_546 = lean_alloc_ctor(0, 2, 0); -} else { - x_546 = x_544; +else +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_33, 0); +x_36 = lean_ctor_get(x_33, 1); +lean_inc(x_36); +lean_inc(x_35); +lean_dec(x_33); +x_37 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_37, 0, x_35); +lean_ctor_set(x_37, 1, x_36); +return x_37; } -lean_ctor_set(x_546, 0, x_545); -lean_ctor_set(x_546, 1, x_543); -return x_546; } +else +{ +uint8_t x_38; +x_38 = !lean_is_exclusive(x_33); +if (x_38 == 0) +{ +return x_33; +} +else +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_33, 0); +x_40 = lean_ctor_get(x_33, 1); +lean_inc(x_40); +lean_inc(x_39); +lean_dec(x_33); +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 { -lean_object* x_547; lean_object* x_548; lean_object* x_549; lean_object* x_550; -lean_dec(x_408); -lean_dec(x_407); -lean_dec(x_405); -lean_dec(x_403); -lean_dec(x_397); -lean_dec(x_395); -lean_dec(x_394); -lean_dec(x_393); -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_547 = lean_ctor_get(x_409, 0); -lean_inc(x_547); -x_548 = lean_ctor_get(x_409, 1); -lean_inc(x_548); -if (lean_is_exclusive(x_409)) { - lean_ctor_release(x_409, 0); - lean_ctor_release(x_409, 1); - x_549 = x_409; +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_42 = lean_ctor_get(x_26, 0); +x_43 = lean_ctor_get(x_26, 1); +x_44 = lean_ctor_get(x_26, 2); +x_45 = lean_ctor_get(x_26, 3); +x_46 = lean_ctor_get(x_26, 4); +x_47 = lean_ctor_get(x_26, 5); +x_48 = lean_ctor_get(x_26, 6); +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_inc(x_42); +lean_dec(x_26); +x_49 = l_Lean_PersistentArray_append___rarg(x_16, x_45); +x_50 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_50, 0, x_42); +lean_ctor_set(x_50, 1, x_43); +lean_ctor_set(x_50, 2, x_44); +lean_ctor_set(x_50, 3, x_49); +lean_ctor_set(x_50, 4, x_46); +lean_ctor_set(x_50, 5, x_47); +lean_ctor_set(x_50, 6, x_48); +x_51 = lean_st_ref_set(x_13, x_50, x_27); +x_52 = lean_ctor_get(x_51, 1); +lean_inc(x_52); +lean_dec(x_51); +x_53 = l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_generate___spec__3(x_22, x_8, x_9, x_10, x_11, x_12, x_13, x_52); +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_22); +if (lean_obj_tag(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_53, 0); +lean_inc(x_54); +x_55 = lean_ctor_get(x_53, 1); +lean_inc(x_55); +if (lean_is_exclusive(x_53)) { + lean_ctor_release(x_53, 0); + lean_ctor_release(x_53, 1); + x_56 = x_53; } else { - lean_dec_ref(x_409); - x_549 = lean_box(0); + lean_dec_ref(x_53); + x_56 = lean_box(0); } -if (lean_is_scalar(x_549)) { - x_550 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_56)) { + x_57 = lean_alloc_ctor(0, 2, 0); } else { - x_550 = x_549; -} -lean_ctor_set(x_550, 0, x_547); -lean_ctor_set(x_550, 1, x_548); -return x_550; -} + x_57 = x_56; } +lean_ctor_set(x_57, 0, x_54); +lean_ctor_set(x_57, 1, x_55); +return x_57; } else { -lean_object* x_551; lean_object* x_552; lean_object* x_553; lean_object* x_554; -lean_dec(x_397); -lean_dec(x_395); -lean_dec(x_394); -lean_dec(x_393); -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_551 = lean_ctor_get(x_402, 0); -lean_inc(x_551); -x_552 = lean_ctor_get(x_402, 1); -lean_inc(x_552); -if (lean_is_exclusive(x_402)) { - lean_ctor_release(x_402, 0); - lean_ctor_release(x_402, 1); - x_553 = x_402; +lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_58 = lean_ctor_get(x_53, 0); +lean_inc(x_58); +x_59 = lean_ctor_get(x_53, 1); +lean_inc(x_59); +if (lean_is_exclusive(x_53)) { + lean_ctor_release(x_53, 0); + lean_ctor_release(x_53, 1); + x_60 = x_53; } else { - lean_dec_ref(x_402); - x_553 = lean_box(0); + lean_dec_ref(x_53); + x_60 = lean_box(0); } -if (lean_is_scalar(x_553)) { - x_554 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_60)) { + x_61 = lean_alloc_ctor(1, 2, 0); } else { - x_554 = x_553; -} -lean_ctor_set(x_554, 0, x_551); -lean_ctor_set(x_554, 1, x_552); -return x_554; + x_61 = x_60; } +lean_ctor_set(x_61, 0, x_58); +lean_ctor_set(x_61, 1, x_59); +return x_61; } } } -LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at_Lean_Meta_SynthInstance_consume___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: +else { -lean_object* x_9; -x_9 = l_Lean_MVarId_isAssigned___at_Lean_Meta_SynthInstance_consume___spec__1(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_object* x_62; double x_63; lean_object* x_64; +x_62 = lean_box(0); +x_63 = lean_unbox_float(x_23); +lean_dec(x_23); +x_64 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_generate___spec__1___lambda__2(x_2, x_22, x_16, x_3, x_4, x_5, x_24, x_63, x_62, x_8, x_9, x_10, x_11, x_12, x_13, x_21); +return x_64; } } -LEAN_EXPORT lean_object* l_List_filterAuxM___at_Lean_Meta_SynthInstance_consume___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: +else { -lean_object* x_10; -x_10 = l_List_filterAuxM___at_Lean_Meta_SynthInstance_consume___spec__2(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); -return x_10; +lean_object* x_65; double x_66; lean_object* x_67; +x_65 = lean_box(0); +x_66 = lean_unbox_float(x_23); +lean_dec(x_23); +x_67 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_generate___spec__1___lambda__2(x_2, x_22, x_16, x_3, x_4, x_5, x_24, x_66, x_65, x_8, x_9, x_10, x_11, x_12, x_13, x_21); +return x_67; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_SynthInstance_consume___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) { -_start: +} +else { -lean_object* x_12; -x_12 = l_Array_mapMUnsafe_map___at_Lean_Meta_SynthInstance_consume___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); -lean_dec(x_6); +uint8_t x_75; +lean_dec(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_5); lean_dec(x_3); -return x_12; -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_SynthInstance_consume___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) { -_start: +lean_dec(x_2); +x_75 = !lean_is_exclusive(x_19); +if (x_75 == 0) { -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 = l_Array_mapMUnsafe_map___at_Lean_Meta_SynthInstance_consume___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); -return x_15; -} +return x_19; } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_consume___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: +else { -size_t x_6; size_t x_7; lean_object* x_8; -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_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_consume___spec__4(x_1, x_2, x_6, x_7, x_5); -lean_dec(x_2); -return x_8; +lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_76 = lean_ctor_get(x_19, 0); +x_77 = lean_ctor_get(x_19, 1); +lean_inc(x_77); +lean_inc(x_76); +lean_dec(x_19); +x_78 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_78, 0, x_76); +lean_ctor_set(x_78, 1, x_77); +return x_78; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_consume___spec__5___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_3); -lean_dec(x_3); -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_consume___spec__5(x_1, x_2, x_6, x_7, x_5); -lean_dec(x_2); -return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_consume___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_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_generate___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) { _start: { -lean_object* x_9; -x_9 = l_Lean_Meta_SynthInstance_consume___lambda__1(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_Meta_SynthInstance_consume___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) { -_start: +lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_12 = lean_ctor_get(x_9, 2); +lean_inc(x_12); +lean_inc(x_1); +x_13 = l_Lean_isTracingEnabledFor___at_Lean_Meta_SynthInstance_newSubgoal___spec__10(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_unbox(x_14); +if (x_15 == 0) { -lean_object* x_9; -x_9 = l_Lean_Meta_SynthInstance_consume___lambda__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_3); +lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); +lean_dec(x_13); +x_17 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__4___closed__1; +x_18 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_12, x_17); +if (x_18 == 0) +{ +lean_object* x_19; +lean_dec(x_14); +lean_dec(x_12); lean_dec(x_2); -return x_9; -} +lean_dec(x_1); +x_19 = lean_apply_7(x_3, x_5, x_6, x_7, x_8, x_9, x_10, x_16); +if (lean_obj_tag(x_19) == 0) +{ +uint8_t x_20; +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) +{ +return x_19; } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_consume___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) { -_start: +else { -lean_object* x_9; -x_9 = l_Lean_Meta_SynthInstance_consume___lambda__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); -lean_dec(x_2); -return x_9; +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(0, 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_SynthInstance_getTop___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: +else { -lean_object* x_7; uint8_t x_8; -x_7 = lean_st_ref_get(x_1, x_6); -x_8 = !lean_is_exclusive(x_7); -if (x_8 == 0) +uint8_t x_24; +x_24 = !lean_is_exclusive(x_19); +if (x_24 == 0) { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_9 = lean_ctor_get(x_7, 0); -x_10 = lean_ctor_get(x_9, 1); -lean_inc(x_10); -lean_dec(x_9); -x_11 = l_Lean_Meta_SynthInstance_instInhabitedGeneratorNode; -x_12 = l_Array_back___rarg(x_11, x_10); -lean_dec(x_10); -lean_ctor_set(x_7, 0, x_12); -return x_7; +return x_19; } 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; -x_13 = lean_ctor_get(x_7, 0); -x_14 = lean_ctor_get(x_7, 1); -lean_inc(x_14); -lean_inc(x_13); -lean_dec(x_7); -x_15 = lean_ctor_get(x_13, 1); -lean_inc(x_15); -lean_dec(x_13); -x_16 = l_Lean_Meta_SynthInstance_instInhabitedGeneratorNode; -x_17 = l_Array_back___rarg(x_16, x_15); -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_14); -return x_18; +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_19, 0); +x_26 = lean_ctor_get(x_19, 1); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_19); +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_28; uint8_t x_29; lean_object* x_30; +x_28 = lean_box(0); +x_29 = lean_unbox(x_14); +lean_dec(x_14); +x_30 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_generate___spec__1___lambda__3(x_3, x_2, x_1, x_4, x_12, x_29, x_28, x_5, x_6, x_7, x_8, x_9, x_10, x_16); +return x_30; +} } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_getTop(lean_object* x_1) { -_start: +else { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_getTop___rarg___boxed), 6, 0); -return x_2; +lean_object* x_31; lean_object* x_32; uint8_t x_33; lean_object* x_34; +x_31 = lean_ctor_get(x_13, 1); +lean_inc(x_31); +lean_dec(x_13); +x_32 = lean_box(0); +x_33 = lean_unbox(x_14); +lean_dec(x_14); +x_34 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_generate___spec__1___lambda__3(x_3, x_2, x_1, x_4, x_12, x_33, x_32, x_5, x_6, x_7, x_8, x_9, x_10, x_31); +return x_34; } } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_getTop___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) { +} +static lean_object* _init_l_Lean_Meta_SynthInstance_generate___lambda__1___closed__1() { _start: { -lean_object* x_7; -x_7 = l_Lean_Meta_SynthInstance_getTop___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_object* x_1; +x_1 = lean_mk_string_from_bytes(" apply ", 7); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_getTop___boxed(lean_object* x_1) { +static lean_object* _init_l_Lean_Meta_SynthInstance_generate___lambda__1___closed__2() { _start: { -lean_object* x_2; -x_2 = l_Lean_Meta_SynthInstance_getTop(x_1); -lean_dec(x_1); +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_SynthInstance_generate___lambda__1___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_modifyTop(lean_object* x_1, lean_object* x_2, 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_Meta_SynthInstance_generate___lambda__1___closed__3() { _start: { -lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_9 = lean_st_ref_take(x_3, 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_is_exclusive(x_10); -if (x_12 == 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_10, 1); -x_14 = lean_array_get_size(x_13); -x_15 = lean_unsigned_to_nat(1u); -x_16 = lean_nat_sub(x_14, x_15); -x_17 = lean_nat_dec_lt(x_16, x_14); -lean_dec(x_14); -if (x_17 == 0) -{ -lean_object* x_18; uint8_t x_19; -lean_dec(x_16); -lean_dec(x_1); -x_18 = lean_st_ref_set(x_3, x_10, x_11); -x_19 = !lean_is_exclusive(x_18); -if (x_19 == 0) -{ -lean_object* x_20; lean_object* x_21; -x_20 = lean_ctor_get(x_18, 0); -lean_dec(x_20); -x_21 = lean_box(0); -lean_ctor_set(x_18, 0, x_21); -return x_18; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(" to ", 4); +return x_1; } -else +} +static lean_object* _init_l_Lean_Meta_SynthInstance_generate___lambda__1___closed__4() { +_start: { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_18, 1); -lean_inc(x_22); -lean_dec(x_18); -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_1; lean_object* x_2; +x_1 = l_Lean_Meta_SynthInstance_generate___lambda__1___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } -else +LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_generate___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_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_array_fget(x_13, x_16); -x_26 = lean_box(0); -x_27 = lean_array_fset(x_13, x_16, x_26); -x_28 = lean_apply_1(x_1, x_25); -x_29 = lean_array_fset(x_27, x_16, x_28); -lean_dec(x_16); -lean_ctor_set(x_10, 1, x_29); -x_30 = lean_st_ref_set(x_3, x_10, x_11); -x_31 = !lean_is_exclusive(x_30); -if (x_31 == 0) +lean_object* x_11; +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +x_11 = lean_infer_type(x_1, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_11) == 0) { -lean_object* x_32; -x_32 = lean_ctor_get(x_30, 0); -lean_dec(x_32); -lean_ctor_set(x_30, 0, x_26); -return x_30; -} -else +lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +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 = l_Lean_instantiateMVars___at_Lean_Meta_SynthInstance_mkTableKeyFor___spec__1(x_12, 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); +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) { -lean_object* x_33; lean_object* x_34; -x_33 = lean_ctor_get(x_30, 1); -lean_inc(x_33); -lean_dec(x_30); -x_34 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_34, 0, x_26); -lean_ctor_set(x_34, 1, x_33); -return x_34; -} -} +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_16 = lean_ctor_get(x_14, 0); +x_17 = l_Lean_exceptOptionEmoji___rarg(x_3); +x_18 = l_Lean_stringToMessageData(x_17); +lean_dec(x_17); +x_19 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; +x_20 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_20, 0, x_19); +lean_ctor_set(x_20, 1, x_18); +x_21 = l_Lean_Meta_SynthInstance_generate___lambda__1___closed__2; +x_22 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_22, 0, x_20); +lean_ctor_set(x_22, 1, x_21); +x_23 = lean_ctor_get(x_2, 0); +lean_inc(x_23); +lean_dec(x_2); +x_24 = l_Lean_MessageData_ofExpr(x_23); +x_25 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_25, 0, x_22); +lean_ctor_set(x_25, 1, x_24); +x_26 = l_Lean_Meta_SynthInstance_generate___lambda__1___closed__4; +x_27 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = l_Lean_MessageData_ofExpr(x_16); +x_29 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +x_30 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_19); +lean_ctor_set(x_14, 0, x_30); +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; uint8_t x_42; -x_35 = lean_ctor_get(x_10, 0); -x_36 = lean_ctor_get(x_10, 1); -x_37 = lean_ctor_get(x_10, 2); -x_38 = lean_ctor_get(x_10, 3); -lean_inc(x_38); -lean_inc(x_37); -lean_inc(x_36); -lean_inc(x_35); -lean_dec(x_10); -x_39 = lean_array_get_size(x_36); -x_40 = lean_unsigned_to_nat(1u); -x_41 = lean_nat_sub(x_39, x_40); -x_42 = lean_nat_dec_lt(x_41, x_39); -lean_dec(x_39); -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_dec(x_41); -lean_dec(x_1); -x_43 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_43, 0, x_35); -lean_ctor_set(x_43, 1, x_36); -lean_ctor_set(x_43, 2, x_37); -lean_ctor_set(x_43, 3, x_38); -x_44 = lean_st_ref_set(x_3, x_43, x_11); -x_45 = lean_ctor_get(x_44, 1); -lean_inc(x_45); -if (lean_is_exclusive(x_44)) { - lean_ctor_release(x_44, 0); - lean_ctor_release(x_44, 1); - x_46 = x_44; -} else { - lean_dec_ref(x_44); - x_46 = lean_box(0); -} -x_47 = lean_box(0); -if (lean_is_scalar(x_46)) { - x_48 = lean_alloc_ctor(0, 2, 0); -} else { - x_48 = x_46; -} -lean_ctor_set(x_48, 0, x_47); -lean_ctor_set(x_48, 1, x_45); -return x_48; +lean_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_31 = lean_ctor_get(x_14, 0); +x_32 = lean_ctor_get(x_14, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_14); +x_33 = l_Lean_exceptOptionEmoji___rarg(x_3); +x_34 = l_Lean_stringToMessageData(x_33); +lean_dec(x_33); +x_35 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; +x_36 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_34); +x_37 = l_Lean_Meta_SynthInstance_generate___lambda__1___closed__2; +x_38 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_38, 0, x_36); +lean_ctor_set(x_38, 1, x_37); +x_39 = lean_ctor_get(x_2, 0); +lean_inc(x_39); +lean_dec(x_2); +x_40 = l_Lean_MessageData_ofExpr(x_39); +x_41 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_41, 0, x_38); +lean_ctor_set(x_41, 1, x_40); +x_42 = l_Lean_Meta_SynthInstance_generate___lambda__1___closed__4; +x_43 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_43, 0, x_41); +lean_ctor_set(x_43, 1, x_42); +x_44 = l_Lean_MessageData_ofExpr(x_31); +x_45 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +x_46 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_35); +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_32); +return x_47; +} } 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; -x_49 = lean_array_fget(x_36, x_41); -x_50 = lean_box(0); -x_51 = lean_array_fset(x_36, x_41, x_50); -x_52 = lean_apply_1(x_1, x_49); -x_53 = lean_array_fset(x_51, x_41, x_52); -lean_dec(x_41); -x_54 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_54, 0, x_35); -lean_ctor_set(x_54, 1, x_53); -lean_ctor_set(x_54, 2, x_37); -lean_ctor_set(x_54, 3, x_38); -x_55 = lean_st_ref_set(x_3, x_54, x_11); -x_56 = lean_ctor_get(x_55, 1); -lean_inc(x_56); -if (lean_is_exclusive(x_55)) { - lean_ctor_release(x_55, 0); - lean_ctor_release(x_55, 1); - x_57 = x_55; -} else { - lean_dec_ref(x_55); - x_57 = lean_box(0); -} -if (lean_is_scalar(x_57)) { - x_58 = lean_alloc_ctor(0, 2, 0); -} else { - x_58 = x_57; +uint8_t x_48; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +x_48 = !lean_is_exclusive(x_11); +if (x_48 == 0) +{ +return x_11; } -lean_ctor_set(x_58, 0, x_50); -lean_ctor_set(x_58, 1, x_56); -return x_58; +else +{ +lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_49 = lean_ctor_get(x_11, 0); +x_50 = lean_ctor_get(x_11, 1); +lean_inc(x_50); +lean_inc(x_49); +lean_dec(x_11); +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_EXPORT lean_object* l_Lean_Meta_SynthInstance_modifyTop___boxed(lean_object* x_1, lean_object* x_2, lean_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_Meta_SynthInstance_generate___lambda__2(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_9; -x_9 = l_Lean_Meta_SynthInstance_modifyTop(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); +uint8_t x_3; +x_3 = !lean_is_exclusive(x_2); +if (x_3 == 0) +{ +lean_object* x_4; +x_4 = lean_ctor_get(x_2, 4); lean_dec(x_4); -lean_dec(x_3); +lean_ctor_set(x_2, 4, x_1); +return x_2; +} +else +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_5 = lean_ctor_get(x_2, 0); +x_6 = lean_ctor_get(x_2, 1); +x_7 = lean_ctor_get(x_2, 2); +x_8 = lean_ctor_get(x_2, 3); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); lean_dec(x_2); +x_9 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_9, 0, x_5); +lean_ctor_set(x_9, 1, x_6); +lean_ctor_set(x_9, 2, x_7); +lean_ctor_set(x_9, 3, x_8); +lean_ctor_set(x_9, 4, x_1); return x_9; } } -LEAN_EXPORT lean_object* l_Lean_withSeconds___at_Lean_Meta_SynthInstance_generate___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_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_9 = lean_io_mono_nanos_now(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_apply_7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_11); -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_io_mono_nanos_now(x_14); -x_16 = !lean_is_exclusive(x_15); -if (x_16 == 0) -{ -lean_object* x_17; lean_object* x_18; uint8_t x_19; lean_object* x_20; double x_21; double x_22; double x_23; lean_object* x_24; lean_object* x_25; -x_17 = lean_ctor_get(x_15, 0); -x_18 = lean_nat_sub(x_17, x_10); -lean_dec(x_10); -lean_dec(x_17); -x_19 = 0; -x_20 = lean_unsigned_to_nat(0u); -x_21 = l_Float_ofScientific(x_18, x_19, x_20); -lean_dec(x_18); -x_22 = l_Lean_withSeconds___at_Lean_Meta_SynthInstance_newSubgoal___spec__14___closed__1; -x_23 = lean_float_div(x_21, x_22); -x_24 = lean_box_float(x_23); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_13); -lean_ctor_set(x_25, 1, x_24); -lean_ctor_set(x_15, 0, x_25); -return x_15; } -else +LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_generate___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) { +_start: { -lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; lean_object* x_30; double x_31; double x_32; double x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_26 = lean_ctor_get(x_15, 0); -x_27 = lean_ctor_get(x_15, 1); -lean_inc(x_27); -lean_inc(x_26); -lean_dec(x_15); -x_28 = lean_nat_sub(x_26, x_10); -lean_dec(x_10); -lean_dec(x_26); -x_29 = 0; -x_30 = lean_unsigned_to_nat(0u); -x_31 = l_Float_ofScientific(x_28, x_29, x_30); -lean_dec(x_28); -x_32 = l_Lean_withSeconds___at_Lean_Meta_SynthInstance_newSubgoal___spec__14___closed__1; -x_33 = lean_float_div(x_31, x_32); -x_34 = lean_box_float(x_33); -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_13); -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_27); -return x_36; -} +lean_object* x_9; lean_object* x_10; +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; } -else -{ -uint8_t x_37; -lean_dec(x_10); -x_37 = !lean_is_exclusive(x_12); -if (x_37 == 0) -{ -return x_12; } -else +static lean_object* _init_l_Lean_Meta_SynthInstance_generate___lambda__4___closed__1() { +_start: { -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_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; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_generate___lambda__3___boxed), 8, 0); +return x_1; } } +static lean_object* _init_l_Lean_Meta_SynthInstance_generate___lambda__4___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +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_withOptProfile___at_Lean_Meta_SynthInstance_generate___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_EXPORT lean_object* l_Lean_Meta_SynthInstance_generate___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) { _start: { -lean_object* x_9; uint8_t x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_9 = l_Lean_withOptProfile___at_Lean_Meta_SynthInstance_newSubgoal___spec__12___closed__2; -x_10 = 0; -x_11 = l_Lean_getBoolOption___at_Lean_Meta_SynthInstance_newSubgoal___spec__13(x_9, x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_unbox(x_12); -lean_dec(x_12); -if (x_13 == 0) -{ -lean_object* x_14; lean_object* x_15; -x_14 = lean_ctor_get(x_11, 1); -lean_inc(x_14); -lean_dec(x_11); -x_15 = lean_apply_7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_14); -if (lean_obj_tag(x_15) == 0) +lean_object* x_12; +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_1); +x_12 = l_Lean_Meta_SynthInstance_tryResolve(x_1, x_2, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_12) == 0) { -uint8_t x_16; -x_16 = !lean_is_exclusive(x_15); -if (x_16 == 0) +lean_object* x_13; +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +if (lean_obj_tag(x_13) == 0) { -lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_17 = lean_ctor_get(x_15, 0); -x_18 = lean_box(0); -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_17); -lean_ctor_set(x_19, 1, x_18); -lean_ctor_set(x_15, 0, x_19); -return x_15; +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +lean_dec(x_3); +lean_dec(x_1); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = l_Lean_Meta_SynthInstance_generate___lambda__4___closed__1; +x_16 = lean_box(0); +x_17 = lean_apply_8(x_15, x_16, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +return x_17; } 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_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_13, 0); +lean_inc(x_18); +lean_dec(x_13); +x_19 = lean_ctor_get(x_12, 1); +lean_inc(x_19); +lean_dec(x_12); +x_20 = lean_ctor_get(x_18, 0); lean_inc(x_20); -lean_dec(x_15); -x_22 = lean_box(0); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_20); -lean_ctor_set(x_23, 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_21); -return x_24; -} -} -else +x_21 = lean_ctor_get(x_18, 1); +lean_inc(x_21); +lean_dec(x_18); +x_22 = lean_unsigned_to_nat(0u); +x_23 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_23, 0, x_1); +lean_ctor_set(x_23, 1, x_3); +lean_ctor_set(x_23, 2, x_20); +lean_ctor_set(x_23, 3, x_21); +lean_ctor_set(x_23, 4, x_22); +x_24 = l_Lean_Meta_SynthInstance_consume(x_23, x_5, x_6, x_7, x_8, x_9, x_10, x_19); +if (lean_obj_tag(x_24) == 0) { uint8_t x_25; -x_25 = !lean_is_exclusive(x_15); +x_25 = !lean_is_exclusive(x_24); if (x_25 == 0) { -return x_15; +lean_object* x_26; lean_object* x_27; +x_26 = lean_ctor_get(x_24, 0); +lean_dec(x_26); +x_27 = l_Lean_Meta_SynthInstance_generate___lambda__4___closed__2; +lean_ctor_set(x_24, 0, x_27); +return x_24; } 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_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_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_24, 1); +lean_inc(x_28); +lean_dec(x_24); +x_29 = l_Lean_Meta_SynthInstance_generate___lambda__4___closed__2; +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; } } else { -lean_object* x_29; lean_object* x_30; -x_29 = lean_ctor_get(x_11, 1); -lean_inc(x_29); -lean_dec(x_11); -x_30 = l_Lean_withSeconds___at_Lean_Meta_SynthInstance_generate___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_29); -if (lean_obj_tag(x_30) == 0) -{ uint8_t x_31; -x_31 = !lean_is_exclusive(x_30); +x_31 = !lean_is_exclusive(x_24); if (x_31 == 0) { -lean_object* x_32; uint8_t x_33; -x_32 = lean_ctor_get(x_30, 0); -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, 1); -x_35 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_35, 0, x_34); -lean_ctor_set(x_32, 1, x_35); -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 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_38, 0, x_37); -x_39 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_39, 0, x_36); -lean_ctor_set(x_39, 1, x_38); -lean_ctor_set(x_30, 0, x_39); -return x_30; -} +return x_24; } 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; -x_40 = lean_ctor_get(x_30, 0); -x_41 = lean_ctor_get(x_30, 1); -lean_inc(x_41); -lean_inc(x_40); -lean_dec(x_30); -x_42 = lean_ctor_get(x_40, 0); -lean_inc(x_42); -x_43 = lean_ctor_get(x_40, 1); -lean_inc(x_43); -if (lean_is_exclusive(x_40)) { - lean_ctor_release(x_40, 0); - lean_ctor_release(x_40, 1); - x_44 = x_40; -} else { - lean_dec_ref(x_40); - x_44 = lean_box(0); +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_24, 0); +x_33 = lean_ctor_get(x_24, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_24); +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; } -x_45 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_45, 0, x_43); -if (lean_is_scalar(x_44)) { - x_46 = lean_alloc_ctor(0, 2, 0); -} else { - x_46 = x_44; } -lean_ctor_set(x_46, 0, x_42); -lean_ctor_set(x_46, 1, 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_41); -return x_47; } } else { -uint8_t x_48; -x_48 = !lean_is_exclusive(x_30); -if (x_48 == 0) +uint8_t x_35; +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_35 = !lean_is_exclusive(x_12); +if (x_35 == 0) { -return x_30; +return x_12; } else { -lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_49 = lean_ctor_get(x_30, 0); -x_50 = lean_ctor_get(x_30, 1); -lean_inc(x_50); -lean_inc(x_49); -lean_dec(x_30); -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_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_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; } } } } -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_generate___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_Lean_Meta_SynthInstance_generate(lean_object* x_1, 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: { -if (lean_obj_tag(x_1) == 0) -{ -lean_object* x_9; lean_object* x_10; -x_9 = lean_ctor_get(x_1, 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; uint8_t x_17; +x_8 = l_Lean_Meta_SynthInstance_getTop___rarg(x_2, x_3, x_4, x_5, x_6, x_7); +x_9 = lean_ctor_get(x_8, 0); lean_inc(x_9); -x_10 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_10, 0, x_9); -lean_ctor_set(x_10, 1, x_8); -return x_10; -} -else -{ -lean_object* x_11; lean_object* x_12; -x_11 = lean_ctor_get(x_1, 0); +x_10 = lean_ctor_get(x_8, 1); +lean_inc(x_10); +lean_dec(x_8); +x_11 = lean_ctor_get(x_9, 0); lean_inc(x_11); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_11); -lean_ctor_set(x_12, 1, x_8); -return x_12; -} -} -} -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_generate___spec__1___lambda__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) { -_start: -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_12 = lean_ctor_get(x_9, 1); lean_inc(x_12); -x_15 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__15(x_1, x_2, x_3, x_6, x_4, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -x_16 = lean_ctor_get(x_15, 1); -lean_inc(x_16); -lean_dec(x_15); -x_17 = l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_generate___spec__4(x_5, x_8, x_9, x_10, x_11, x_12, x_13, x_16); -lean_dec(x_12); -lean_dec(x_8); -return x_17; -} -} -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_generate___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) { -_start: -{ -lean_object* x_12; lean_object* x_13; uint8_t x_14; -lean_inc(x_1); -x_12 = l_Lean_isTracingEnabledFor___at_Lean_Meta_SynthInstance_newSubgoal___spec__10(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -x_13 = lean_ctor_get(x_12, 0); +x_13 = lean_ctor_get(x_9, 2); lean_inc(x_13); -x_14 = lean_unbox(x_13); -lean_dec(x_13); -if (x_14 == 0) -{ -lean_object* x_15; lean_object* x_16; -lean_dec(x_2); -lean_dec(x_1); -x_15 = lean_ctor_get(x_12, 1); +x_14 = lean_ctor_get(x_9, 3); +lean_inc(x_14); +x_15 = lean_ctor_get(x_9, 4); lean_inc(x_15); -lean_dec(x_12); -x_16 = lean_apply_7(x_3, x_5, x_6, x_7, x_8, x_9, x_10, x_15); -return x_16; -} -else +lean_dec(x_9); +x_16 = lean_unsigned_to_nat(0u); +x_17 = lean_nat_dec_eq(x_15, 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; -x_17 = lean_ctor_get(x_12, 1); -lean_inc(x_17); -lean_dec(x_12); -x_18 = lean_ctor_get(x_9, 5); -lean_inc(x_18); -x_19 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_SynthInstance_newSubgoal___spec__11___rarg(x_10, x_17); -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; uint8_t x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_18 = lean_unsigned_to_nat(1u); +x_19 = lean_nat_sub(x_15, x_18); +lean_dec(x_15); +x_20 = l_Lean_Meta_SynthInstance_instInhabitedInstance; +x_21 = lean_array_get(x_20, x_14, x_19); +lean_dec(x_14); lean_inc(x_21); -lean_dec(x_19); -x_22 = lean_alloc_closure((void*)(l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__1), 8, 1); -lean_closure_set(x_22, 0, x_3); -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_23 = l_Lean_withOptProfile___at_Lean_Meta_SynthInstance_generate___spec__2(x_22, x_5, x_6, x_7, x_8, x_9, x_10, x_21); -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; -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_24, 0); -lean_inc(x_26); -x_27 = lean_ctor_get(x_24, 1); -lean_inc(x_27); -lean_dec(x_24); -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_26); -x_28 = lean_apply_8(x_2, x_26, x_5, x_6, x_7, x_8, x_9, x_10, x_25); -if (lean_obj_tag(x_28) == 0) +lean_inc(x_11); +x_22 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_generate___lambda__1___boxed), 10, 2); +lean_closure_set(x_22, 0, x_11); +lean_closure_set(x_22, 1, x_21); +x_23 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_generate___lambda__2), 2, 1); +lean_closure_set(x_23, 0, x_19); +x_24 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_modifyTop___boxed), 8, 1); +lean_closure_set(x_24, 0, x_23); +x_25 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_generate___lambda__4___boxed), 11, 3); +lean_closure_set(x_25, 0, x_11); +lean_closure_set(x_25, 1, x_21); +lean_closure_set(x_25, 2, x_12); +x_26 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_SynthInstance_newSubgoal___spec__7___rarg), 9, 2); +lean_closure_set(x_26, 0, x_24); +lean_closure_set(x_26, 1, x_25); +x_27 = l_Lean_Meta_SynthInstance_newSubgoal___closed__1; +x_28 = 1; +x_29 = lean_box(x_28); +x_30 = lean_alloc_closure((void*)(l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_generate___spec__1___boxed), 11, 4); +lean_closure_set(x_30, 0, x_27); +lean_closure_set(x_30, 1, x_22); +lean_closure_set(x_30, 2, x_26); +lean_closure_set(x_30, 3, x_29); +x_31 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__16___rarg(x_13, x_30, x_1, x_2, x_3, x_4, x_5, x_6, x_10); +if (lean_obj_tag(x_31) == 0) { -if (lean_obj_tag(x_27) == 0) +uint8_t x_32; +x_32 = !lean_is_exclusive(x_31); +if (x_32 == 0) { -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; -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_box(0); -x_32 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_generate___spec__1___lambda__1(x_20, x_1, x_18, x_4, x_26, x_29, x_31, x_5, x_6, x_7, x_8, x_9, x_10, x_30); -lean_dec(x_10); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_26); -return x_32; +lean_object* x_33; lean_object* x_34; +x_33 = lean_ctor_get(x_31, 0); +lean_dec(x_33); +x_34 = lean_box(0); +lean_ctor_set(x_31, 0, x_34); +return x_31; } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; double x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; 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_28, 0); -lean_inc(x_33); -x_34 = lean_ctor_get(x_28, 1); -lean_inc(x_34); -lean_dec(x_28); -x_35 = lean_ctor_get(x_27, 0); +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_31, 1); lean_inc(x_35); -lean_dec(x_27); -x_36 = lean_unbox_float(x_35); -lean_dec(x_35); -x_37 = lean_float_to_string(x_36); -x_38 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_38, 0, x_37); -x_39 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_39, 0, x_38); -x_40 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___closed__2; -x_41 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_41, 0, x_40); -lean_ctor_set(x_41, 1, x_39); -x_42 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___closed__4; -x_43 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_43, 0, x_41); -lean_ctor_set(x_43, 1, x_42); -x_44 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_44, 0, x_43); -lean_ctor_set(x_44, 1, x_33); -x_45 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; -x_46 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -x_47 = lean_box(0); -x_48 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_generate___spec__1___lambda__1(x_20, x_1, x_18, x_4, x_26, x_46, x_47, x_5, x_6, x_7, x_8, x_9, x_10, x_34); -lean_dec(x_10); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_26); -return x_48; +lean_dec(x_31); +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_35); +return x_37; } } else { -uint8_t x_49; -lean_dec(x_27); -lean_dec(x_26); -lean_dec(x_20); -lean_dec(x_18); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); +uint8_t x_38; +x_38 = !lean_is_exclusive(x_31); +if (x_38 == 0) +{ +return x_31; +} +else +{ +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_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; +} +} +} +else +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); lean_dec(x_6); lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); lean_dec(x_1); -x_49 = !lean_is_exclusive(x_28); +x_42 = lean_st_ref_take(x_2, x_10); +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_is_exclusive(x_43); +if (x_45 == 0) +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; +x_46 = lean_ctor_get(x_43, 1); +x_47 = lean_array_pop(x_46); +lean_ctor_set(x_43, 1, x_47); +x_48 = lean_st_ref_set(x_2, x_43, x_44); +lean_dec(x_2); +x_49 = !lean_is_exclusive(x_48); if (x_49 == 0) { -return x_28; +lean_object* x_50; lean_object* x_51; +x_50 = lean_ctor_get(x_48, 0); +lean_dec(x_50); +x_51 = lean_box(0); +lean_ctor_set(x_48, 0, x_51); +return x_48; } else { -lean_object* x_50; lean_object* x_51; lean_object* x_52; -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, 2, 0); -lean_ctor_set(x_52, 0, x_50); -lean_ctor_set(x_52, 1, x_51); -return x_52; -} +lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_52 = lean_ctor_get(x_48, 1); +lean_inc(x_52); +lean_dec(x_48); +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_52); +return x_54; } } else { -uint8_t x_53; -lean_dec(x_20); -lean_dec(x_18); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +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_55 = lean_ctor_get(x_43, 0); +x_56 = lean_ctor_get(x_43, 1); +x_57 = lean_ctor_get(x_43, 2); +x_58 = lean_ctor_get(x_43, 3); +lean_inc(x_58); +lean_inc(x_57); +lean_inc(x_56); +lean_inc(x_55); +lean_dec(x_43); +x_59 = lean_array_pop(x_56); +x_60 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_60, 0, x_55); +lean_ctor_set(x_60, 1, x_59); +lean_ctor_set(x_60, 2, x_57); +lean_ctor_set(x_60, 3, x_58); +x_61 = lean_st_ref_set(x_2, x_60, x_44); +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); +} +x_64 = lean_box(0); +if (lean_is_scalar(x_63)) { + x_65 = lean_alloc_ctor(0, 2, 0); +} else { + x_65 = x_63; +} +lean_ctor_set(x_65, 0, x_64); +lean_ctor_set(x_65, 1, x_62); +return x_65; +} +} +} +} +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_generate___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_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_generate___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); lean_dec(x_2); lean_dec(x_1); -x_53 = !lean_is_exclusive(x_23); -if (x_53 == 0) +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_generate___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) { +_start: { -return x_23; +uint8_t x_15; lean_object* x_16; +x_15 = lean_unbox(x_4); +lean_dec(x_4); +x_16 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_generate___spec__1___lambda__1(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); +lean_dec(x_13); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_7); +lean_dec(x_5); +return x_16; } -else +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_generate___spec__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, 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_54; lean_object* x_55; lean_object* x_56; -x_54 = lean_ctor_get(x_23, 0); -x_55 = lean_ctor_get(x_23, 1); -lean_inc(x_55); -lean_inc(x_54); -lean_dec(x_23); -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; +uint8_t x_17; uint8_t x_18; double x_19; lean_object* x_20; +x_17 = lean_unbox(x_5); +lean_dec(x_5); +x_18 = lean_unbox(x_7); +lean_dec(x_7); +x_19 = lean_unbox_float(x_8); +lean_dec(x_8); +x_20 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_generate___spec__1___lambda__2(x_1, x_2, x_3, x_4, x_17, x_6, x_18, x_19, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +return x_20; +} } +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_generate___spec__1___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) { +_start: +{ +uint8_t x_15; uint8_t x_16; lean_object* x_17; +x_15 = lean_unbox(x_4); +lean_dec(x_4); +x_16 = lean_unbox(x_6); +lean_dec(x_6); +x_17 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_generate___spec__1___lambda__3(x_1, x_2, x_3, x_15, x_5, x_16, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +return x_17; } } +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_generate___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) { +_start: +{ +uint8_t x_12; lean_object* x_13; +x_12 = lean_unbox(x_4); +lean_dec(x_4); +x_13 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_generate___spec__1(x_1, x_2, x_3, x_12, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_13; } } -static lean_object* _init_l_Lean_Meta_SynthInstance_generate___lambda__1___closed__1() { +LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_generate___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_1; -x_1 = lean_mk_string_from_bytes(" apply ", 7); -return x_1; +lean_object* x_11; +x_11 = l_Lean_Meta_SynthInstance_generate___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_5); +lean_dec(x_4); +lean_dec(x_3); +return x_11; } } -static lean_object* _init_l_Lean_Meta_SynthInstance_generate___lambda__1___closed__2() { +LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_generate___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) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_SynthInstance_generate___lambda__1___closed__1; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +lean_object* x_9; +x_9 = l_Lean_Meta_SynthInstance_generate___lambda__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); +lean_dec(x_2); +lean_dec(x_1); +return x_9; } } -static lean_object* _init_l_Lean_Meta_SynthInstance_generate___lambda__1___closed__3() { +LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_generate___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) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes(" to ", 4); -return x_1; +lean_object* x_12; +x_12 = l_Lean_Meta_SynthInstance_generate___lambda__4(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; } } -static lean_object* _init_l_Lean_Meta_SynthInstance_generate___lambda__1___closed__4() { +static lean_object* _init_l_Lean_Meta_SynthInstance_getNextToResume___rarg___closed__1() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_SynthInstance_generate___lambda__1___closed__3; -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_Lean_Meta_SynthInstance_instInhabitedConsumerNode; +x_2 = l_Lean_Meta_SynthInstance_instInhabitedAnswer; +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_SynthInstance_generate___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_Meta_SynthInstance_getNextToResume___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_11; -lean_inc(x_9); +lean_object* x_7; 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; uint8_t x_16; +x_7 = lean_st_ref_get(x_1, x_6); +x_8 = lean_ctor_get(x_7, 0); lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -x_11 = lean_infer_type(x_1, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_11) == 0) -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; -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 = l_Lean_instantiateMVars___at_Lean_Meta_SynthInstance_mkTableKeyFor___spec__1(x_12, x_4, x_5, x_6, x_7, x_8, x_9, x_13); -lean_dec(x_9); -lean_dec(x_8); +x_9 = lean_ctor_get(x_7, 1); +lean_inc(x_9); lean_dec(x_7); -lean_dec(x_6); -x_15 = !lean_is_exclusive(x_14); -if (x_15 == 0) +x_10 = lean_ctor_get(x_8, 2); +lean_inc(x_10); +lean_dec(x_8); +x_11 = l_Lean_Meta_SynthInstance_getNextToResume___rarg___closed__1; +x_12 = l_Array_back___rarg(x_11, x_10); +lean_dec(x_10); +x_13 = lean_st_ref_take(x_1, 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 = !lean_is_exclusive(x_14); +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; -x_16 = lean_ctor_get(x_14, 0); -x_17 = l_Lean_exceptOptionEmoji___rarg(x_3); -x_18 = l_Lean_stringToMessageData(x_17); -lean_dec(x_17); -x_19 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; -x_20 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_18); -x_21 = l_Lean_Meta_SynthInstance_generate___lambda__1___closed__2; -x_22 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_22, 0, x_20); -lean_ctor_set(x_22, 1, x_21); -x_23 = lean_ctor_get(x_2, 0); -lean_inc(x_23); -lean_dec(x_2); -x_24 = l_Lean_MessageData_ofExpr(x_23); -x_25 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_25, 0, x_22); -lean_ctor_set(x_25, 1, x_24); -x_26 = l_Lean_Meta_SynthInstance_generate___lambda__1___closed__4; -x_27 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_27, 0, x_25); -lean_ctor_set(x_27, 1, x_26); -x_28 = l_Lean_MessageData_ofExpr(x_16); -x_29 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_29, 0, x_27); -lean_ctor_set(x_29, 1, x_28); -x_30 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_30, 0, x_29); -lean_ctor_set(x_30, 1, x_19); -lean_ctor_set(x_14, 0, x_30); -return x_14; +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_17 = lean_ctor_get(x_14, 2); +x_18 = lean_array_pop(x_17); +lean_ctor_set(x_14, 2, x_18); +x_19 = lean_st_ref_set(x_1, x_14, x_15); +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) +{ +lean_object* x_21; +x_21 = lean_ctor_get(x_19, 0); +lean_dec(x_21); +lean_ctor_set(x_19, 0, x_12); +return x_19; } 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; lean_object* x_40; lean_object* x_41; 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_31 = lean_ctor_get(x_14, 0); -x_32 = lean_ctor_get(x_14, 1); -lean_inc(x_32); -lean_inc(x_31); -lean_dec(x_14); -x_33 = l_Lean_exceptOptionEmoji___rarg(x_3); -x_34 = l_Lean_stringToMessageData(x_33); -lean_dec(x_33); -x_35 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; -x_36 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_36, 0, x_35); -lean_ctor_set(x_36, 1, x_34); -x_37 = l_Lean_Meta_SynthInstance_generate___lambda__1___closed__2; -x_38 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_38, 0, x_36); -lean_ctor_set(x_38, 1, x_37); -x_39 = lean_ctor_get(x_2, 0); -lean_inc(x_39); -lean_dec(x_2); -x_40 = l_Lean_MessageData_ofExpr(x_39); -x_41 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_41, 0, x_38); -lean_ctor_set(x_41, 1, x_40); -x_42 = l_Lean_Meta_SynthInstance_generate___lambda__1___closed__4; -x_43 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_43, 0, x_41); -lean_ctor_set(x_43, 1, x_42); -x_44 = l_Lean_MessageData_ofExpr(x_31); -x_45 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_45, 0, x_43); -lean_ctor_set(x_45, 1, x_44); -x_46 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_46, 0, x_45); -lean_ctor_set(x_46, 1, x_35); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_32); -return x_47; +lean_object* x_22; lean_object* x_23; +x_22 = lean_ctor_get(x_19, 1); +lean_inc(x_22); +lean_dec(x_19); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_12); +lean_ctor_set(x_23, 1, x_22); +return x_23; } } else { -uint8_t x_48; -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_2); -x_48 = !lean_is_exclusive(x_11); -if (x_48 == 0) -{ -return x_11; +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; 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_24 = lean_ctor_get(x_14, 0); +x_25 = lean_ctor_get(x_14, 1); +x_26 = lean_ctor_get(x_14, 2); +x_27 = lean_ctor_get(x_14, 3); +lean_inc(x_27); +lean_inc(x_26); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_14); +x_28 = lean_array_pop(x_26); +x_29 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_29, 0, x_24); +lean_ctor_set(x_29, 1, x_25); +lean_ctor_set(x_29, 2, x_28); +lean_ctor_set(x_29, 3, x_27); +x_30 = lean_st_ref_set(x_1, x_29, x_15); +x_31 = lean_ctor_get(x_30, 1); +lean_inc(x_31); +if (lean_is_exclusive(x_30)) { + lean_ctor_release(x_30, 0); + lean_ctor_release(x_30, 1); + x_32 = x_30; +} else { + lean_dec_ref(x_30); + x_32 = lean_box(0); } -else -{ -lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_49 = lean_ctor_get(x_11, 0); -x_50 = lean_ctor_get(x_11, 1); -lean_inc(x_50); -lean_inc(x_49); -lean_dec(x_11); -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; +if (lean_is_scalar(x_32)) { + x_33 = lean_alloc_ctor(0, 2, 0); +} else { + x_33 = x_32; } +lean_ctor_set(x_33, 0, x_12); +lean_ctor_set(x_33, 1, x_31); +return x_33; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_generate___lambda__2(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_getNextToResume(lean_object* x_1) { _start: { -uint8_t x_3; -x_3 = !lean_is_exclusive(x_2); -if (x_3 == 0) -{ -lean_object* x_4; -x_4 = lean_ctor_get(x_2, 4); -lean_dec(x_4); -lean_ctor_set(x_2, 4, x_1); +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_getNextToResume___rarg___boxed), 6, 0); return x_2; } -else +} +LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_getNextToResume___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_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; -x_5 = lean_ctor_get(x_2, 0); -x_6 = lean_ctor_get(x_2, 1); -x_7 = lean_ctor_get(x_2, 2); -x_8 = lean_ctor_get(x_2, 3); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); +lean_object* x_7; +x_7 = l_Lean_Meta_SynthInstance_getNextToResume___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); -x_9 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_9, 0, x_5); -lean_ctor_set(x_9, 1, x_6); -lean_ctor_set(x_9, 2, x_7); -lean_ctor_set(x_9, 3, x_8); -lean_ctor_set(x_9, 4, x_1); -return x_9; +lean_dec(x_1); +return x_7; } } +LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_getNextToResume___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_Meta_SynthInstance_getNextToResume(x_1); +lean_dec(x_1); +return x_2; +} } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_generate___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_EXPORT lean_object* l_panic___at_Lean_Meta_SynthInstance_resume___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_9; lean_object* x_10; -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_9; lean_object* x_10; lean_object* x_11; +x_9 = l_panic___at_Lean_Meta_SynthInstance_getEntry___spec__1___closed__1; +x_10 = lean_panic_fn(x_9, x_1); +x_11 = lean_apply_7(x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +return x_11; } } -static lean_object* _init_l_Lean_Meta_SynthInstance_generate___lambda__4___closed__1() { +static lean_object* _init_l_Lean_Meta_SynthInstance_resume___lambda__1___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_generate___lambda__3___boxed), 8, 0); +x_1 = lean_mk_string_from_bytes("propagating ", 12); return x_1; } } -static lean_object* _init_l_Lean_Meta_SynthInstance_generate___lambda__4___closed__2() { +static lean_object* _init_l_Lean_Meta_SynthInstance_resume___lambda__1___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = lean_box(0); -x_2 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_2, 0, x_1); +x_1 = l_Lean_Meta_SynthInstance_resume___lambda__1___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_generate___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) { +static lean_object* _init_l_Lean_Meta_SynthInstance_resume___lambda__1___closed__3() { _start: { -lean_object* x_12; -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_1); -x_12 = l_Lean_Meta_SynthInstance_tryResolve(x_1, x_2, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_12) == 0) -{ -lean_object* x_13; -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -if (lean_obj_tag(x_13) == 0) -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; -lean_dec(x_3); -lean_dec(x_1); -x_14 = lean_ctor_get(x_12, 1); -lean_inc(x_14); -lean_dec(x_12); -x_15 = l_Lean_Meta_SynthInstance_generate___lambda__4___closed__1; -x_16 = lean_box(0); -x_17 = lean_apply_8(x_15, x_16, x_5, x_6, x_7, x_8, x_9, x_10, x_14); -return x_17; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(" to subgoal ", 12); +return x_1; } -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_13, 0); -lean_inc(x_18); -lean_dec(x_13); -x_19 = lean_ctor_get(x_12, 1); -lean_inc(x_19); -lean_dec(x_12); -x_20 = lean_ctor_get(x_18, 0); -lean_inc(x_20); -x_21 = lean_ctor_get(x_18, 1); -lean_inc(x_21); -lean_dec(x_18); -x_22 = lean_unsigned_to_nat(0u); -x_23 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_23, 0, x_1); -lean_ctor_set(x_23, 1, x_3); -lean_ctor_set(x_23, 2, x_20); -lean_ctor_set(x_23, 3, x_21); -lean_ctor_set(x_23, 4, x_22); -x_24 = l_Lean_Meta_SynthInstance_consume(x_23, x_5, x_6, x_7, x_8, x_9, x_10, x_19); -if (lean_obj_tag(x_24) == 0) -{ -uint8_t x_25; -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) -{ -lean_object* x_26; lean_object* x_27; -x_26 = lean_ctor_get(x_24, 0); -lean_dec(x_26); -x_27 = l_Lean_Meta_SynthInstance_generate___lambda__4___closed__2; -lean_ctor_set(x_24, 0, x_27); -return x_24; } -else +static lean_object* _init_l_Lean_Meta_SynthInstance_resume___lambda__1___closed__4() { +_start: { -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_24, 1); -lean_inc(x_28); -lean_dec(x_24); -x_29 = l_Lean_Meta_SynthInstance_generate___lambda__4___closed__2; -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_1; lean_object* x_2; +x_1 = l_Lean_Meta_SynthInstance_resume___lambda__1___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } -else -{ -uint8_t x_31; -x_31 = !lean_is_exclusive(x_24); -if (x_31 == 0) -{ -return x_24; } -else +static lean_object* _init_l_Lean_Meta_SynthInstance_resume___lambda__1___closed__5() { +_start: { -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_24, 0); -x_33 = lean_ctor_get(x_24, 1); -lean_inc(x_33); -lean_inc(x_32); -lean_dec(x_24); -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_object* x_1; +x_1 = lean_mk_string_from_bytes(" of ", 4); +return x_1; } -else -{ -uint8_t x_35; -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_35 = !lean_is_exclusive(x_12); -if (x_35 == 0) -{ -return x_12; } -else +static lean_object* _init_l_Lean_Meta_SynthInstance_resume___lambda__1___closed__6() { +_start: { -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_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; -} -} +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_SynthInstance_resume___lambda__1___closed__5; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_generate(lean_object* x_1, lean_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_Meta_SynthInstance_resume___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_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 = l_Lean_Meta_SynthInstance_getTop___rarg(x_2, x_3, x_4, x_5, x_6, 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_dec(x_8); -x_11 = lean_ctor_get(x_9, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_9, 1); +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_11 = l_Lean_instantiateMVars___at_Lean_Meta_SynthInstance_mkTableKeyFor___spec__1(x_1, 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_9, 2); +x_13 = lean_ctor_get(x_11, 1); lean_inc(x_13); -x_14 = lean_ctor_get(x_9, 3); -lean_inc(x_14); -x_15 = lean_ctor_get(x_9, 4); -lean_inc(x_15); -lean_dec(x_9); -x_16 = lean_unsigned_to_nat(0u); -x_17 = lean_nat_dec_eq(x_15, 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; uint8_t x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_18 = lean_unsigned_to_nat(1u); -x_19 = lean_nat_sub(x_15, x_18); -lean_dec(x_15); -x_20 = l_Lean_Meta_SynthInstance_instInhabitedInstance; -x_21 = lean_array_get(x_20, x_14, x_19); -lean_dec(x_14); -lean_inc(x_21); -lean_inc(x_11); -x_22 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_generate___lambda__1___boxed), 10, 2); -lean_closure_set(x_22, 0, x_11); -lean_closure_set(x_22, 1, x_21); -x_23 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_generate___lambda__2), 2, 1); -lean_closure_set(x_23, 0, x_19); -x_24 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_modifyTop___boxed), 8, 1); -lean_closure_set(x_24, 0, x_23); -x_25 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_generate___lambda__4___boxed), 11, 3); -lean_closure_set(x_25, 0, x_11); -lean_closure_set(x_25, 1, x_21); -lean_closure_set(x_25, 2, x_12); -x_26 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_SynthInstance_newSubgoal___spec__7___rarg), 9, 2); -lean_closure_set(x_26, 0, x_24); -lean_closure_set(x_26, 1, x_25); -x_27 = l_Lean_Meta_SynthInstance_newSubgoal___closed__1; -x_28 = 1; -x_29 = lean_box(x_28); -x_30 = lean_alloc_closure((void*)(l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_generate___spec__1___boxed), 11, 4); -lean_closure_set(x_30, 0, x_27); -lean_closure_set(x_30, 1, x_22); -lean_closure_set(x_30, 2, x_26); -lean_closure_set(x_30, 3, x_29); -x_31 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__18___rarg(x_13, x_30, x_1, x_2, x_3, x_4, x_5, x_6, x_10); -if (lean_obj_tag(x_31) == 0) -{ -uint8_t x_32; -x_32 = !lean_is_exclusive(x_31); -if (x_32 == 0) +lean_dec(x_11); +x_14 = l_Lean_instantiateMVars___at_Lean_Meta_SynthInstance_mkTableKeyFor___spec__1(x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 0) { -lean_object* x_33; lean_object* x_34; -x_33 = lean_ctor_get(x_31, 0); -lean_dec(x_33); -x_34 = lean_box(0); -lean_ctor_set(x_31, 0, x_34); -return x_31; +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_14, 0); +x_17 = l_Lean_MessageData_ofExpr(x_3); +x_18 = l_Lean_Meta_SynthInstance_resume___lambda__1___closed__2; +x_19 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_17); +x_20 = l_Lean_Meta_SynthInstance_resume___lambda__1___closed__4; +x_21 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); +x_22 = l_Lean_MessageData_ofExpr(x_12); +x_23 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +x_24 = l_Lean_Meta_SynthInstance_resume___lambda__1___closed__6; +x_25 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +x_26 = l_Lean_MessageData_ofExpr(x_16); +x_27 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; +x_29 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +lean_ctor_set(x_14, 0, x_29); +return x_14; } else { -lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_35 = lean_ctor_get(x_31, 1); -lean_inc(x_35); -lean_dec(x_31); -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_35); -return x_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; lean_object* x_37; lean_object* x_38; lean_object* x_39; 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_30 = lean_ctor_get(x_14, 0); +x_31 = lean_ctor_get(x_14, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_14); +x_32 = l_Lean_MessageData_ofExpr(x_3); +x_33 = l_Lean_Meta_SynthInstance_resume___lambda__1___closed__2; +x_34 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_32); +x_35 = l_Lean_Meta_SynthInstance_resume___lambda__1___closed__4; +x_36 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +x_37 = l_Lean_MessageData_ofExpr(x_12); +x_38 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_38, 0, x_36); +lean_ctor_set(x_38, 1, x_37); +x_39 = l_Lean_Meta_SynthInstance_resume___lambda__1___closed__6; +x_40 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_40, 0, x_38); +lean_ctor_set(x_40, 1, x_39); +x_41 = l_Lean_MessageData_ofExpr(x_30); +x_42 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_42, 0, x_40); +lean_ctor_set(x_42, 1, x_41); +x_43 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; +x_44 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_44, 0, x_42); +lean_ctor_set(x_44, 1, 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_31); +return x_45; } } -else -{ -uint8_t x_38; -x_38 = !lean_is_exclusive(x_31); -if (x_38 == 0) -{ -return x_31; } -else +LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_resume___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_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_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; -} +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_1, 1); +lean_inc(x_13); +lean_dec(x_1); +x_14 = lean_alloc_closure((void*)(l_Lean_instantiateMVars___at_Lean_Meta_SynthInstance_mkTableKeyFor___spec__1___boxed), 8, 1); +lean_closure_set(x_14, 0, x_13); +x_15 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_resume___lambda__1___boxed), 10, 2); +lean_closure_set(x_15, 0, x_2); +lean_closure_set(x_15, 1, x_3); +x_16 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_SynthInstance_newSubgoal___spec__7___rarg), 9, 2); +lean_closure_set(x_16, 0, x_14); +lean_closure_set(x_16, 1, x_15); +x_17 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__16___rarg(x_4, x_16, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_17; } } -else +LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_resume___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) { +_start: { -lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -x_42 = lean_st_ref_take(x_2, x_10); -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_is_exclusive(x_43); -if (x_45 == 0) +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_15 = lean_ctor_get(x_1, 2); +x_16 = lean_nat_add(x_2, x_15); +x_17 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_17, 0, x_3); +lean_ctor_set(x_17, 1, x_4); +lean_ctor_set(x_17, 2, x_5); +lean_ctor_set(x_17, 3, x_6); +lean_ctor_set(x_17, 4, x_16); +x_18 = l_Lean_Meta_SynthInstance_consume(x_17, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_resume___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, uint8_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) { +_start: { -lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; -x_46 = lean_ctor_get(x_43, 1); -x_47 = lean_array_pop(x_46); -lean_ctor_set(x_43, 1, x_47); -x_48 = lean_st_ref_set(x_2, x_43, x_44); -lean_dec(x_2); -x_49 = !lean_is_exclusive(x_48); -if (x_49 == 0) +if (x_8 == 0) { -lean_object* x_50; lean_object* x_51; -x_50 = lean_ctor_get(x_48, 0); -lean_dec(x_50); -x_51 = lean_box(0); -lean_ctor_set(x_48, 0, x_51); -return x_48; +lean_object* x_16; lean_object* x_17; +lean_dec(x_7); +x_16 = lean_box(0); +x_17 = l_Lean_Meta_SynthInstance_resume___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_16, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +return x_17; } else { -lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_52 = lean_ctor_get(x_48, 1); -lean_inc(x_52); -lean_dec(x_48); -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_52); -return x_54; +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_1, 2); +x_19 = lean_nat_add(x_2, x_18); +x_20 = l_Nat_repr(x_19); +x_21 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_21, 0, x_20); +x_22 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_22, 0, x_21); +x_23 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___lambda__2___closed__4; +x_24 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_22); +x_25 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; +x_26 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +x_27 = l_Lean_addTrace___at_Lean_Meta_SynthInstance_wakeUp___spec__1(x_7, x_26, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +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_Meta_SynthInstance_resume___lambda__3(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; } } -else +} +static lean_object* _init_l_Lean_Meta_SynthInstance_resume___lambda__5___closed__1() { +_start: { -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_55 = lean_ctor_get(x_43, 0); -x_56 = lean_ctor_get(x_43, 1); -x_57 = lean_ctor_get(x_43, 2); -x_58 = lean_ctor_get(x_43, 3); -lean_inc(x_58); -lean_inc(x_57); -lean_inc(x_56); -lean_inc(x_55); -lean_dec(x_43); -x_59 = lean_array_pop(x_56); -x_60 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_60, 0, x_55); -lean_ctor_set(x_60, 1, x_59); -lean_ctor_set(x_60, 2, x_57); -lean_ctor_set(x_60, 3, x_58); -x_61 = lean_st_ref_set(x_2, x_60, x_44); -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); +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("resume", 6); +return x_1; } -x_64 = lean_box(0); -if (lean_is_scalar(x_63)) { - x_65 = lean_alloc_ctor(0, 2, 0); -} else { - x_65 = x_63; } -lean_ctor_set(x_65, 0, x_64); -lean_ctor_set(x_65, 1, x_62); -return x_65; +static lean_object* _init_l_Lean_Meta_SynthInstance_resume___lambda__5___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_6____closed__8; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_6____closed__1; +x_3 = l_Lean_Meta_SynthInstance_resume___lambda__5___closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} } +static lean_object* _init_l_Lean_Meta_SynthInstance_resume___lambda__5___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_SynthInstance_resume___lambda__5___closed__2; +x_2 = lean_alloc_closure((void*)(l_Lean_isTracingEnabledFor___at_Lean_Meta_SynthInstance_newSubgoal___spec__10___boxed), 8, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; } } +LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_resume___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_17; +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +x_17 = lean_infer_type(x_1, x_12, x_13, x_14, x_15, x_16); +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; lean_object* x_24; uint8_t x_25; lean_object* x_26; +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); +lean_inc(x_2); +x_20 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_resume___lambda__2___boxed), 12, 4); +lean_closure_set(x_20, 0, x_2); +lean_closure_set(x_20, 1, x_18); +lean_closure_set(x_20, 2, x_9); +lean_closure_set(x_20, 3, x_3); +x_21 = l_Lean_Meta_SynthInstance_resume___lambda__5___closed__2; +x_22 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_resume___lambda__4___boxed), 15, 7); +lean_closure_set(x_22, 0, x_2); +lean_closure_set(x_22, 1, x_4); +lean_closure_set(x_22, 2, x_5); +lean_closure_set(x_22, 3, x_6); +lean_closure_set(x_22, 4, x_7); +lean_closure_set(x_22, 5, x_8); +lean_closure_set(x_22, 6, x_21); +x_23 = l_Lean_Meta_SynthInstance_resume___lambda__5___closed__3; +x_24 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_SynthInstance_newSubgoal___spec__7___rarg), 9, 2); +lean_closure_set(x_24, 0, x_23); +lean_closure_set(x_24, 1, x_22); +x_25 = 1; +x_26 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_addAnswer___spec__3(x_21, x_20, x_24, x_25, x_10, x_11, x_12, x_13, x_14, x_15, x_19); +return x_26; } -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_generate___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: +else { -lean_object* x_9; -x_9 = l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_generate___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +uint8_t x_27; +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); 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_9; -} -} -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_generate___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) { -_start: +x_27 = !lean_is_exclusive(x_17); +if (x_27 == 0) { -uint8_t x_15; lean_object* x_16; -x_15 = lean_unbox(x_4); -lean_dec(x_4); -x_16 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_generate___spec__1___lambda__1(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); -lean_dec(x_13); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_7); -lean_dec(x_5); -return x_16; -} +return x_17; } -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_generate___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) { -_start: +else { -uint8_t x_12; lean_object* x_13; -x_12 = lean_unbox(x_4); -lean_dec(x_4); -x_13 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_generate___spec__1(x_1, x_2, x_3, x_12, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_13; +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_17, 0); +x_29 = lean_ctor_get(x_17, 1); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_17); +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_EXPORT lean_object* l_Lean_Meta_SynthInstance_generate___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_Lean_Meta_SynthInstance_generate___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_5); -lean_dec(x_4); -lean_dec(x_3); -return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_generate___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) { +static lean_object* _init_l_Lean_Meta_SynthInstance_resume___closed__1() { _start: { -lean_object* x_9; -x_9 = l_Lean_Meta_SynthInstance_generate___lambda__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); -lean_dec(x_2); -lean_dec(x_1); -return x_9; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Lean.Meta.SynthInstance.resume", 30); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_generate___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) { +static lean_object* _init_l_Lean_Meta_SynthInstance_resume___closed__2() { _start: { -lean_object* x_12; -x_12 = l_Lean_Meta_SynthInstance_generate___lambda__4(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; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("resume found no remaining subgoals", 34); +return x_1; } } -static lean_object* _init_l_Lean_Meta_SynthInstance_getNextToResume___rarg___closed__1() { +static lean_object* _init_l_Lean_Meta_SynthInstance_resume___closed__3() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_SynthInstance_instInhabitedConsumerNode; -x_2 = l_Lean_Meta_SynthInstance_instInhabitedAnswer; -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_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_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_getInstances___spec__7___closed__1; +x_2 = l_Lean_Meta_SynthInstance_resume___closed__1; +x_3 = lean_unsigned_to_nat(540u); +x_4 = lean_unsigned_to_nat(18u); +x_5 = l_Lean_Meta_SynthInstance_resume___closed__2; +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_Meta_SynthInstance_getNextToResume___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_EXPORT lean_object* l_Lean_Meta_SynthInstance_resume(lean_object* x_1, 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; 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; uint8_t x_16; -x_7 = lean_st_ref_get(x_1, x_6); -x_8 = lean_ctor_get(x_7, 0); -lean_inc(x_8); -x_9 = lean_ctor_get(x_7, 1); +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_8 = l_Lean_Meta_SynthInstance_getNextToResume___rarg(x_2, x_3, x_4, x_5, x_6, x_7); +x_9 = lean_ctor_get(x_8, 0); lean_inc(x_9); -lean_dec(x_7); -x_10 = lean_ctor_get(x_8, 2); +x_10 = lean_ctor_get(x_9, 0); lean_inc(x_10); -lean_dec(x_8); -x_11 = l_Lean_Meta_SynthInstance_getNextToResume___rarg___closed__1; -x_12 = l_Array_back___rarg(x_11, x_10); +x_11 = lean_ctor_get(x_10, 3); +lean_inc(x_11); +if (lean_obj_tag(x_11) == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_dec(x_10); -x_13 = lean_st_ref_take(x_1, x_9); -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); +lean_dec(x_9); +x_12 = lean_ctor_get(x_8, 1); +lean_inc(x_12); +lean_dec(x_8); +x_13 = l_Lean_Meta_SynthInstance_resume___closed__3; +x_14 = l_panic___at_Lean_Meta_SynthInstance_resume___spec__1(x_13, x_1, x_2, x_3, x_4, x_5, x_6, 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; +x_15 = lean_ctor_get(x_8, 1); lean_inc(x_15); -lean_dec(x_13); -x_16 = !lean_is_exclusive(x_14); -if (x_16 == 0) +lean_dec(x_8); +x_16 = lean_ctor_get(x_9, 1); +lean_inc(x_16); +lean_dec(x_9); +x_17 = lean_ctor_get(x_10, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_10, 1); +lean_inc(x_18); +x_19 = lean_ctor_get(x_10, 2); +lean_inc(x_19); +x_20 = lean_ctor_get(x_10, 4); +lean_inc(x_20); +lean_dec(x_10); +x_21 = lean_ctor_get(x_11, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_11, 1); +lean_inc(x_22); +lean_dec(x_11); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +lean_inc(x_16); +lean_inc(x_21); +lean_inc(x_19); +x_23 = l_Lean_Meta_SynthInstance_tryAnswer(x_19, x_21, x_16, x_1, x_2, x_3, x_4, x_5, x_6, x_15); +if (lean_obj_tag(x_23) == 0) { -lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; -x_17 = lean_ctor_get(x_14, 2); -x_18 = lean_array_pop(x_17); -lean_ctor_set(x_14, 2, x_18); -x_19 = lean_st_ref_set(x_1, x_14, x_15); -x_20 = !lean_is_exclusive(x_19); -if (x_20 == 0) +lean_object* x_24; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +if (lean_obj_tag(x_24) == 0) { -lean_object* x_21; -x_21 = lean_ctor_get(x_19, 0); +uint8_t x_25; +lean_dec(x_22); lean_dec(x_21); -lean_ctor_set(x_19, 0, x_12); -return x_19; +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +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_25 = !lean_is_exclusive(x_23); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; +x_26 = lean_ctor_get(x_23, 0); +lean_dec(x_26); +x_27 = lean_box(0); +lean_ctor_set(x_23, 0, x_27); +return x_23; +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_23, 1); +lean_inc(x_28); +lean_dec(x_23); +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; +} +} +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; +x_31 = lean_ctor_get(x_23, 1); +lean_inc(x_31); +lean_dec(x_23); +x_32 = lean_ctor_get(x_24, 0); +lean_inc(x_32); +lean_dec(x_24); +lean_inc(x_17); +x_33 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_mkTableKeyFor___lambda__1___boxed), 8, 1); +lean_closure_set(x_33, 0, x_17); +lean_inc(x_32); +x_34 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_resume___lambda__5), 16, 8); +lean_closure_set(x_34, 0, x_21); +lean_closure_set(x_34, 1, x_16); +lean_closure_set(x_34, 2, x_19); +lean_closure_set(x_34, 3, x_20); +lean_closure_set(x_34, 4, x_17); +lean_closure_set(x_34, 5, x_18); +lean_closure_set(x_34, 6, x_32); +lean_closure_set(x_34, 7, x_22); +x_35 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_SynthInstance_newSubgoal___spec__7___rarg), 9, 2); +lean_closure_set(x_35, 0, x_33); +lean_closure_set(x_35, 1, x_34); +x_36 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__16___rarg(x_32, x_35, x_1, x_2, x_3, x_4, x_5, x_6, x_31); +return x_36; +} } else { -lean_object* x_22; lean_object* x_23; -x_22 = lean_ctor_get(x_19, 1); -lean_inc(x_22); +uint8_t x_37; +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_20); lean_dec(x_19); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_12); -lean_ctor_set(x_23, 1, x_22); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +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_is_exclusive(x_23); +if (x_37 == 0) +{ 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; -x_24 = lean_ctor_get(x_14, 0); -x_25 = lean_ctor_get(x_14, 1); -x_26 = lean_ctor_get(x_14, 2); -x_27 = lean_ctor_get(x_14, 3); -lean_inc(x_27); -lean_inc(x_26); -lean_inc(x_25); -lean_inc(x_24); -lean_dec(x_14); -x_28 = lean_array_pop(x_26); -x_29 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_29, 0, x_24); -lean_ctor_set(x_29, 1, x_25); -lean_ctor_set(x_29, 2, x_28); -lean_ctor_set(x_29, 3, x_27); -x_30 = lean_st_ref_set(x_1, x_29, x_15); -x_31 = lean_ctor_get(x_30, 1); -lean_inc(x_31); -if (lean_is_exclusive(x_30)) { - lean_ctor_release(x_30, 0); - lean_ctor_release(x_30, 1); - x_32 = x_30; -} else { - lean_dec_ref(x_30); - x_32 = lean_box(0); +lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_38 = lean_ctor_get(x_23, 0); +x_39 = lean_ctor_get(x_23, 1); +lean_inc(x_39); +lean_inc(x_38); +lean_dec(x_23); +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; } -if (lean_is_scalar(x_32)) { - x_33 = lean_alloc_ctor(0, 2, 0); -} else { - x_33 = x_32; } -lean_ctor_set(x_33, 0, x_12); -lean_ctor_set(x_33, 1, x_31); -return x_33; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_getNextToResume(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_resume___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_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_getNextToResume___rarg___boxed), 6, 0); -return x_2; +lean_object* x_11; +x_11 = l_Lean_Meta_SynthInstance_resume___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_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_getNextToResume___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_resume___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_7; -x_7 = l_Lean_Meta_SynthInstance_getNextToResume___rarg(x_1, x_2, x_3, x_4, x_5, x_6); +lean_object* x_13; +x_13 = l_Lean_Meta_SynthInstance_resume___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_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_7; +return x_13; } } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_getNextToResume___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_resume___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) { _start: { -lean_object* x_2; -x_2 = l_Lean_Meta_SynthInstance_getNextToResume(x_1); +lean_object* x_15; +x_15 = l_Lean_Meta_SynthInstance_resume___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, x_14); +lean_dec(x_7); +lean_dec(x_2); lean_dec(x_1); -return x_2; +return x_15; } } -LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_SynthInstance_resume___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_Meta_SynthInstance_resume___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, lean_object* x_15) { _start: { -lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_9 = l_panic___at_Lean_Meta_SynthInstance_getEntry___spec__1___closed__1; -x_10 = lean_panic_fn(x_9, x_1); -x_11 = lean_apply_7(x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -return x_11; +uint8_t x_16; lean_object* x_17; +x_16 = lean_unbox(x_8); +lean_dec(x_8); +x_17 = l_Lean_Meta_SynthInstance_resume___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_16, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_2); +lean_dec(x_1); +return x_17; } } -static lean_object* _init_l_Lean_Meta_SynthInstance_resume___lambda__1___closed__1() { +LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_step(lean_object* x_1, 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_1; -x_1 = lean_mk_string_from_bytes("propagating ", 12); -return x_1; -} +lean_object* x_8; +x_8 = l_Lean_Meta_SynthInstance_checkMaxHeartbeats(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_9 = lean_ctor_get(x_8, 1); +lean_inc(x_9); +lean_dec(x_8); +x_10 = lean_st_ref_get(x_2, x_9); +x_11 = !lean_is_exclusive(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_ctor_get(x_10, 0); +x_13 = lean_ctor_get(x_10, 1); +x_14 = lean_ctor_get(x_12, 2); +lean_inc(x_14); +x_15 = l_Array_isEmpty___rarg(x_14); +lean_dec(x_14); +if (x_15 == 0) +{ +lean_object* x_16; +lean_free_object(x_10); +lean_dec(x_12); +x_16 = l_Lean_Meta_SynthInstance_resume(x_1, x_2, x_3, x_4, x_5, x_6, x_13); +if (lean_obj_tag(x_16) == 0) +{ +uint8_t x_17; +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) +{ +lean_object* x_18; uint8_t x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_16, 0); +lean_dec(x_18); +x_19 = 1; +x_20 = lean_box(x_19); +lean_ctor_set(x_16, 0, x_20); +return x_16; } -static lean_object* _init_l_Lean_Meta_SynthInstance_resume___lambda__1___closed__2() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_SynthInstance_resume___lambda__1___closed__1; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +lean_object* x_21; uint8_t x_22; lean_object* x_23; lean_object* x_24; +x_21 = lean_ctor_get(x_16, 1); +lean_inc(x_21); +lean_dec(x_16); +x_22 = 1; +x_23 = lean_box(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_21); +return x_24; } } -static lean_object* _init_l_Lean_Meta_SynthInstance_resume___lambda__1___closed__3() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes(" to subgoal ", 12); -return x_1; -} +uint8_t x_25; +x_25 = !lean_is_exclusive(x_16); +if (x_25 == 0) +{ +return x_16; } -static lean_object* _init_l_Lean_Meta_SynthInstance_resume___lambda__1___closed__4() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_SynthInstance_resume___lambda__1___closed__3; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_16, 0); +x_27 = lean_ctor_get(x_16, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_16); +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; } } -static lean_object* _init_l_Lean_Meta_SynthInstance_resume___lambda__1___closed__5() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes(" of ", 4); -return x_1; } +else +{ +lean_object* x_29; uint8_t x_30; +x_29 = lean_ctor_get(x_12, 1); +lean_inc(x_29); +lean_dec(x_12); +x_30 = l_Array_isEmpty___rarg(x_29); +lean_dec(x_29); +if (x_30 == 0) +{ +lean_object* x_31; +lean_free_object(x_10); +x_31 = l_Lean_Meta_SynthInstance_generate(x_1, x_2, x_3, x_4, x_5, x_6, x_13); +if (lean_obj_tag(x_31) == 0) +{ +uint8_t x_32; +x_32 = !lean_is_exclusive(x_31); +if (x_32 == 0) +{ +lean_object* x_33; uint8_t x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_31, 0); +lean_dec(x_33); +x_34 = 1; +x_35 = lean_box(x_34); +lean_ctor_set(x_31, 0, x_35); +return x_31; } -static lean_object* _init_l_Lean_Meta_SynthInstance_resume___lambda__1___closed__6() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_SynthInstance_resume___lambda__1___closed__5; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +lean_object* x_36; uint8_t x_37; lean_object* x_38; lean_object* x_39; +x_36 = lean_ctor_get(x_31, 1); +lean_inc(x_36); +lean_dec(x_31); +x_37 = 1; +x_38 = lean_box(x_37); +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_36); +return x_39; } } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_resume___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: +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_instantiateMVars___at_Lean_Meta_SynthInstance_mkTableKeyFor___spec__1(x_1, 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 = l_Lean_instantiateMVars___at_Lean_Meta_SynthInstance_mkTableKeyFor___spec__1(x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_13); -x_15 = !lean_is_exclusive(x_14); -if (x_15 == 0) +uint8_t x_40; +x_40 = !lean_is_exclusive(x_31); +if (x_40 == 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_14, 0); -x_17 = l_Lean_MessageData_ofExpr(x_3); -x_18 = l_Lean_Meta_SynthInstance_resume___lambda__1___closed__2; -x_19 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_19, 0, x_18); -lean_ctor_set(x_19, 1, x_17); -x_20 = l_Lean_Meta_SynthInstance_resume___lambda__1___closed__4; -x_21 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_21, 0, x_19); -lean_ctor_set(x_21, 1, x_20); -x_22 = l_Lean_MessageData_ofExpr(x_12); -x_23 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_23, 0, x_21); -lean_ctor_set(x_23, 1, x_22); -x_24 = l_Lean_Meta_SynthInstance_resume___lambda__1___closed__6; -x_25 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_25, 0, x_23); -lean_ctor_set(x_25, 1, x_24); -x_26 = l_Lean_MessageData_ofExpr(x_16); -x_27 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_27, 0, x_25); -lean_ctor_set(x_27, 1, x_26); -x_28 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; -x_29 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_29, 0, x_27); -lean_ctor_set(x_29, 1, x_28); -lean_ctor_set(x_14, 0, x_29); -return x_14; +return x_31; } else { -lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_30 = lean_ctor_get(x_14, 0); -x_31 = lean_ctor_get(x_14, 1); -lean_inc(x_31); -lean_inc(x_30); -lean_dec(x_14); -x_32 = l_Lean_MessageData_ofExpr(x_3); -x_33 = l_Lean_Meta_SynthInstance_resume___lambda__1___closed__2; -x_34 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_34, 0, x_33); -lean_ctor_set(x_34, 1, x_32); -x_35 = l_Lean_Meta_SynthInstance_resume___lambda__1___closed__4; -x_36 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_36, 0, x_34); -lean_ctor_set(x_36, 1, x_35); -x_37 = l_Lean_MessageData_ofExpr(x_12); -x_38 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_38, 0, x_36); -lean_ctor_set(x_38, 1, x_37); -x_39 = l_Lean_Meta_SynthInstance_resume___lambda__1___closed__6; -x_40 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set(x_40, 1, x_39); -x_41 = l_Lean_MessageData_ofExpr(x_30); -x_42 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_42, 0, x_40); -lean_ctor_set(x_42, 1, x_41); -x_43 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; -x_44 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_44, 0, x_42); -lean_ctor_set(x_44, 1, 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_31); -return x_45; +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; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_resume___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: +else { -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_1, 1); -lean_inc(x_13); +uint8_t x_44; lean_object* x_45; +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_closure((void*)(l_Lean_instantiateMVars___at_Lean_Meta_SynthInstance_mkTableKeyFor___spec__1___boxed), 8, 1); -lean_closure_set(x_14, 0, x_13); -x_15 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_resume___lambda__1___boxed), 10, 2); -lean_closure_set(x_15, 0, x_2); -lean_closure_set(x_15, 1, x_3); -x_16 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_SynthInstance_newSubgoal___spec__7___rarg), 9, 2); -lean_closure_set(x_16, 0, x_14); -lean_closure_set(x_16, 1, x_15); -x_17 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__18___rarg(x_4, x_16, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_17; +x_44 = 0; +x_45 = lean_box(x_44); +lean_ctor_set(x_10, 0, x_45); +return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_resume___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) { -_start: +} +else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_15 = lean_ctor_get(x_1, 2); -x_16 = lean_nat_add(x_2, x_15); -x_17 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_17, 0, x_3); -lean_ctor_set(x_17, 1, x_4); -lean_ctor_set(x_17, 2, x_5); -lean_ctor_set(x_17, 3, x_6); -lean_ctor_set(x_17, 4, x_16); -x_18 = l_Lean_Meta_SynthInstance_consume(x_17, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -return x_18; +lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; +x_46 = lean_ctor_get(x_10, 0); +x_47 = lean_ctor_get(x_10, 1); +lean_inc(x_47); +lean_inc(x_46); +lean_dec(x_10); +x_48 = lean_ctor_get(x_46, 2); +lean_inc(x_48); +x_49 = l_Array_isEmpty___rarg(x_48); +lean_dec(x_48); +if (x_49 == 0) +{ +lean_object* x_50; +lean_dec(x_46); +x_50 = l_Lean_Meta_SynthInstance_resume(x_1, x_2, x_3, x_4, x_5, x_6, x_47); +if (lean_obj_tag(x_50) == 0) +{ +lean_object* x_51; lean_object* x_52; uint8_t x_53; lean_object* x_54; lean_object* x_55; +x_51 = lean_ctor_get(x_50, 1); +lean_inc(x_51); +if (lean_is_exclusive(x_50)) { + lean_ctor_release(x_50, 0); + lean_ctor_release(x_50, 1); + x_52 = x_50; +} else { + lean_dec_ref(x_50); + x_52 = lean_box(0); +} +x_53 = 1; +x_54 = lean_box(x_53); +if (lean_is_scalar(x_52)) { + x_55 = lean_alloc_ctor(0, 2, 0); +} else { + x_55 = x_52; } +lean_ctor_set(x_55, 0, x_54); +lean_ctor_set(x_55, 1, x_51); +return x_55; } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_resume___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, uint8_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) { -_start: -{ -if (x_8 == 0) +else { -lean_object* x_16; lean_object* x_17; -lean_dec(x_7); -x_16 = lean_box(0); -x_17 = l_Lean_Meta_SynthInstance_resume___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_16, x_9, x_10, x_11, x_12, x_13, x_14, x_15); -return x_17; +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_56 = lean_ctor_get(x_50, 0); +lean_inc(x_56); +x_57 = lean_ctor_get(x_50, 1); +lean_inc(x_57); +if (lean_is_exclusive(x_50)) { + lean_ctor_release(x_50, 0); + lean_ctor_release(x_50, 1); + x_58 = x_50; +} else { + lean_dec_ref(x_50); + x_58 = lean_box(0); +} +if (lean_is_scalar(x_58)) { + x_59 = lean_alloc_ctor(1, 2, 0); +} else { + x_59 = x_58; +} +lean_ctor_set(x_59, 0, x_56); +lean_ctor_set(x_59, 1, x_57); +return x_59; +} } 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; -x_18 = lean_ctor_get(x_1, 2); -x_19 = lean_nat_add(x_2, x_18); -x_20 = l_Nat_repr(x_19); -x_21 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_21, 0, x_20); -x_22 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_22, 0, x_21); -x_23 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___lambda__2___closed__4; -x_24 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_22); -x_25 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; -x_26 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_26, 0, x_24); -lean_ctor_set(x_26, 1, x_25); -x_27 = l_Lean_addTrace___at_Lean_Meta_SynthInstance_wakeUp___spec__1(x_7, x_26, x_9, x_10, x_11, x_12, x_13, x_14, x_15); -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_Meta_SynthInstance_resume___lambda__3(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; +lean_object* x_60; uint8_t x_61; +x_60 = lean_ctor_get(x_46, 1); +lean_inc(x_60); +lean_dec(x_46); +x_61 = l_Array_isEmpty___rarg(x_60); +lean_dec(x_60); +if (x_61 == 0) +{ +lean_object* x_62; +x_62 = l_Lean_Meta_SynthInstance_generate(x_1, x_2, x_3, x_4, x_5, x_6, x_47); +if (lean_obj_tag(x_62) == 0) +{ +lean_object* x_63; lean_object* x_64; uint8_t 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_62); + x_64 = lean_box(0); } +x_65 = 1; +x_66 = lean_box(x_65); +if (lean_is_scalar(x_64)) { + x_67 = lean_alloc_ctor(0, 2, 0); +} else { + x_67 = x_64; } +lean_ctor_set(x_67, 0, x_66); +lean_ctor_set(x_67, 1, x_63); +return x_67; } -static lean_object* _init_l_Lean_Meta_SynthInstance_resume___lambda__5___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("resume", 6); -return x_1; +lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; +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_62); + x_70 = lean_box(0); } +if (lean_is_scalar(x_70)) { + x_71 = lean_alloc_ctor(1, 2, 0); +} else { + x_71 = x_70; } -static lean_object* _init_l_Lean_Meta_SynthInstance_resume___lambda__5___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_6____closed__8; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_6____closed__1; -x_3 = l_Lean_Meta_SynthInstance_resume___lambda__5___closed__1; -x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); -return x_4; +lean_ctor_set(x_71, 0, x_68); +lean_ctor_set(x_71, 1, x_69); +return x_71; } } -static lean_object* _init_l_Lean_Meta_SynthInstance_resume___lambda__5___closed__3() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_SynthInstance_resume___lambda__5___closed__2; -x_2 = lean_alloc_closure((void*)(l_Lean_isTracingEnabledFor___at_Lean_Meta_SynthInstance_newSubgoal___spec__10___boxed), 8, 1); -lean_closure_set(x_2, 0, x_1); -return x_2; +uint8_t x_72; lean_object* x_73; lean_object* x_74; +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_72 = 0; +x_73 = lean_box(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_47); +return x_74; +} } } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_resume___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_17; -lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -x_17 = lean_infer_type(x_1, x_12, x_13, x_14, x_15, x_16); -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; lean_object* x_24; uint8_t x_25; lean_object* x_26; -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); -lean_inc(x_2); -x_20 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_resume___lambda__2___boxed), 12, 4); -lean_closure_set(x_20, 0, x_2); -lean_closure_set(x_20, 1, x_18); -lean_closure_set(x_20, 2, x_9); -lean_closure_set(x_20, 3, x_3); -x_21 = l_Lean_Meta_SynthInstance_resume___lambda__5___closed__2; -x_22 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_resume___lambda__4___boxed), 15, 7); -lean_closure_set(x_22, 0, x_2); -lean_closure_set(x_22, 1, x_4); -lean_closure_set(x_22, 2, x_5); -lean_closure_set(x_22, 3, x_6); -lean_closure_set(x_22, 4, x_7); -lean_closure_set(x_22, 5, x_8); -lean_closure_set(x_22, 6, x_21); -x_23 = l_Lean_Meta_SynthInstance_resume___lambda__5___closed__3; -x_24 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_SynthInstance_newSubgoal___spec__7___rarg), 9, 2); -lean_closure_set(x_24, 0, x_23); -lean_closure_set(x_24, 1, x_22); -x_25 = 1; -x_26 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_addAnswer___spec__3(x_21, x_20, x_24, x_25, x_10, x_11, x_12, x_13, x_14, x_15, x_19); -return x_26; } else { -uint8_t x_27; -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); -lean_dec(x_7); +uint8_t x_75; lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_27 = !lean_is_exclusive(x_17); -if (x_27 == 0) +lean_dec(x_1); +x_75 = !lean_is_exclusive(x_8); +if (x_75 == 0) { -return x_17; +return x_8; } else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_17, 0); -x_29 = lean_ctor_get(x_17, 1); -lean_inc(x_29); -lean_inc(x_28); -lean_dec(x_17); -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_76; lean_object* x_77; lean_object* x_78; +x_76 = lean_ctor_get(x_8, 0); +x_77 = lean_ctor_get(x_8, 1); +lean_inc(x_77); +lean_inc(x_76); +lean_dec(x_8); +x_78 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_78, 0, x_76); +lean_ctor_set(x_78, 1, x_77); +return x_78; } } } } -static lean_object* _init_l_Lean_Meta_SynthInstance_resume___closed__1() { +LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_getResult___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_1; -x_1 = lean_mk_string_from_bytes("Lean.Meta.SynthInstance.resume", 30); -return x_1; +lean_object* x_7; uint8_t x_8; +x_7 = lean_st_ref_get(x_1, x_6); +x_8 = !lean_is_exclusive(x_7); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; +x_9 = lean_ctor_get(x_7, 0); +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +lean_dec(x_9); +lean_ctor_set(x_7, 0, x_10); +return x_7; +} +else +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_11 = lean_ctor_get(x_7, 0); +x_12 = lean_ctor_get(x_7, 1); +lean_inc(x_12); +lean_inc(x_11); +lean_dec(x_7); +x_13 = lean_ctor_get(x_11, 0); +lean_inc(x_13); +lean_dec(x_11); +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; } } -static lean_object* _init_l_Lean_Meta_SynthInstance_resume___closed__2() { +} +LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_getResult(lean_object* x_1) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("resume found no remaining subgoals", 34); -return x_1; +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_getResult___rarg___boxed), 6, 0); +return x_2; } } -static lean_object* _init_l_Lean_Meta_SynthInstance_resume___closed__3() { +LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_getResult___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_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_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_getInstances___spec__7___closed__1; -x_2 = l_Lean_Meta_SynthInstance_resume___closed__1; -x_3 = lean_unsigned_to_nat(545u); -x_4 = lean_unsigned_to_nat(18u); -x_5 = l_Lean_Meta_SynthInstance_resume___closed__2; -x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); -return x_6; +lean_object* x_7; +x_7 = l_Lean_Meta_SynthInstance_getResult___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_Meta_SynthInstance_resume(lean_object* x_1, lean_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_Meta_SynthInstance_getResult___boxed(lean_object* x_1) { _start: { -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_8 = l_Lean_Meta_SynthInstance_getNextToResume___rarg(x_2, x_3, x_4, x_5, x_6, x_7); -x_9 = lean_ctor_get(x_8, 0); -lean_inc(x_9); -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -x_11 = lean_ctor_get(x_10, 3); -lean_inc(x_11); -if (lean_obj_tag(x_11) == 0) -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; -lean_dec(x_10); -lean_dec(x_9); -x_12 = lean_ctor_get(x_8, 1); -lean_inc(x_12); -lean_dec(x_8); -x_13 = l_Lean_Meta_SynthInstance_resume___closed__3; -x_14 = l_panic___at_Lean_Meta_SynthInstance_resume___spec__1(x_13, x_1, x_2, x_3, x_4, x_5, x_6, x_12); -return x_14; +lean_object* x_2; +x_2 = l_Lean_Meta_SynthInstance_getResult(x_1); +lean_dec(x_1); +return x_2; } -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; -x_15 = lean_ctor_get(x_8, 1); -lean_inc(x_15); -lean_dec(x_8); -x_16 = lean_ctor_get(x_9, 1); -lean_inc(x_16); -lean_dec(x_9); -x_17 = lean_ctor_get(x_10, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_10, 1); -lean_inc(x_18); -x_19 = lean_ctor_get(x_10, 2); -lean_inc(x_19); -x_20 = lean_ctor_get(x_10, 4); -lean_inc(x_20); -lean_dec(x_10); -x_21 = lean_ctor_get(x_11, 0); -lean_inc(x_21); -x_22 = lean_ctor_get(x_11, 1); -lean_inc(x_22); -lean_dec(x_11); +} +LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_synth(lean_object* x_1, 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_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); lean_inc(x_1); -lean_inc(x_16); -lean_inc(x_21); -lean_inc(x_19); -x_23 = l_Lean_Meta_SynthInstance_tryAnswer(x_19, x_21, x_16, x_1, x_2, x_3, x_4, x_5, x_6, x_15); -if (lean_obj_tag(x_23) == 0) +x_8 = l_Lean_Meta_SynthInstance_step(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_8) == 0) { -lean_object* x_24; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -if (lean_obj_tag(x_24) == 0) +lean_object* x_9; uint8_t x_10; +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_unbox(x_9); +lean_dec(x_9); +if (x_10 == 0) { -uint8_t x_25; -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_16); +uint8_t x_11; 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_25 = !lean_is_exclusive(x_23); -if (x_25 == 0) +x_11 = !lean_is_exclusive(x_8); +if (x_11 == 0) { -lean_object* x_26; lean_object* x_27; -x_26 = lean_ctor_get(x_23, 0); -lean_dec(x_26); -x_27 = lean_box(0); -lean_ctor_set(x_23, 0, x_27); -return x_23; +lean_object* x_12; lean_object* x_13; +x_12 = lean_ctor_get(x_8, 0); +lean_dec(x_12); +x_13 = lean_box(0); +lean_ctor_set(x_8, 0, x_13); +return x_8; } else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_23, 1); -lean_inc(x_28); -lean_dec(x_23); -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_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_8, 1); +lean_inc(x_14); +lean_dec(x_8); +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_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_23, 1); -lean_inc(x_31); -lean_dec(x_23); -x_32 = lean_ctor_get(x_24, 0); -lean_inc(x_32); -lean_dec(x_24); +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_8, 1); lean_inc(x_17); -x_33 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_mkTableKeyFor___lambda__1___boxed), 8, 1); -lean_closure_set(x_33, 0, x_17); -lean_inc(x_32); -x_34 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_resume___lambda__5), 16, 8); -lean_closure_set(x_34, 0, x_21); -lean_closure_set(x_34, 1, x_16); -lean_closure_set(x_34, 2, x_19); -lean_closure_set(x_34, 3, x_20); -lean_closure_set(x_34, 4, x_17); -lean_closure_set(x_34, 5, x_18); -lean_closure_set(x_34, 6, x_32); -lean_closure_set(x_34, 7, x_22); -x_35 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_SynthInstance_newSubgoal___spec__7___rarg), 9, 2); -lean_closure_set(x_35, 0, x_33); -lean_closure_set(x_35, 1, x_34); -x_36 = l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_newSubgoal___spec__18___rarg(x_32, x_35, x_1, x_2, x_3, x_4, x_5, x_6, x_31); -return x_36; -} +lean_dec(x_8); +x_18 = l_Lean_Meta_SynthInstance_getResult___rarg(x_2, x_3, x_4, x_5, x_6, x_17); +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_20; +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_7 = x_20; +goto _start; } else { -uint8_t x_37; -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_16); +uint8_t x_22; 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_is_exclusive(x_23); -if (x_37 == 0) +x_22 = !lean_is_exclusive(x_18); +if (x_22 == 0) { -return x_23; +lean_object* x_23; uint8_t x_24; +x_23 = lean_ctor_get(x_18, 0); +lean_dec(x_23); +x_24 = !lean_is_exclusive(x_19); +if (x_24 == 0) +{ +return x_18; } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = lean_ctor_get(x_23, 0); -x_39 = lean_ctor_get(x_23, 1); -lean_inc(x_39); -lean_inc(x_38); -lean_dec(x_23); -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_object* x_25; lean_object* x_26; +x_25 = lean_ctor_get(x_19, 0); +lean_inc(x_25); +lean_dec(x_19); +x_26 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_18, 0, x_26); +return x_18; } } +else +{ +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_18, 1); +lean_inc(x_27); +lean_dec(x_18); +x_28 = lean_ctor_get(x_19, 0); +lean_inc(x_28); +if (lean_is_exclusive(x_19)) { + lean_ctor_release(x_19, 0); + x_29 = x_19; +} else { + lean_dec_ref(x_19); + x_29 = lean_box(0); } +if (lean_is_scalar(x_29)) { + x_30 = lean_alloc_ctor(1, 1, 0); +} else { + x_30 = x_29; } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_resume___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_Lean_Meta_SynthInstance_resume___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_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_11; +lean_ctor_set(x_30, 0, x_28); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_27); +return x_31; } } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_resume___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_13; -x_13 = l_Lean_Meta_SynthInstance_resume___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_5); -return x_13; } } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_resume___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) { -_start: +else { -lean_object* x_15; -x_15 = l_Lean_Meta_SynthInstance_resume___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, x_14); -lean_dec(x_7); +uint8_t x_32; +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_15; -} +x_32 = !lean_is_exclusive(x_8); +if (x_32 == 0) +{ +return x_8; } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_resume___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, lean_object* x_15) { -_start: +else { -uint8_t x_16; lean_object* x_17; -x_16 = lean_unbox(x_8); +lean_object* x_33; lean_object* x_34; lean_object* x_35; +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_17 = l_Lean_Meta_SynthInstance_resume___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_16, x_9, x_10, x_11, x_12, x_13, x_14, x_15); -lean_dec(x_2); -lean_dec(x_1); -return x_17; +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_EXPORT lean_object* l_Lean_Meta_SynthInstance_step(lean_object* x_1, lean_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_Meta_SynthInstance_mkTableKey___at_Lean_Meta_SynthInstance_main___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_8; -x_8 = l_Lean_Meta_SynthInstance_checkMaxHeartbeats(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_8) == 0) -{ -lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_9 = lean_ctor_get(x_8, 1); +lean_object* x_7; 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; uint8_t x_21; +x_7 = lean_st_ref_get(x_3, x_6); +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 = lean_ctor_get(x_8, 0); +lean_inc(x_10); lean_dec(x_8); -x_10 = lean_st_ref_get(x_2, x_9); -x_11 = !lean_is_exclusive(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_ctor_get(x_10, 0); -x_13 = lean_ctor_get(x_10, 1); -x_14 = lean_ctor_get(x_12, 2); -lean_inc(x_14); -x_15 = l_Array_isEmpty___rarg(x_14); +x_11 = lean_unsigned_to_nat(0u); +x_12 = l_Lean_Meta_SynthInstance_MkTableKey_State_lmap___default___closed__1; +x_13 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_13, 0, x_11); +lean_ctor_set(x_13, 1, x_12); +lean_ctor_set(x_13, 2, x_12); +lean_ctor_set(x_13, 3, x_10); +x_14 = l_Lean_Meta_SynthInstance_MkTableKey_normExpr(x_1, x_13); +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); -if (x_15 == 0) -{ -lean_object* x_16; -lean_free_object(x_10); -lean_dec(x_12); -x_16 = l_Lean_Meta_SynthInstance_resume(x_1, x_2, x_3, x_4, x_5, x_6, x_13); -if (lean_obj_tag(x_16) == 0) +x_17 = lean_ctor_get(x_16, 3); +lean_inc(x_17); +lean_dec(x_16); +x_18 = lean_st_ref_take(x_3, x_9); +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_is_exclusive(x_19); +if (x_21 == 0) { -uint8_t x_17; -x_17 = !lean_is_exclusive(x_16); -if (x_17 == 0) +lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_22 = lean_ctor_get(x_19, 0); +lean_dec(x_22); +lean_ctor_set(x_19, 0, x_17); +x_23 = lean_st_ref_set(x_3, x_19, x_20); +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) { -lean_object* x_18; uint8_t x_19; lean_object* x_20; -x_18 = lean_ctor_get(x_16, 0); -lean_dec(x_18); -x_19 = 1; -x_20 = lean_box(x_19); -lean_ctor_set(x_16, 0, x_20); -return x_16; +lean_object* x_25; +x_25 = lean_ctor_get(x_23, 0); +lean_dec(x_25); +lean_ctor_set(x_23, 0, x_15); +return x_23; } else { -lean_object* x_21; uint8_t x_22; lean_object* x_23; lean_object* x_24; -x_21 = lean_ctor_get(x_16, 1); -lean_inc(x_21); -lean_dec(x_16); -x_22 = 1; -x_23 = lean_box(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_21); -return x_24; +lean_object* x_26; lean_object* x_27; +x_26 = lean_ctor_get(x_23, 1); +lean_inc(x_26); +lean_dec(x_23); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_15); +lean_ctor_set(x_27, 1, x_26); +return x_27; } } else { -uint8_t x_25; -x_25 = !lean_is_exclusive(x_16); -if (x_25 == 0) -{ -return x_16; +lean_object* x_28; lean_object* x_29; 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_19, 1); +x_29 = lean_ctor_get(x_19, 2); +x_30 = lean_ctor_get(x_19, 3); +lean_inc(x_30); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_19); +x_31 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_31, 0, x_17); +lean_ctor_set(x_31, 1, x_28); +lean_ctor_set(x_31, 2, x_29); +lean_ctor_set(x_31, 3, x_30); +x_32 = lean_st_ref_set(x_3, x_31, x_20); +x_33 = lean_ctor_get(x_32, 1); +lean_inc(x_33); +if (lean_is_exclusive(x_32)) { + lean_ctor_release(x_32, 0); + lean_ctor_release(x_32, 1); + x_34 = x_32; +} else { + lean_dec_ref(x_32); + x_34 = lean_box(0); } -else -{ -lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_26 = lean_ctor_get(x_16, 0); -x_27 = lean_ctor_get(x_16, 1); -lean_inc(x_27); -lean_inc(x_26); -lean_dec(x_16); -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; +if (lean_is_scalar(x_34)) { + x_35 = lean_alloc_ctor(0, 2, 0); +} else { + x_35 = x_34; } +lean_ctor_set(x_35, 0, x_15); +lean_ctor_set(x_35, 1, x_33); +return x_35; } } -else -{ -lean_object* x_29; uint8_t x_30; -x_29 = lean_ctor_get(x_12, 1); -lean_inc(x_29); -lean_dec(x_12); -x_30 = l_Array_isEmpty___rarg(x_29); -lean_dec(x_29); -if (x_30 == 0) -{ -lean_object* x_31; -lean_free_object(x_10); -x_31 = l_Lean_Meta_SynthInstance_generate(x_1, x_2, x_3, x_4, x_5, x_6, x_13); -if (lean_obj_tag(x_31) == 0) +} +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_SynthInstance_main___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: { -uint8_t x_32; -x_32 = !lean_is_exclusive(x_31); -if (x_32 == 0) +lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_7 = lean_ctor_get(x_4, 5); +x_8 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_2, x_3, x_4, x_5, x_6); +x_9 = !lean_is_exclusive(x_8); +if (x_9 == 0) { -lean_object* x_33; uint8_t x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_31, 0); -lean_dec(x_33); -x_34 = 1; -x_35 = lean_box(x_34); -lean_ctor_set(x_31, 0, x_35); -return x_31; +lean_object* x_10; lean_object* x_11; +x_10 = lean_ctor_get(x_8, 0); +lean_inc(x_7); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_7); +lean_ctor_set(x_11, 1, x_10); +lean_ctor_set_tag(x_8, 1); +lean_ctor_set(x_8, 0, x_11); +return x_8; } else { -lean_object* x_36; uint8_t x_37; lean_object* x_38; lean_object* x_39; -x_36 = lean_ctor_get(x_31, 1); -lean_inc(x_36); -lean_dec(x_31); -x_37 = 1; -x_38 = lean_box(x_37); -x_39 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_39, 1, x_36); -return x_39; +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_12 = lean_ctor_get(x_8, 0); +x_13 = lean_ctor_get(x_8, 1); +lean_inc(x_13); +lean_inc(x_12); +lean_dec(x_8); +lean_inc(x_7); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_7); +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; } } -else +} +LEAN_EXPORT lean_object* l_Lean_Core_withCurrHeartbeats___at_Lean_Meta_SynthInstance_main___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: { -uint8_t x_40; -x_40 = !lean_is_exclusive(x_31); -if (x_40 == 0) +lean_object* x_7; lean_object* x_8; +x_7 = lean_apply_2(x_1, x_2, x_3); +x_8 = l___private_Lean_CoreM_0__Lean_Core_withCurrHeartbeatsImp___rarg(x_7, x_4, x_5, x_6); +if (lean_obj_tag(x_8) == 0) { -return x_31; -} -else +uint8_t x_9; +x_9 = !lean_is_exclusive(x_8); +if (x_9 == 0) { -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; -} -} +return x_8; } else { -uint8_t x_44; lean_object* x_45; -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 = 0; -x_45 = lean_box(x_44); -lean_ctor_set(x_10, 0, x_45); -return x_10; -} +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = lean_ctor_get(x_8, 0); +x_11 = lean_ctor_get(x_8, 1); +lean_inc(x_11); +lean_inc(x_10); +lean_dec(x_8); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_10); +lean_ctor_set(x_12, 1, x_11); +return x_12; } } else { -lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; -x_46 = lean_ctor_get(x_10, 0); -x_47 = lean_ctor_get(x_10, 1); -lean_inc(x_47); -lean_inc(x_46); -lean_dec(x_10); -x_48 = lean_ctor_get(x_46, 2); -lean_inc(x_48); -x_49 = l_Array_isEmpty___rarg(x_48); -lean_dec(x_48); -if (x_49 == 0) -{ -lean_object* x_50; -lean_dec(x_46); -x_50 = l_Lean_Meta_SynthInstance_resume(x_1, x_2, x_3, x_4, x_5, x_6, x_47); -if (lean_obj_tag(x_50) == 0) +uint8_t x_13; +x_13 = !lean_is_exclusive(x_8); +if (x_13 == 0) { -lean_object* x_51; lean_object* x_52; uint8_t x_53; lean_object* x_54; lean_object* x_55; -x_51 = lean_ctor_get(x_50, 1); -lean_inc(x_51); -if (lean_is_exclusive(x_50)) { - lean_ctor_release(x_50, 0); - lean_ctor_release(x_50, 1); - x_52 = x_50; -} else { - lean_dec_ref(x_50); - x_52 = lean_box(0); +return x_8; +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_8, 0); +x_15 = lean_ctor_get(x_8, 1); +lean_inc(x_15); +lean_inc(x_14); +lean_dec(x_8); +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; } -x_53 = 1; -x_54 = lean_box(x_53); -if (lean_is_scalar(x_52)) { - x_55 = lean_alloc_ctor(0, 2, 0); -} else { - x_55 = x_52; } -lean_ctor_set(x_55, 0, x_54); -lean_ctor_set(x_55, 1, x_51); -return x_55; } -else +} +static lean_object* _init_l_Lean_Meta_SynthInstance_main___lambda__1___closed__1() { +_start: { -lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_56 = lean_ctor_get(x_50, 0); -lean_inc(x_56); -x_57 = lean_ctor_get(x_50, 1); -lean_inc(x_57); -if (lean_is_exclusive(x_50)) { - lean_ctor_release(x_50, 0); - lean_ctor_release(x_50, 1); - x_58 = x_50; -} else { - lean_dec_ref(x_50); - x_58 = lean_box(0); +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("failed to synthesize", 20); +return x_1; } -if (lean_is_scalar(x_58)) { - x_59 = lean_alloc_ctor(1, 2, 0); -} else { - x_59 = x_58; } -lean_ctor_set(x_59, 0, x_56); -lean_ctor_set(x_59, 1, x_57); -return x_59; +static lean_object* _init_l_Lean_Meta_SynthInstance_main___lambda__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_SynthInstance_main___lambda__1___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } -else +static lean_object* _init_l_Lean_Meta_SynthInstance_main___lambda__1___closed__3() { +_start: { -lean_object* x_60; uint8_t x_61; -x_60 = lean_ctor_get(x_46, 1); -lean_inc(x_60); -lean_dec(x_46); -x_61 = l_Array_isEmpty___rarg(x_60); -lean_dec(x_60); -if (x_61 == 0) +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("\n", 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_SynthInstance_main___lambda__1___closed__4() { +_start: { -lean_object* x_62; -x_62 = l_Lean_Meta_SynthInstance_generate(x_1, x_2, x_3, x_4, x_5, x_6, x_47); -if (lean_obj_tag(x_62) == 0) +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_SynthInstance_main___lambda__1___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_SynthInstance_main___lambda__1___closed__5() { +_start: { -lean_object* x_63; lean_object* x_64; uint8_t 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_62); - x_64 = lean_box(0); +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_SynthInstance_instInhabitedInstance___closed__2; +x_3 = l_Lean_Meta_SynthInstance_MkTableKey_State_lmap___default___closed__1; +x_4 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_2); +lean_ctor_set(x_4, 3, x_3); +return x_4; } -x_65 = 1; -x_66 = lean_box(x_65); -if (lean_is_scalar(x_64)) { - x_67 = lean_alloc_ctor(0, 2, 0); +} +LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_main___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: +{ +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_object* x_18; lean_object* x_19; 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_9 = 0; +x_10 = lean_box(0); +lean_inc(x_4); +x_11 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkFreshExprMVarImpl(x_1, x_9, x_10, x_4, x_5, x_6, x_7, x_8); +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_2); +x_14 = l_Lean_Meta_SynthInstance_mkTableKey___at_Lean_Meta_SynthInstance_main___spec__1(x_2, x_4, x_5, x_6, x_7, x_13); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +if (lean_is_exclusive(x_14)) { + lean_ctor_release(x_14, 0); + lean_ctor_release(x_14, 1); + x_17 = x_14; } else { - x_67 = x_64; + lean_dec_ref(x_14); + x_17 = lean_box(0); } -lean_ctor_set(x_67, 0, x_66); -lean_ctor_set(x_67, 1, x_63); -return x_67; +x_33 = lean_ctor_get(x_6, 2); +lean_inc(x_33); +x_34 = l_Lean_Meta_SynthInstance_getMaxHeartbeats(x_33); +lean_dec(x_33); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_3); +lean_ctor_set(x_35, 1, x_34); +x_36 = l_Lean_Meta_SynthInstance_main___lambda__1___closed__5; +x_37 = lean_st_mk_ref(x_36, x_16); +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_st_ref_get(x_5, 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_ctor_get(x_41, 0); +lean_inc(x_43); +lean_dec(x_41); +x_44 = lean_box(1); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_38); +lean_inc(x_35); +x_45 = l_Lean_Meta_SynthInstance_newSubgoal(x_43, x_15, x_12, x_44, x_35, x_38, x_4, x_5, x_6, x_7, x_42); +if (lean_obj_tag(x_45) == 0) +{ +lean_object* x_46; lean_object* x_47; +x_46 = lean_ctor_get(x_45, 1); +lean_inc(x_46); +lean_dec(x_45); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_38); +x_47 = l_Lean_Meta_SynthInstance_synth(x_35, x_38, x_4, x_5, x_6, x_7, x_46); +if (lean_obj_tag(x_47) == 0) +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; uint8_t x_51; +lean_dec(x_17); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +x_48 = lean_ctor_get(x_47, 0); +lean_inc(x_48); +x_49 = lean_ctor_get(x_47, 1); +lean_inc(x_49); +lean_dec(x_47); +x_50 = lean_st_ref_get(x_38, x_49); +lean_dec(x_38); +x_51 = !lean_is_exclusive(x_50); +if (x_51 == 0) +{ +lean_object* x_52; +x_52 = lean_ctor_get(x_50, 0); +lean_dec(x_52); +lean_ctor_set(x_50, 0, x_48); +return x_50; } else { -lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; -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_62); - x_70 = lean_box(0); +lean_object* x_53; lean_object* x_54; +x_53 = lean_ctor_get(x_50, 1); +lean_inc(x_53); +lean_dec(x_50); +x_54 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_54, 0, x_48); +lean_ctor_set(x_54, 1, x_53); +return x_54; } -if (lean_is_scalar(x_70)) { - x_71 = lean_alloc_ctor(1, 2, 0); -} else { - x_71 = x_70; } -lean_ctor_set(x_71, 0, x_68); -lean_ctor_set(x_71, 1, x_69); -return x_71; +else +{ +lean_object* x_55; lean_object* x_56; +lean_dec(x_38); +x_55 = lean_ctor_get(x_47, 0); +lean_inc(x_55); +x_56 = lean_ctor_get(x_47, 1); +lean_inc(x_56); +lean_dec(x_47); +x_18 = x_55; +x_19 = x_56; +goto block_32; } } else { -uint8_t x_72; lean_object* x_73; lean_object* x_74; +lean_object* x_57; lean_object* x_58; +lean_dec(x_38); +lean_dec(x_35); +x_57 = lean_ctor_get(x_45, 0); +lean_inc(x_57); +x_58 = lean_ctor_get(x_45, 1); +lean_inc(x_58); +lean_dec(x_45); +x_18 = x_57; +x_19 = x_58; +goto block_32; +} +block_32: +{ +uint8_t x_20; +x_20 = l_Lean_Exception_isMaxHeartbeat(x_18); +if (x_20 == 0) +{ +lean_object* x_21; +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_72 = 0; -x_73 = lean_box(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_47); -return x_74; -} -} +if (lean_is_scalar(x_17)) { + x_21 = lean_alloc_ctor(1, 2, 0); +} else { + x_21 = x_17; + lean_ctor_set_tag(x_21, 1); } +lean_ctor_set(x_21, 0, x_18); +lean_ctor_set(x_21, 1, x_19); +return x_21; } else { -uint8_t x_75; +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_dec(x_17); +x_22 = l_Lean_indentExpr(x_2); +x_23 = l_Lean_Meta_SynthInstance_main___lambda__1___closed__2; +x_24 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_22); +x_25 = l_Lean_Meta_SynthInstance_main___lambda__1___closed__4; +x_26 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +x_27 = l_Lean_Exception_toMessageData(x_18); +x_28 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 1, x_27); +x_29 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; +x_30 = lean_alloc_ctor(7, 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_Meta_SynthInstance_main___spec__2(x_30, x_4, x_5, x_6, x_7, x_19); +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_75 = !lean_is_exclusive(x_8); -if (x_75 == 0) -{ -return x_8; -} -else -{ -lean_object* x_76; lean_object* x_77; lean_object* x_78; -x_76 = lean_ctor_get(x_8, 0); -x_77 = lean_ctor_get(x_8, 1); -lean_inc(x_77); -lean_inc(x_76); -lean_dec(x_8); -x_78 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_78, 0, x_76); -lean_ctor_set(x_78, 1, x_77); -return x_78; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_getResult___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; uint8_t x_8; -x_7 = lean_st_ref_get(x_1, x_6); -x_8 = !lean_is_exclusive(x_7); -if (x_8 == 0) -{ -lean_object* x_9; lean_object* x_10; -x_9 = lean_ctor_get(x_7, 0); -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -lean_dec(x_9); -lean_ctor_set(x_7, 0, x_10); -return x_7; +return x_31; } -else -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_11 = lean_ctor_get(x_7, 0); -x_12 = lean_ctor_get(x_7, 1); -lean_inc(x_12); -lean_inc(x_11); -lean_dec(x_7); -x_13 = lean_ctor_get(x_11, 0); -lean_inc(x_13); -lean_dec(x_11); -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_Meta_SynthInstance_getResult(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_main(lean_object* x_1, 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_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_getResult___rarg___boxed), 6, 0); -return x_2; +lean_object* x_8; lean_object* x_9; lean_object* x_10; +lean_inc(x_1); +x_8 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_8, 0, x_1); +x_9 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_main___lambda__1), 8, 3); +lean_closure_set(x_9, 0, x_8); +lean_closure_set(x_9, 1, x_1); +lean_closure_set(x_9, 2, x_2); +x_10 = l_Lean_Core_withCurrHeartbeats___at_Lean_Meta_SynthInstance_main___spec__3(x_9, x_3, x_4, x_5, x_6, x_7); +return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_getResult___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_mkTableKey___at_Lean_Meta_SynthInstance_main___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) { _start: { lean_object* x_7; -x_7 = l_Lean_Meta_SynthInstance_getResult___rarg(x_1, x_2, x_3, x_4, x_5, x_6); +x_7 = l_Lean_Meta_SynthInstance_mkTableKey___at_Lean_Meta_SynthInstance_main___spec__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); lean_dec(x_2); -lean_dec(x_1); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_getResult___boxed(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = l_Lean_Meta_SynthInstance_getResult(x_1); -lean_dec(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_synth(lean_object* x_1, lean_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_throwError___at_Lean_Meta_SynthInstance_main___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_8; -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_8 = l_Lean_Meta_SynthInstance_step(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_8) == 0) -{ -lean_object* x_9; uint8_t x_10; -x_9 = lean_ctor_get(x_8, 0); -lean_inc(x_9); -x_10 = lean_unbox(x_9); -lean_dec(x_9); -if (x_10 == 0) -{ -uint8_t x_11; -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_11 = !lean_is_exclusive(x_8); -if (x_11 == 0) -{ -lean_object* x_12; lean_object* x_13; -x_12 = lean_ctor_get(x_8, 0); -lean_dec(x_12); -x_13 = lean_box(0); -lean_ctor_set(x_8, 0, x_13); -return x_8; -} -else -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_8, 1); -lean_inc(x_14); -lean_dec(x_8); -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_17; lean_object* x_18; lean_object* x_19; -x_17 = lean_ctor_get(x_8, 1); -lean_inc(x_17); -lean_dec(x_8); -x_18 = l_Lean_Meta_SynthInstance_getResult___rarg(x_2, x_3, x_4, x_5, x_6, x_17); -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -if (lean_obj_tag(x_19) == 0) -{ -lean_object* x_20; -x_20 = lean_ctor_get(x_18, 1); -lean_inc(x_20); -lean_dec(x_18); -x_7 = x_20; -goto _start; -} -else -{ -uint8_t x_22; -lean_dec(x_6); +lean_object* x_7; +x_7 = l_Lean_throwError___at_Lean_Meta_SynthInstance_main___spec__2(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_22 = !lean_is_exclusive(x_18); -if (x_22 == 0) -{ -lean_object* x_23; uint8_t x_24; -x_23 = lean_ctor_get(x_18, 0); -lean_dec(x_23); -x_24 = !lean_is_exclusive(x_19); -if (x_24 == 0) -{ -return x_18; -} -else -{ -lean_object* x_25; lean_object* x_26; -x_25 = lean_ctor_get(x_19, 0); -lean_inc(x_25); -lean_dec(x_19); -x_26 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_18, 0, x_26); -return x_18; -} -} -else -{ -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_18, 1); -lean_inc(x_27); -lean_dec(x_18); -x_28 = lean_ctor_get(x_19, 0); -lean_inc(x_28); -if (lean_is_exclusive(x_19)) { - lean_ctor_release(x_19, 0); - x_29 = x_19; -} else { - lean_dec_ref(x_19); - x_29 = lean_box(0); -} -if (lean_is_scalar(x_29)) { - x_30 = lean_alloc_ctor(1, 1, 0); -} else { - x_30 = x_29; -} -lean_ctor_set(x_30, 0, x_28); -x_31 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_27); -return x_31; -} +return x_7; } } +LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocess___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; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_8 = lean_whnf(x_2, x_3, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_9; lean_object* x_10; uint8_t x_11; uint8_t x_12; uint8_t x_13; lean_object* x_14; +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_dec(x_8); +x_11 = 0; +x_12 = 1; +x_13 = 1; +x_14 = l_Lean_Meta_mkForallFVars(x_1, x_9, x_11, x_12, x_13, x_3, x_4, x_5, x_6, x_10); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_14; } else { -uint8_t x_32; +uint8_t x_15; 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_8); -if (x_32 == 0) +x_15 = !lean_is_exclusive(x_8); +if (x_15 == 0) { return x_8; } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; -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_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_8, 0); +x_17 = lean_ctor_get(x_8, 1); +lean_inc(x_17); +lean_inc(x_16); lean_dec(x_8); -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; +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_Lean_Meta_SynthInstance_mkTableKey___at_Lean_Meta_SynthInstance_main___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) { +static lean_object* _init_l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocess___closed__1() { _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; lean_object* x_20; uint8_t x_21; -x_7 = lean_st_ref_get(x_3, x_6); -x_8 = lean_ctor_get(x_7, 0); -lean_inc(x_8); -x_9 = lean_ctor_get(x_7, 1); +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocess___lambda__1), 7, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocess(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___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocess___closed__1; +x_8 = l_Lean_Meta_forallTelescopeReducing___at_Lean_Meta_getParamNames___spec__2___rarg(x_1, x_7, x_2, x_3, x_4, x_5, x_6); +return x_8; +} +} +LEAN_EXPORT lean_object* l_List_forIn_loop___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessLevels___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) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_8; +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_2); +lean_ctor_set(x_8, 1, x_7); +return x_8; +} +else +{ +lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_9 = lean_ctor_get(x_1, 0); lean_inc(x_9); -lean_dec(x_7); -x_10 = lean_ctor_get(x_8, 0); +x_10 = lean_ctor_get(x_1, 1); lean_inc(x_10); -lean_dec(x_8); -x_11 = lean_unsigned_to_nat(0u); -x_12 = l_Lean_Meta_SynthInstance_MkTableKey_State_lmap___default___closed__1; -x_13 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_13, 0, x_11); -lean_ctor_set(x_13, 1, x_12); -lean_ctor_set(x_13, 2, x_12); -lean_ctor_set(x_13, 3, x_10); -x_14 = l_Lean_Meta_SynthInstance_MkTableKey_normExpr(x_1, x_13); +lean_dec(x_1); +x_11 = !lean_is_exclusive(x_2); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_12 = lean_ctor_get(x_2, 0); +x_13 = lean_ctor_get(x_2, 1); +x_14 = l_Lean_instantiateLevelMVars___at_Lean_Meta_normalizeLevel___spec__1(x_9, x_3, x_4, x_5, x_6, x_7); 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_ctor_get(x_16, 3); -lean_inc(x_17); -lean_dec(x_16); -x_18 = lean_st_ref_take(x_3, x_9); -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_is_exclusive(x_19); -if (x_21 == 0) -{ -lean_object* x_22; lean_object* x_23; uint8_t x_24; -x_22 = lean_ctor_get(x_19, 0); -lean_dec(x_22); -lean_ctor_set(x_19, 0, x_17); -x_23 = lean_st_ref_set(x_3, x_19, x_20); -x_24 = !lean_is_exclusive(x_23); -if (x_24 == 0) +x_17 = l_Lean_Level_hasMVar(x_15); +if (x_17 == 0) { -lean_object* x_25; -x_25 = lean_ctor_get(x_23, 0); -lean_dec(x_25); -lean_ctor_set(x_23, 0, x_15); -return x_23; +lean_object* x_18; +x_18 = lean_array_push(x_13, x_15); +lean_ctor_set(x_2, 1, x_18); +x_1 = x_10; +x_7 = x_16; +goto _start; } else { -lean_object* x_26; lean_object* x_27; -x_26 = lean_ctor_get(x_23, 1); -lean_inc(x_26); -lean_dec(x_23); -x_27 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_27, 0, x_15); -lean_ctor_set(x_27, 1, x_26); -return x_27; +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_dec(x_15); +lean_dec(x_12); +x_20 = l_Lean_Meta_mkFreshLevelMVar(x_3, x_4, x_5, x_6, 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); +x_23 = lean_array_push(x_13, x_21); +x_24 = 1; +x_25 = lean_box(x_24); +lean_ctor_set(x_2, 1, x_23); +lean_ctor_set(x_2, 0, x_25); +x_1 = x_10; +x_7 = x_22; +goto _start; } } 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; -x_28 = lean_ctor_get(x_19, 1); -x_29 = lean_ctor_get(x_19, 2); -x_30 = lean_ctor_get(x_19, 3); -lean_inc(x_30); -lean_inc(x_29); +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 = lean_ctor_get(x_2, 0); +x_28 = lean_ctor_get(x_2, 1); lean_inc(x_28); -lean_dec(x_19); -x_31 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_31, 0, x_17); -lean_ctor_set(x_31, 1, x_28); -lean_ctor_set(x_31, 2, x_29); -lean_ctor_set(x_31, 3, x_30); -x_32 = lean_st_ref_set(x_3, x_31, x_20); -x_33 = lean_ctor_get(x_32, 1); -lean_inc(x_33); -if (lean_is_exclusive(x_32)) { - lean_ctor_release(x_32, 0); - lean_ctor_release(x_32, 1); - x_34 = x_32; -} else { - lean_dec_ref(x_32); - x_34 = lean_box(0); +lean_inc(x_27); +lean_dec(x_2); +x_29 = l_Lean_instantiateLevelMVars___at_Lean_Meta_normalizeLevel___spec__1(x_9, x_3, x_4, x_5, x_6, x_7); +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 = l_Lean_Level_hasMVar(x_30); +if (x_32 == 0) +{ +lean_object* x_33; lean_object* x_34; +x_33 = lean_array_push(x_28, x_30); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_27); +lean_ctor_set(x_34, 1, x_33); +x_1 = x_10; +x_2 = x_34; +x_7 = x_31; +goto _start; } -if (lean_is_scalar(x_34)) { - x_35 = lean_alloc_ctor(0, 2, 0); -} else { - x_35 = x_34; +else +{ +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_dec(x_30); +lean_dec(x_27); +x_36 = l_Lean_Meta_mkFreshLevelMVar(x_3, x_4, x_5, x_6, x_31); +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_array_push(x_28, x_37); +x_40 = 1; +x_41 = lean_box(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_39); +x_1 = x_10; +x_2 = x_42; +x_7 = x_38; +goto _start; } -lean_ctor_set(x_35, 0, x_15); -lean_ctor_set(x_35, 1, x_33); -return x_35; } } } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_SynthInstance_main___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) { +} +static lean_object* _init_l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessLevels___closed__1() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 0; +x_2 = l_Lean_Meta_SynthInstance_instInhabitedInstance___closed__2; +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; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessLevels(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 = lean_ctor_get(x_4, 5); -x_8 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_2, x_3, x_4, x_5, x_6); +x_7 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessLevels___closed__1; +x_8 = l_List_forIn_loop___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessLevels___spec__1(x_1, x_7, x_2, x_3, x_4, x_5, x_6); x_9 = !lean_is_exclusive(x_8); if (x_9 == 0) { -lean_object* x_10; lean_object* x_11; +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; x_10 = lean_ctor_get(x_8, 0); -lean_inc(x_7); -x_11 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_11, 0, x_7); -lean_ctor_set(x_11, 1, x_10); -lean_ctor_set_tag(x_8, 1); -lean_ctor_set(x_8, 0, x_11); +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_array_to_list(lean_box(0), x_12); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_11); +lean_ctor_set(x_8, 0, x_14); return x_8; } else { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_12 = lean_ctor_get(x_8, 0); -x_13 = lean_ctor_get(x_8, 1); -lean_inc(x_13); -lean_inc(x_12); +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_ctor_get(x_8, 0); +x_16 = lean_ctor_get(x_8, 1); +lean_inc(x_16); +lean_inc(x_15); lean_dec(x_8); -lean_inc(x_7); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_7); -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; +x_17 = lean_ctor_get(x_15, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_15, 1); +lean_inc(x_18); +lean_dec(x_15); +x_19 = lean_array_to_list(lean_box(0), 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_17); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_16); +return x_21; } } } -LEAN_EXPORT lean_object* l_Lean_Core_withCurrHeartbeats___at_Lean_Meta_SynthInstance_main___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_EXPORT lean_object* l_List_forIn_loop___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessLevels___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) { _start: { -lean_object* x_7; lean_object* x_8; -x_7 = lean_apply_2(x_1, x_2, x_3); -x_8 = l___private_Lean_CoreM_0__Lean_Core_withCurrHeartbeatsImp___rarg(x_7, x_4, x_5, x_6); -if (lean_obj_tag(x_8) == 0) +lean_object* x_8; +x_8 = l_List_forIn_loop___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessLevels___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_8; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessLevels___boxed(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_9; -x_9 = !lean_is_exclusive(x_8); -if (x_9 == 0) +lean_object* x_7; +x_7 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessLevels(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); +return x_7; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessArgs___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: { -return x_8; +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +lean_inc(x_5); +x_11 = lean_array_fset(x_1, x_2, x_5); +x_12 = lean_expr_instantiate1(x_3, x_5); +lean_dec(x_5); +lean_dec(x_3); +x_13 = lean_unsigned_to_nat(1u); +x_14 = lean_nat_add(x_2, x_13); +lean_dec(x_2); +x_15 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessArgs(x_12, x_14, x_11, x_4, x_6, x_7, x_8, x_9, x_10); +return x_15; } -else +} +static lean_object* _init_l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessArgs___closed__1() { +_start: { -lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = lean_ctor_get(x_8, 0); -x_11 = lean_ctor_get(x_8, 1); -lean_inc(x_11); -lean_inc(x_10); +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("type class resolution failed, insufficient number of arguments", 62); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessArgs___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessArgs___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessArgs(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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; +x_10 = lean_array_get_size(x_3); +x_11 = lean_nat_dec_lt(x_2, x_10); +lean_dec(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_2); +lean_dec(x_1); x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_10); -lean_ctor_set(x_12, 1, x_11); +lean_ctor_set(x_12, 0, x_3); +lean_ctor_set(x_12, 1, x_9); return x_12; } +else +{ +lean_object* x_13; +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_13 = lean_whnf(x_1, x_5, x_6, x_7, x_8, x_9); +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) == 7) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +x_17 = lean_ctor_get(x_14, 2); +lean_inc(x_17); +lean_dec(x_14); +x_18 = lean_array_fget(x_3, x_2); +x_19 = l_Array_contains___at___private_Lean_Meta_FunInfo_0__Lean_Meta_collectDeps_visit___spec__2(x_4, x_2); +if (x_19 == 0) +{ +lean_object* x_20; +lean_dec(x_16); +x_20 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessArgs___lambda__1(x_3, x_2, x_17, x_4, x_18, x_5, x_6, x_7, x_8, x_15); +return x_20; } else { -uint8_t x_13; -x_13 = !lean_is_exclusive(x_8); -if (x_13 == 0) +lean_object* x_21; uint8_t 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_18); +x_21 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_21, 0, x_16); +x_22 = 0; +x_23 = lean_box(0); +lean_inc(x_5); +x_24 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkFreshExprMVarImpl(x_21, x_22, x_23, x_5, x_6, x_7, x_8, 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_Meta_SynthInstance_0__Lean_Meta_preprocessArgs___lambda__1(x_3, x_2, x_17, x_4, x_25, x_5, x_6, x_7, x_8, x_26); +return x_27; +} +} +else { -return x_8; +lean_object* x_28; lean_object* x_29; lean_object* x_30; +lean_dec(x_14); +lean_dec(x_3); +lean_dec(x_2); +x_28 = lean_ctor_get(x_13, 1); +lean_inc(x_28); +lean_dec(x_13); +x_29 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessArgs___closed__2; +x_30 = l_Lean_throwError___at_Lean_Meta_collectForwardDeps___spec__1(x_29, 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); +return x_30; +} } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_8, 0); -x_15 = lean_ctor_get(x_8, 1); -lean_inc(x_15); -lean_inc(x_14); +uint8_t x_31; lean_dec(x_8); -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_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +x_31 = !lean_is_exclusive(x_13); +if (x_31 == 0) +{ +return x_13; } +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_13, 0); +x_33 = lean_ctor_get(x_13, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_13); +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; } } } -static lean_object* _init_l_Lean_Meta_SynthInstance_main___lambda__1___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("failed to synthesize", 20); -return x_1; } } -static lean_object* _init_l_Lean_Meta_SynthInstance_main___lambda__1___closed__2() { +LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessArgs___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_1; lean_object* x_2; -x_1 = l_Lean_Meta_SynthInstance_main___lambda__1___closed__1; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +lean_object* x_11; +x_11 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessArgs___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_4); +return x_11; } } -static lean_object* _init_l_Lean_Meta_SynthInstance_main___lambda__1___closed__3() { +LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessArgs___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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_1; -x_1 = lean_mk_string_from_bytes("\n", 1); -return x_1; +lean_object* x_10; +x_10 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessArgs(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_4); +return x_10; } } -static lean_object* _init_l_Lean_Meta_SynthInstance_main___lambda__1___closed__4() { +static lean_object* _init_l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessOutParam___lambda__1___closed__1() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_SynthInstance_main___lambda__1___closed__3; -x_2 = l_Lean_stringToMessageData(x_1); +x_1 = l_Lean_levelZero; +x_2 = l_Lean_Expr_sort___override(x_1); return x_2; } } -static lean_object* _init_l_Lean_Meta_SynthInstance_main___lambda__1___closed__5() { +LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessOutParam___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_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = lean_box(0); -x_2 = l_Lean_Meta_SynthInstance_instInhabitedInstance___closed__2; -x_3 = l_Lean_Meta_SynthInstance_MkTableKey_State_lmap___default___closed__1; -x_4 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_4, 0, x_1); -lean_ctor_set(x_4, 1, x_2); -lean_ctor_set(x_4, 2, x_2); -lean_ctor_set(x_4, 3, x_3); -return x_4; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_main___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; +x_9 = l_Lean_Expr_getAppFn(x_3); +if (lean_obj_tag(x_9) == 4) { -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_object* x_18; lean_object* x_19; 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_9 = 0; -x_10 = lean_box(0); -lean_inc(x_4); -x_11 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkFreshExprMVarImpl(x_1, x_9, x_10, x_4, x_5, x_6, x_7, x_8); -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_2); -x_14 = l_Lean_Meta_SynthInstance_mkTableKey___at_Lean_Meta_SynthInstance_main___spec__1(x_2, x_4, x_5, x_6, x_7, x_13); -x_15 = lean_ctor_get(x_14, 0); +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_st_ref_get(x_7, x_8); +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); -x_16 = lean_ctor_get(x_14, 1); -lean_inc(x_16); -if (lean_is_exclusive(x_14)) { - lean_ctor_release(x_14, 0); - lean_ctor_release(x_14, 1); - x_17 = x_14; -} else { - lean_dec_ref(x_14); - x_17 = lean_box(0); +lean_dec(x_13); +x_16 = l_Lean_getOutParamPositions_x3f(x_15, x_10); +if (lean_obj_tag(x_16) == 0) +{ +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_ctor_set(x_11, 0, x_1); +return x_11; } -x_33 = lean_ctor_get(x_6, 2); -lean_inc(x_33); -x_34 = l_Lean_Meta_SynthInstance_getMaxHeartbeats(x_33); -lean_dec(x_33); -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_3); -lean_ctor_set(x_35, 1, x_34); -x_36 = l_Lean_Meta_SynthInstance_main___lambda__1___closed__5; -x_37 = lean_st_mk_ref(x_36, x_16); -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_st_ref_get(x_5, 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_ctor_get(x_41, 0); -lean_inc(x_43); -lean_dec(x_41); -x_44 = lean_box(1); +else +{ +lean_object* x_17; uint8_t x_18; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +lean_dec(x_16); +x_18 = l_Array_isEmpty___rarg(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_free_object(x_11); +lean_dec(x_1); +x_19 = lean_unsigned_to_nat(0u); +x_20 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_3, x_19); +x_21 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessOutParam___lambda__1___closed__1; +lean_inc(x_20); +x_22 = lean_mk_array(x_20, x_21); +x_23 = lean_unsigned_to_nat(1u); +x_24 = lean_nat_sub(x_20, x_23); +lean_dec(x_20); +x_25 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_3, x_22, x_24); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_38); -lean_inc(x_35); -x_45 = l_Lean_Meta_SynthInstance_newSubgoal(x_43, x_15, x_12, x_44, x_35, x_38, x_4, x_5, x_6, x_7, x_42); -if (lean_obj_tag(x_45) == 0) +lean_inc(x_9); +x_26 = lean_infer_type(x_9, x_4, x_5, x_6, x_7, x_14); +if (lean_obj_tag(x_26) == 0) { -lean_object* x_46; lean_object* x_47; -x_46 = lean_ctor_get(x_45, 1); -lean_inc(x_46); -lean_dec(x_45); +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); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_38); -x_47 = l_Lean_Meta_SynthInstance_synth(x_35, x_38, x_4, x_5, x_6, x_7, x_46); -if (lean_obj_tag(x_47) == 0) -{ -lean_object* x_48; lean_object* x_49; lean_object* x_50; uint8_t x_51; +x_29 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessArgs(x_27, x_19, x_25, x_17, x_4, x_5, x_6, x_7, x_28); lean_dec(x_17); +if (lean_obj_tag(x_29) == 0) +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; uint8_t x_34; uint8_t x_35; lean_object* x_36; +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 = l_Lean_mkAppN(x_9, x_30); +x_33 = 0; +x_34 = 1; +x_35 = 1; +x_36 = l_Lean_Meta_mkForallFVars(x_2, x_32, x_33, x_34, x_35, x_4, x_5, x_6, x_7, x_31); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_2); -x_48 = lean_ctor_get(x_47, 0); -lean_inc(x_48); -x_49 = lean_ctor_get(x_47, 1); -lean_inc(x_49); -lean_dec(x_47); -x_50 = lean_st_ref_get(x_38, x_49); -lean_dec(x_38); -x_51 = !lean_is_exclusive(x_50); -if (x_51 == 0) +if (lean_obj_tag(x_36) == 0) { -lean_object* x_52; -x_52 = lean_ctor_get(x_50, 0); -lean_dec(x_52); -lean_ctor_set(x_50, 0, x_48); -return x_50; +uint8_t x_37; +x_37 = !lean_is_exclusive(x_36); +if (x_37 == 0) +{ +return x_36; } else { -lean_object* x_53; lean_object* x_54; -x_53 = lean_ctor_get(x_50, 1); -lean_inc(x_53); -lean_dec(x_50); -x_54 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_54, 0, x_48); -lean_ctor_set(x_54, 1, x_53); -return x_54; +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_inc(x_38); +lean_dec(x_36); +x_40 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_40, 0, x_38); +lean_ctor_set(x_40, 1, x_39); +return x_40; } } else { -lean_object* x_55; lean_object* x_56; -lean_dec(x_38); -x_55 = lean_ctor_get(x_47, 0); -lean_inc(x_55); -x_56 = lean_ctor_get(x_47, 1); -lean_inc(x_56); -lean_dec(x_47); -x_18 = x_55; -x_19 = x_56; -goto block_32; -} +uint8_t x_41; +x_41 = !lean_is_exclusive(x_36); +if (x_41 == 0) +{ +return x_36; } else { -lean_object* x_57; lean_object* x_58; -lean_dec(x_38); -lean_dec(x_35); -x_57 = lean_ctor_get(x_45, 0); -lean_inc(x_57); -x_58 = lean_ctor_get(x_45, 1); -lean_inc(x_58); -lean_dec(x_45); -x_18 = x_57; -x_19 = x_58; -goto block_32; +lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_36, 0); +x_43 = lean_ctor_get(x_36, 1); +lean_inc(x_43); +lean_inc(x_42); +lean_dec(x_36); +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; } -block_32: -{ -uint8_t x_20; -x_20 = l_Lean_Exception_isMaxHeartbeat(x_18); -if (x_20 == 0) +} +} +else { -lean_object* x_21; +uint8_t x_45; +lean_dec(x_9); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); -if (lean_is_scalar(x_17)) { - x_21 = lean_alloc_ctor(1, 2, 0); -} else { - x_21 = x_17; - lean_ctor_set_tag(x_21, 1); +x_45 = !lean_is_exclusive(x_29); +if (x_45 == 0) +{ +return x_29; +} +else +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_46 = lean_ctor_get(x_29, 0); +x_47 = lean_ctor_get(x_29, 1); +lean_inc(x_47); +lean_inc(x_46); +lean_dec(x_29); +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_ctor_set(x_21, 0, x_18); -lean_ctor_set(x_21, 1, x_19); -return x_21; } 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_49; +lean_dec(x_25); lean_dec(x_17); -x_22 = l_Lean_indentExpr(x_2); -x_23 = l_Lean_Meta_SynthInstance_main___lambda__1___closed__2; -x_24 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_22); -x_25 = l_Lean_Meta_SynthInstance_main___lambda__1___closed__4; -x_26 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_26, 0, x_24); -lean_ctor_set(x_26, 1, x_25); -x_27 = l_Lean_Exception_toMessageData(x_18); -x_28 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -x_29 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; -x_30 = lean_alloc_ctor(7, 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_Meta_SynthInstance_main___spec__2(x_30, x_4, x_5, x_6, x_7, x_19); +lean_dec(x_9); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -return x_31; -} -} -} +lean_dec(x_2); +x_49 = !lean_is_exclusive(x_26); +if (x_49 == 0) +{ +return x_26; } -LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_main(lean_object* x_1, 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: +else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; -lean_inc(x_1); -x_8 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_8, 0, x_1); -x_9 = lean_alloc_closure((void*)(l_Lean_Meta_SynthInstance_main___lambda__1), 8, 3); -lean_closure_set(x_9, 0, x_8); -lean_closure_set(x_9, 1, x_1); -lean_closure_set(x_9, 2, x_2); -x_10 = l_Lean_Core_withCurrHeartbeats___at_Lean_Meta_SynthInstance_main___spec__3(x_9, x_3, x_4, x_5, x_6, x_7); -return x_10; +lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_50 = lean_ctor_get(x_26, 0); +x_51 = lean_ctor_get(x_26, 1); +lean_inc(x_51); +lean_inc(x_50); +lean_dec(x_26); +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_EXPORT lean_object* l_Lean_Meta_SynthInstance_mkTableKey___at_Lean_Meta_SynthInstance_main___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) { -_start: +} +else { -lean_object* x_7; -x_7 = l_Lean_Meta_SynthInstance_mkTableKey___at_Lean_Meta_SynthInstance_main___spec__1(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_17); +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_7; +lean_ctor_set(x_11, 0, x_1); +return x_11; } } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_SynthInstance_main___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: +} +else { -lean_object* x_7; -x_7 = l_Lean_throwError___at_Lean_Meta_SynthInstance_main___spec__2(x_1, x_2, x_3, x_4, x_5, x_6); +lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_53 = lean_ctor_get(x_11, 0); +x_54 = lean_ctor_get(x_11, 1); +lean_inc(x_54); +lean_inc(x_53); +lean_dec(x_11); +x_55 = lean_ctor_get(x_53, 0); +lean_inc(x_55); +lean_dec(x_53); +x_56 = l_Lean_getOutParamPositions_x3f(x_55, x_10); +if (lean_obj_tag(x_56) == 0) +{ +lean_object* x_57; +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_7; -} +x_57 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_57, 0, x_1); +lean_ctor_set(x_57, 1, x_54); +return x_57; } -LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocess___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: +else { -lean_object* x_8; +lean_object* x_58; uint8_t x_59; +x_58 = lean_ctor_get(x_56, 0); +lean_inc(x_58); +lean_dec(x_56); +x_59 = l_Array_isEmpty___rarg(x_58); +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_dec(x_1); +x_60 = lean_unsigned_to_nat(0u); +x_61 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_3, x_60); +x_62 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessOutParam___lambda__1___closed__1; +lean_inc(x_61); +x_63 = lean_mk_array(x_61, x_62); +x_64 = lean_unsigned_to_nat(1u); +x_65 = lean_nat_sub(x_61, x_64); +lean_dec(x_61); +x_66 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_3, x_63, x_65); +lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_3); -x_8 = lean_whnf(x_2, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_8) == 0) -{ -lean_object* x_9; lean_object* x_10; uint8_t x_11; uint8_t x_12; uint8_t x_13; lean_object* x_14; -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_dec(x_8); -x_11 = 0; -x_12 = 1; -x_13 = 1; -x_14 = l_Lean_Meta_mkForallFVars(x_1, x_9, x_11, x_12, x_13, x_3, x_4, x_5, x_6, x_10); +x_67 = lean_infer_type(x_9, x_4, x_5, x_6, x_7, x_54); +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_67, 0); +lean_inc(x_68); +x_69 = lean_ctor_get(x_67, 1); +lean_inc(x_69); +lean_dec(x_67); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_70 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessArgs(x_68, x_60, x_66, x_58, x_4, x_5, x_6, x_7, x_69); +lean_dec(x_58); +if (lean_obj_tag(x_70) == 0) +{ +lean_object* x_71; lean_object* x_72; lean_object* x_73; uint8_t x_74; uint8_t x_75; uint8_t x_76; lean_object* x_77; +x_71 = lean_ctor_get(x_70, 0); +lean_inc(x_71); +x_72 = lean_ctor_get(x_70, 1); +lean_inc(x_72); +lean_dec(x_70); +x_73 = l_Lean_mkAppN(x_9, x_71); +x_74 = 0; +x_75 = 1; +x_76 = 1; +x_77 = l_Lean_Meta_mkForallFVars(x_2, x_73, x_74, x_75, x_76, x_4, x_5, x_6, x_7, x_72); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -return x_14; +if (lean_obj_tag(x_77) == 0) +{ +lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_78 = lean_ctor_get(x_77, 0); +lean_inc(x_78); +x_79 = lean_ctor_get(x_77, 1); +lean_inc(x_79); +if (lean_is_exclusive(x_77)) { + lean_ctor_release(x_77, 0); + lean_ctor_release(x_77, 1); + x_80 = x_77; +} else { + lean_dec_ref(x_77); + x_80 = lean_box(0); +} +if (lean_is_scalar(x_80)) { + x_81 = lean_alloc_ctor(0, 2, 0); +} else { + x_81 = x_80; +} +lean_ctor_set(x_81, 0, x_78); +lean_ctor_set(x_81, 1, x_79); +return x_81; } else { -uint8_t x_15; -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -x_15 = !lean_is_exclusive(x_8); -if (x_15 == 0) -{ -return x_8; +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; +x_82 = lean_ctor_get(x_77, 0); +lean_inc(x_82); +x_83 = lean_ctor_get(x_77, 1); +lean_inc(x_83); +if (lean_is_exclusive(x_77)) { + lean_ctor_release(x_77, 0); + lean_ctor_release(x_77, 1); + x_84 = x_77; +} else { + lean_dec_ref(x_77); + x_84 = lean_box(0); +} +if (lean_is_scalar(x_84)) { + x_85 = lean_alloc_ctor(1, 2, 0); +} else { + x_85 = x_84; +} +lean_ctor_set(x_85, 0, x_82); +lean_ctor_set(x_85, 1, x_83); +return x_85; +} } else { -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_8, 0); -x_17 = lean_ctor_get(x_8, 1); -lean_inc(x_17); -lean_inc(x_16); -lean_dec(x_8); -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_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; +lean_dec(x_9); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +x_86 = lean_ctor_get(x_70, 0); +lean_inc(x_86); +x_87 = lean_ctor_get(x_70, 1); +lean_inc(x_87); +if (lean_is_exclusive(x_70)) { + lean_ctor_release(x_70, 0); + lean_ctor_release(x_70, 1); + x_88 = x_70; +} else { + lean_dec_ref(x_70); + x_88 = lean_box(0); } +if (lean_is_scalar(x_88)) { + x_89 = lean_alloc_ctor(1, 2, 0); +} else { + x_89 = x_88; } +lean_ctor_set(x_89, 0, x_86); +lean_ctor_set(x_89, 1, x_87); +return x_89; } } -static lean_object* _init_l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocess___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocess___lambda__1), 7, 0); -return x_1; -} +lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; +lean_dec(x_66); +lean_dec(x_58); +lean_dec(x_9); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +x_90 = lean_ctor_get(x_67, 0); +lean_inc(x_90); +x_91 = lean_ctor_get(x_67, 1); +lean_inc(x_91); +if (lean_is_exclusive(x_67)) { + lean_ctor_release(x_67, 0); + lean_ctor_release(x_67, 1); + x_92 = x_67; +} else { + lean_dec_ref(x_67); + x_92 = lean_box(0); } -LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocess(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___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocess___closed__1; -x_8 = l_Lean_Meta_forallTelescopeReducing___at_Lean_Meta_getParamNames___spec__2___rarg(x_1, x_7, x_2, x_3, x_4, x_5, x_6); -return x_8; +if (lean_is_scalar(x_92)) { + x_93 = lean_alloc_ctor(1, 2, 0); +} else { + x_93 = x_92; } +lean_ctor_set(x_93, 0, x_90); +lean_ctor_set(x_93, 1, x_91); +return x_93; } -LEAN_EXPORT lean_object* l_List_forIn_loop___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessLevels___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) { -_start: -{ -if (lean_obj_tag(x_1) == 0) -{ -lean_object* x_8; -x_8 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_8, 0, x_2); -lean_ctor_set(x_8, 1, x_7); -return x_8; } else { -lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_9 = lean_ctor_get(x_1, 0); -lean_inc(x_9); -x_10 = lean_ctor_get(x_1, 1); -lean_inc(x_10); -lean_dec(x_1); -x_11 = !lean_is_exclusive(x_2); -if (x_11 == 0) -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_12 = lean_ctor_get(x_2, 0); -x_13 = lean_ctor_get(x_2, 1); -x_14 = l_Lean_instantiateLevelMVars___at_Lean_Meta_normalizeLevel___spec__1(x_9, x_3, x_4, x_5, x_6, x_7); -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_Level_hasMVar(x_15); -if (x_17 == 0) -{ -lean_object* x_18; -x_18 = lean_array_push(x_13, x_15); -lean_ctor_set(x_2, 1, x_18); -x_1 = x_10; -x_7 = x_16; -goto _start; +lean_object* x_94; +lean_dec(x_58); +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); +x_94 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_94, 0, x_1); +lean_ctor_set(x_94, 1, x_54); +return x_94; } -else -{ -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; lean_object* x_25; -lean_dec(x_15); -lean_dec(x_12); -x_20 = l_Lean_Meta_mkFreshLevelMVar(x_3, x_4, x_5, x_6, 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); -x_23 = lean_array_push(x_13, x_21); -x_24 = 1; -x_25 = lean_box(x_24); -lean_ctor_set(x_2, 1, x_23); -lean_ctor_set(x_2, 0, x_25); -x_1 = x_10; -x_7 = x_22; -goto _start; } } -else -{ -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 = lean_ctor_get(x_2, 0); -x_28 = lean_ctor_get(x_2, 1); -lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_2); -x_29 = l_Lean_instantiateLevelMVars___at_Lean_Meta_normalizeLevel___spec__1(x_9, x_3, x_4, x_5, x_6, x_7); -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 = l_Lean_Level_hasMVar(x_30); -if (x_32 == 0) -{ -lean_object* x_33; lean_object* x_34; -x_33 = lean_array_push(x_28, x_30); -x_34 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_34, 0, x_27); -lean_ctor_set(x_34, 1, x_33); -x_1 = x_10; -x_2 = x_34; -x_7 = x_31; -goto _start; } else { -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_dec(x_30); -lean_dec(x_27); -x_36 = l_Lean_Meta_mkFreshLevelMVar(x_3, x_4, x_5, x_6, x_31); -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_array_push(x_28, x_37); -x_40 = 1; -x_41 = lean_box(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_39); -x_1 = x_10; -x_2 = x_42; -x_7 = x_38; -goto _start; -} -} +lean_object* x_95; +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); +x_95 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_95, 0, x_1); +lean_ctor_set(x_95, 1, x_8); +return x_95; } } } -static lean_object* _init_l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessLevels___closed__1() { +LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessOutParam(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_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = 0; -x_2 = l_Lean_Meta_SynthInstance_instInhabitedInstance___closed__2; -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; +lean_object* x_7; lean_object* x_8; +lean_inc(x_1); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessOutParam___lambda__1), 8, 1); +lean_closure_set(x_7, 0, x_1); +x_8 = l_Lean_Meta_forallTelescope___at___private_Lean_Meta_InferType_0__Lean_Meta_inferForallType___spec__2___rarg(x_1, x_7, x_2, x_3, x_4, x_5, x_6); +return x_8; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessLevels(lean_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_Meta_synthInstance_x3f_assignOutParams___lambda__1(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) { _start: { -lean_object* x_7; lean_object* x_8; uint8_t x_9; -x_7 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessLevels___closed__1; -x_8 = l_List_forIn_loop___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessLevels___spec__1(x_1, x_7, x_2, x_3, x_4, x_5, x_6); -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; -x_10 = lean_ctor_get(x_8, 0); -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_array_to_list(lean_box(0), x_12); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_13); -lean_ctor_set(x_14, 1, x_11); -lean_ctor_set(x_8, 0, x_14); -return x_8; -} -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; -x_15 = lean_ctor_get(x_8, 0); -x_16 = lean_ctor_get(x_8, 1); -lean_inc(x_16); -lean_inc(x_15); -lean_dec(x_8); -x_17 = lean_ctor_get(x_15, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_15, 1); -lean_inc(x_18); -lean_dec(x_15); -x_19 = lean_array_to_list(lean_box(0), 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_17); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_20); -lean_ctor_set(x_21, 1, x_16); -return x_21; -} +lean_object* x_8; lean_object* x_9; +x_8 = lean_box(x_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); +return x_9; } } -LEAN_EXPORT lean_object* l_List_forIn_loop___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessLevels___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) { +static lean_object* _init_l_Lean_Meta_synthInstance_x3f_assignOutParams___closed__1() { _start: { -lean_object* x_8; -x_8 = l_List_forIn_loop___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessLevels___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_8; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(" result type", 12); +return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessLevels___boxed(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_Meta_synthInstance_x3f_assignOutParams___closed__2() { _start: { -lean_object* x_7; -x_7 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessLevels(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); -return x_7; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_synthInstance_x3f_assignOutParams___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessArgs___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) { +static lean_object* _init_l_Lean_Meta_synthInstance_x3f_assignOutParams___closed__3() { _start: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; -lean_inc(x_5); -x_11 = lean_array_fset(x_1, x_2, x_5); -x_12 = lean_expr_instantiate1(x_3, x_5); -lean_dec(x_5); -lean_dec(x_3); -x_13 = lean_unsigned_to_nat(1u); -x_14 = lean_nat_add(x_2, x_13); -lean_dec(x_2); -x_15 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessArgs(x_12, x_14, x_11, x_4, x_6, x_7, x_8, x_9, x_10); -return x_15; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__4; +x_2 = l_Lean_Meta_synthInstance_x3f_assignOutParams___closed__2; +x_3 = lean_alloc_ctor(7, 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_SynthInstance_0__Lean_Meta_preprocessArgs___closed__1() { +static lean_object* _init_l_Lean_Meta_synthInstance_x3f_assignOutParams___closed__4() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("type class resolution failed, insufficient number of arguments", 62); +x_1 = lean_mk_string_from_bytes("\nis not definitionally equal to", 31); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessArgs___closed__2() { +static lean_object* _init_l_Lean_Meta_synthInstance_x3f_assignOutParams___closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessArgs___closed__1; +x_1 = l_Lean_Meta_synthInstance_x3f_assignOutParams___closed__4; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessArgs(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_Meta_synthInstance_x3f_assignOutParams(lean_object* x_1, 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; uint8_t x_11; -x_10 = lean_array_get_size(x_3); -x_11 = lean_nat_dec_lt(x_2, x_10); -lean_dec(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_2); -lean_dec(x_1); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_3); -lean_ctor_set(x_12, 1, x_9); -return x_12; -} -else -{ -lean_object* x_13; -lean_inc(x_8); -lean_inc(x_7); +lean_object* x_8; lean_inc(x_6); lean_inc(x_5); -x_13 = lean_whnf(x_1, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_13) == 0) +lean_inc(x_4); +lean_inc(x_3); +x_8 = lean_infer_type(x_2, x_3, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_8) == 0) { -lean_object* x_14; -x_14 = lean_ctor_get(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; lean_object* x_16; uint8_t x_17; +x_9 = lean_ctor_get(x_3, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_8, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_8, 1); +lean_inc(x_11); +lean_dec(x_8); +x_12 = lean_ctor_get(x_3, 1); +lean_inc(x_12); +x_13 = lean_ctor_get(x_3, 2); +lean_inc(x_13); +x_14 = lean_ctor_get(x_3, 3); lean_inc(x_14); -if (lean_obj_tag(x_14) == 7) -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_15 = lean_ctor_get(x_13, 1); +x_15 = lean_ctor_get(x_3, 4); lean_inc(x_15); -lean_dec(x_13); -x_16 = lean_ctor_get(x_14, 1); +x_16 = lean_ctor_get(x_3, 5); lean_inc(x_16); -x_17 = lean_ctor_get(x_14, 2); -lean_inc(x_17); -lean_dec(x_14); -x_18 = lean_array_fget(x_3, x_2); -x_19 = l_Array_contains___at___private_Lean_Meta_FunInfo_0__Lean_Meta_collectDeps_visit___spec__2(x_4, x_2); -if (x_19 == 0) -{ -lean_object* x_20; -lean_dec(x_16); -x_20 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessArgs___lambda__1(x_3, x_2, x_17, x_4, x_18, x_5, x_6, x_7, x_8, x_15); -return x_20; -} -else +x_17 = !lean_is_exclusive(x_9); +if (x_17 == 0) { -lean_object* x_21; uint8_t 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_18); -x_21 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_21, 0, x_16); -x_22 = 0; -x_23 = lean_box(0); +uint8_t x_18; uint8_t x_19; lean_object* x_20; lean_object* x_21; +x_18 = 1; +x_19 = 1; +lean_ctor_set_uint8(x_9, 5, x_18); +lean_ctor_set_uint8(x_9, 10, x_19); +x_20 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_20, 0, x_9); +lean_ctor_set(x_20, 1, x_12); +lean_ctor_set(x_20, 2, x_13); +lean_ctor_set(x_20, 3, x_14); +lean_ctor_set(x_20, 4, x_15); +lean_ctor_set(x_20, 5, x_16); +lean_inc(x_6); lean_inc(x_5); -x_24 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkFreshExprMVarImpl(x_21, x_22, x_23, x_5, x_6, x_7, x_8, 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_Meta_SynthInstance_0__Lean_Meta_preprocessArgs___lambda__1(x_3, x_2, x_17, x_4, x_25, x_5, x_6, x_7, x_8, x_26); -return x_27; -} -} -else +lean_inc(x_4); +lean_inc(x_10); +lean_inc(x_1); +x_21 = l_Lean_Meta_isExprDefEq(x_1, x_10, x_20, x_4, x_5, x_6, x_11); +if (lean_obj_tag(x_21) == 0) { -lean_object* x_28; lean_object* x_29; lean_object* x_30; -lean_dec(x_14); -lean_dec(x_3); -lean_dec(x_2); -x_28 = lean_ctor_get(x_13, 1); -lean_inc(x_28); -lean_dec(x_13); -x_29 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessArgs___closed__2; -x_30 = l_Lean_throwError___at_Lean_Meta_collectForwardDeps___spec__1(x_29, 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); -return x_30; -} -} -else +lean_object* x_22; uint8_t x_23; +x_22 = lean_ctor_get(x_21, 0); +lean_inc(x_22); +x_23 = lean_unbox(x_22); +if (x_23 == 0) { -uint8_t x_31; -lean_dec(x_8); -lean_dec(x_7); +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_24 = lean_ctor_get(x_21, 1); +lean_inc(x_24); +lean_dec(x_21); +x_25 = l_Lean_Meta_SynthInstance_newSubgoal___closed__1; +x_26 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_25, x_3, x_4, x_5, x_6, 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) +{ +uint8_t x_29; +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_31 = !lean_is_exclusive(x_13); -if (x_31 == 0) +lean_dec(x_1); +x_29 = !lean_is_exclusive(x_26); +if (x_29 == 0) { -return x_13; +lean_object* x_30; +x_30 = lean_ctor_get(x_26, 0); +lean_dec(x_30); +lean_ctor_set(x_26, 0, x_22); +return x_26; } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_13, 0); -x_33 = lean_ctor_get(x_13, 1); -lean_inc(x_33); -lean_inc(x_32); -lean_dec(x_13); -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_object* x_31; lean_object* x_32; +x_31 = lean_ctor_get(x_26, 1); +lean_inc(x_31); +lean_dec(x_26); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_22); +lean_ctor_set(x_32, 1, x_31); +return x_32; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessArgs___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: +else { -lean_object* x_11; -x_11 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessArgs___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +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; uint8_t x_44; +x_33 = lean_ctor_get(x_26, 1); +lean_inc(x_33); +lean_dec(x_26); +x_34 = l_Lean_indentExpr(x_10); +x_35 = l_Lean_Meta_synthInstance_x3f_assignOutParams___closed__3; +x_36 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_34); +x_37 = l_Lean_Meta_synthInstance_x3f_assignOutParams___closed__5; +x_38 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_38, 0, x_36); +lean_ctor_set(x_38, 1, x_37); +x_39 = l_Lean_indentExpr(x_1); +x_40 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_40, 0, x_38); +lean_ctor_set(x_40, 1, x_39); +x_41 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; +x_42 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_42, 0, x_40); +lean_ctor_set(x_42, 1, x_41); +x_43 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_25, x_42, x_3, x_4, x_5, x_6, x_33); +lean_dec(x_6); +lean_dec(x_5); lean_dec(x_4); -return x_11; -} +lean_dec(x_3); +x_44 = !lean_is_exclusive(x_43); +if (x_44 == 0) +{ +lean_object* x_45; +x_45 = lean_ctor_get(x_43, 0); +lean_dec(x_45); +lean_ctor_set(x_43, 0, x_22); +return x_43; } -LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessArgs___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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: +else { -lean_object* x_10; -x_10 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessArgs(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_4); -return x_10; -} +lean_object* x_46; lean_object* x_47; +x_46 = lean_ctor_get(x_43, 1); +lean_inc(x_46); +lean_dec(x_43); +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_22); +lean_ctor_set(x_47, 1, x_46); +return x_47; } -static lean_object* _init_l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessOutParam___lambda__1___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_levelZero; -x_2 = l_Lean_Expr_sort___override(x_1); -return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessOutParam___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; -x_9 = l_Lean_Expr_getAppFn(x_3); -if (lean_obj_tag(x_9) == 4) -{ -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_st_ref_get(x_7, x_8); -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); -x_16 = l_Lean_getOutParamPositions_x3f(x_15, x_10); -if (lean_obj_tag(x_16) == 0) +else { -lean_dec(x_9); -lean_dec(x_7); +uint8_t x_48; +lean_dec(x_10); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -lean_ctor_set(x_11, 0, x_1); -return x_11; +lean_dec(x_1); +x_48 = !lean_is_exclusive(x_21); +if (x_48 == 0) +{ +lean_object* x_49; +x_49 = lean_ctor_get(x_21, 0); +lean_dec(x_49); +return x_21; } else { -lean_object* x_17; uint8_t x_18; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -lean_dec(x_16); -x_18 = l_Array_isEmpty___rarg(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_free_object(x_11); -lean_dec(x_1); -x_19 = lean_unsigned_to_nat(0u); -x_20 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_3, x_19); -x_21 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessOutParam___lambda__1___closed__1; -lean_inc(x_20); -x_22 = lean_mk_array(x_20, x_21); -x_23 = lean_unsigned_to_nat(1u); -x_24 = lean_nat_sub(x_20, x_23); -lean_dec(x_20); -x_25 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_3, x_22, x_24); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_9); -x_26 = lean_infer_type(x_9, x_4, x_5, x_6, x_7, x_14); -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); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_29 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessArgs(x_27, x_19, x_25, x_17, x_4, x_5, x_6, x_7, x_28); -lean_dec(x_17); -if (lean_obj_tag(x_29) == 0) +lean_object* x_50; lean_object* x_51; +x_50 = lean_ctor_get(x_21, 1); +lean_inc(x_50); +lean_dec(x_21); +x_51 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_51, 0, x_22); +lean_ctor_set(x_51, 1, x_50); +return x_51; +} +} +} +else { -lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; uint8_t x_34; uint8_t x_35; lean_object* x_36; -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 = l_Lean_mkAppN(x_9, x_30); -x_33 = 0; -x_34 = 1; -x_35 = 1; -x_36 = l_Lean_Meta_mkForallFVars(x_2, x_32, x_33, x_34, x_35, x_4, x_5, x_6, x_7, x_31); -lean_dec(x_7); +uint8_t x_52; +lean_dec(x_10); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -if (lean_obj_tag(x_36) == 0) -{ -uint8_t x_37; -x_37 = !lean_is_exclusive(x_36); -if (x_37 == 0) +lean_dec(x_3); +lean_dec(x_1); +x_52 = !lean_is_exclusive(x_21); +if (x_52 == 0) { -return x_36; +return x_21; } 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_inc(x_38); -lean_dec(x_36); -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set(x_40, 1, x_39); -return x_40; +lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_53 = lean_ctor_get(x_21, 0); +x_54 = lean_ctor_get(x_21, 1); +lean_inc(x_54); +lean_inc(x_53); +lean_dec(x_21); +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 { -uint8_t x_41; -x_41 = !lean_is_exclusive(x_36); -if (x_41 == 0) +uint8_t x_56; uint8_t x_57; uint8_t x_58; uint8_t x_59; uint8_t x_60; uint8_t x_61; uint8_t x_62; uint8_t x_63; uint8_t x_64; uint8_t x_65; uint8_t x_66; uint8_t x_67; uint8_t x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; +x_56 = lean_ctor_get_uint8(x_9, 0); +x_57 = lean_ctor_get_uint8(x_9, 1); +x_58 = lean_ctor_get_uint8(x_9, 2); +x_59 = lean_ctor_get_uint8(x_9, 3); +x_60 = lean_ctor_get_uint8(x_9, 4); +x_61 = lean_ctor_get_uint8(x_9, 6); +x_62 = lean_ctor_get_uint8(x_9, 7); +x_63 = lean_ctor_get_uint8(x_9, 8); +x_64 = lean_ctor_get_uint8(x_9, 9); +x_65 = lean_ctor_get_uint8(x_9, 11); +x_66 = lean_ctor_get_uint8(x_9, 12); +lean_dec(x_9); +x_67 = 1; +x_68 = 1; +x_69 = lean_alloc_ctor(0, 0, 13); +lean_ctor_set_uint8(x_69, 0, x_56); +lean_ctor_set_uint8(x_69, 1, x_57); +lean_ctor_set_uint8(x_69, 2, x_58); +lean_ctor_set_uint8(x_69, 3, x_59); +lean_ctor_set_uint8(x_69, 4, x_60); +lean_ctor_set_uint8(x_69, 5, x_67); +lean_ctor_set_uint8(x_69, 6, x_61); +lean_ctor_set_uint8(x_69, 7, x_62); +lean_ctor_set_uint8(x_69, 8, x_63); +lean_ctor_set_uint8(x_69, 9, x_64); +lean_ctor_set_uint8(x_69, 10, x_68); +lean_ctor_set_uint8(x_69, 11, x_65); +lean_ctor_set_uint8(x_69, 12, x_66); +x_70 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_70, 0, x_69); +lean_ctor_set(x_70, 1, x_12); +lean_ctor_set(x_70, 2, x_13); +lean_ctor_set(x_70, 3, x_14); +lean_ctor_set(x_70, 4, x_15); +lean_ctor_set(x_70, 5, x_16); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_10); +lean_inc(x_1); +x_71 = l_Lean_Meta_isExprDefEq(x_1, x_10, x_70, x_4, x_5, x_6, x_11); +if (lean_obj_tag(x_71) == 0) { -return x_36; -} -else +lean_object* x_72; uint8_t x_73; +x_72 = lean_ctor_get(x_71, 0); +lean_inc(x_72); +x_73 = lean_unbox(x_72); +if (x_73 == 0) { -lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_42 = lean_ctor_get(x_36, 0); -x_43 = lean_ctor_get(x_36, 1); -lean_inc(x_43); -lean_inc(x_42); -lean_dec(x_36); -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_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; uint8_t x_78; +x_74 = lean_ctor_get(x_71, 1); +lean_inc(x_74); +lean_dec(x_71); +x_75 = l_Lean_Meta_SynthInstance_newSubgoal___closed__1; +x_76 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_75, x_3, x_4, x_5, x_6, x_74); +x_77 = lean_ctor_get(x_76, 0); +lean_inc(x_77); +x_78 = lean_unbox(x_77); +lean_dec(x_77); +if (x_78 == 0) +{ +lean_object* x_79; lean_object* x_80; lean_object* x_81; +lean_dec(x_10); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_79 = lean_ctor_get(x_76, 1); +lean_inc(x_79); +if (lean_is_exclusive(x_76)) { + lean_ctor_release(x_76, 0); + lean_ctor_release(x_76, 1); + x_80 = x_76; +} else { + lean_dec_ref(x_76); + x_80 = lean_box(0); } +if (lean_is_scalar(x_80)) { + x_81 = lean_alloc_ctor(0, 2, 0); +} else { + x_81 = x_80; } +lean_ctor_set(x_81, 0, x_72); +lean_ctor_set(x_81, 1, x_79); +return x_81; } else { -uint8_t x_45; -lean_dec(x_9); -lean_dec(x_7); +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_82 = lean_ctor_get(x_76, 1); +lean_inc(x_82); +lean_dec(x_76); +x_83 = l_Lean_indentExpr(x_10); +x_84 = l_Lean_Meta_synthInstance_x3f_assignOutParams___closed__3; +x_85 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_85, 0, x_84); +lean_ctor_set(x_85, 1, x_83); +x_86 = l_Lean_Meta_synthInstance_x3f_assignOutParams___closed__5; +x_87 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_87, 0, x_85); +lean_ctor_set(x_87, 1, x_86); +x_88 = l_Lean_indentExpr(x_1); +x_89 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_89, 0, x_87); +lean_ctor_set(x_89, 1, x_88); +x_90 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; +x_91 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_91, 0, x_89); +lean_ctor_set(x_91, 1, x_90); +x_92 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_75, x_91, x_3, x_4, x_5, x_6, x_82); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_2); -x_45 = !lean_is_exclusive(x_29); -if (x_45 == 0) -{ -return x_29; +lean_dec(x_3); +x_93 = lean_ctor_get(x_92, 1); +lean_inc(x_93); +if (lean_is_exclusive(x_92)) { + lean_ctor_release(x_92, 0); + lean_ctor_release(x_92, 1); + x_94 = x_92; +} else { + lean_dec_ref(x_92); + x_94 = lean_box(0); } -else -{ -lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_46 = lean_ctor_get(x_29, 0); -x_47 = lean_ctor_get(x_29, 1); -lean_inc(x_47); -lean_inc(x_46); -lean_dec(x_29); -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; +if (lean_is_scalar(x_94)) { + x_95 = lean_alloc_ctor(0, 2, 0); +} else { + x_95 = x_94; } +lean_ctor_set(x_95, 0, x_72); +lean_ctor_set(x_95, 1, x_93); +return x_95; } } else { -uint8_t x_49; -lean_dec(x_25); -lean_dec(x_17); -lean_dec(x_9); -lean_dec(x_7); +lean_object* x_96; 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_2); -x_49 = !lean_is_exclusive(x_26); -if (x_49 == 0) -{ -return x_26; +lean_dec(x_3); +lean_dec(x_1); +x_96 = lean_ctor_get(x_71, 1); +lean_inc(x_96); +if (lean_is_exclusive(x_71)) { + lean_ctor_release(x_71, 0); + lean_ctor_release(x_71, 1); + x_97 = x_71; +} else { + lean_dec_ref(x_71); + x_97 = lean_box(0); } -else -{ -lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_50 = lean_ctor_get(x_26, 0); -x_51 = lean_ctor_get(x_26, 1); -lean_inc(x_51); -lean_inc(x_50); -lean_dec(x_26); -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; +if (lean_is_scalar(x_97)) { + x_98 = lean_alloc_ctor(0, 2, 0); +} else { + x_98 = x_97; } +lean_ctor_set(x_98, 0, x_72); +lean_ctor_set(x_98, 1, x_96); +return x_98; } } else { -lean_dec(x_17); -lean_dec(x_9); -lean_dec(x_7); +lean_object* x_99; lean_object* x_100; 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); -lean_ctor_set(x_11, 0, x_1); -return x_11; +lean_dec(x_1); +x_99 = lean_ctor_get(x_71, 0); +lean_inc(x_99); +x_100 = lean_ctor_get(x_71, 1); +lean_inc(x_100); +if (lean_is_exclusive(x_71)) { + lean_ctor_release(x_71, 0); + lean_ctor_release(x_71, 1); + x_101 = x_71; +} else { + lean_dec_ref(x_71); + x_101 = lean_box(0); +} +if (lean_is_scalar(x_101)) { + x_102 = lean_alloc_ctor(1, 2, 0); +} else { + x_102 = x_101; +} +lean_ctor_set(x_102, 0, x_99); +lean_ctor_set(x_102, 1, x_100); +return x_102; } } } else { -lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; -x_53 = lean_ctor_get(x_11, 0); -x_54 = lean_ctor_get(x_11, 1); -lean_inc(x_54); -lean_inc(x_53); -lean_dec(x_11); -x_55 = lean_ctor_get(x_53, 0); -lean_inc(x_55); -lean_dec(x_53); -x_56 = l_Lean_getOutParamPositions_x3f(x_55, x_10); -if (lean_obj_tag(x_56) == 0) -{ -lean_object* x_57; -lean_dec(x_9); -lean_dec(x_7); +uint8_t x_103; lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -x_57 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_57, 0, x_1); -lean_ctor_set(x_57, 1, x_54); -return x_57; -} -else -{ -lean_object* x_58; uint8_t x_59; -x_58 = lean_ctor_get(x_56, 0); -lean_inc(x_58); -lean_dec(x_56); -x_59 = l_Array_isEmpty___rarg(x_58); -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_dec(x_1); -x_60 = lean_unsigned_to_nat(0u); -x_61 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_3, x_60); -x_62 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessOutParam___lambda__1___closed__1; -lean_inc(x_61); -x_63 = lean_mk_array(x_61, x_62); -x_64 = lean_unsigned_to_nat(1u); -x_65 = lean_nat_sub(x_61, x_64); -lean_dec(x_61); -x_66 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_3, x_63, x_65); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_9); -x_67 = lean_infer_type(x_9, x_4, x_5, x_6, x_7, x_54); -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_67, 0); -lean_inc(x_68); -x_69 = lean_ctor_get(x_67, 1); -lean_inc(x_69); -lean_dec(x_67); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_70 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessArgs(x_68, x_60, x_66, x_58, x_4, x_5, x_6, x_7, x_69); -lean_dec(x_58); -if (lean_obj_tag(x_70) == 0) -{ -lean_object* x_71; lean_object* x_72; lean_object* x_73; uint8_t x_74; uint8_t x_75; uint8_t x_76; lean_object* x_77; -x_71 = lean_ctor_get(x_70, 0); -lean_inc(x_71); -x_72 = lean_ctor_get(x_70, 1); -lean_inc(x_72); -lean_dec(x_70); -x_73 = l_Lean_mkAppN(x_9, x_71); -x_74 = 0; -x_75 = 1; -x_76 = 1; -x_77 = l_Lean_Meta_mkForallFVars(x_2, x_73, x_74, x_75, x_76, x_4, x_5, x_6, x_7, x_72); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -if (lean_obj_tag(x_77) == 0) +x_103 = !lean_is_exclusive(x_8); +if (x_103 == 0) { -lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; -x_78 = lean_ctor_get(x_77, 0); -lean_inc(x_78); -x_79 = lean_ctor_get(x_77, 1); -lean_inc(x_79); -if (lean_is_exclusive(x_77)) { - lean_ctor_release(x_77, 0); - lean_ctor_release(x_77, 1); - x_80 = x_77; -} else { - lean_dec_ref(x_77); - x_80 = lean_box(0); -} -if (lean_is_scalar(x_80)) { - x_81 = lean_alloc_ctor(0, 2, 0); -} else { - x_81 = x_80; -} -lean_ctor_set(x_81, 0, x_78); -lean_ctor_set(x_81, 1, x_79); -return x_81; +return x_8; } else { -lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; -x_82 = lean_ctor_get(x_77, 0); -lean_inc(x_82); -x_83 = lean_ctor_get(x_77, 1); -lean_inc(x_83); -if (lean_is_exclusive(x_77)) { - lean_ctor_release(x_77, 0); - lean_ctor_release(x_77, 1); - x_84 = x_77; -} else { - lean_dec_ref(x_77); - x_84 = lean_box(0); +lean_object* x_104; lean_object* x_105; lean_object* x_106; +x_104 = lean_ctor_get(x_8, 0); +x_105 = lean_ctor_get(x_8, 1); +lean_inc(x_105); +lean_inc(x_104); +lean_dec(x_8); +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; } -if (lean_is_scalar(x_84)) { - x_85 = lean_alloc_ctor(1, 2, 0); -} else { - x_85 = x_84; } -lean_ctor_set(x_85, 0, x_82); -lean_ctor_set(x_85, 1, x_83); -return x_85; } } -else +LEAN_EXPORT lean_object* l_Lean_Meta_synthInstance_x3f_assignOutParams___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_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; -lean_dec(x_9); -lean_dec(x_7); +uint8_t x_8; lean_object* x_9; +x_8 = lean_unbox(x_1); +lean_dec(x_1); +x_9 = l_Lean_Meta_synthInstance_x3f_assignOutParams___lambda__1(x_8, x_2, x_3, x_4, x_5, x_6, x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); +lean_dec(x_3); lean_dec(x_2); -x_86 = lean_ctor_get(x_70, 0); -lean_inc(x_86); -x_87 = lean_ctor_get(x_70, 1); -lean_inc(x_87); -if (lean_is_exclusive(x_70)) { - lean_ctor_release(x_70, 0); - lean_ctor_release(x_70, 1); - x_88 = x_70; -} else { - lean_dec_ref(x_70); - x_88 = lean_box(0); +return x_9; } -if (lean_is_scalar(x_88)) { - x_89 = lean_alloc_ctor(1, 2, 0); -} else { - x_89 = x_88; } -lean_ctor_set(x_89, 0, x_86); -lean_ctor_set(x_89, 1, x_87); -return x_89; +LEAN_EXPORT uint8_t l_Array_isEqvAux___at_Lean_Meta_synthInstance_x3f___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; 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; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_10 = lean_array_fget(x_4, x_6); +x_11 = lean_array_fget(x_5, x_6); +x_12 = lean_ctor_get(x_10, 1); +lean_inc(x_12); +lean_dec(x_10); +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +lean_dec(x_11); +x_14 = lean_expr_eqv(x_12, x_13); +lean_dec(x_13); +lean_dec(x_12); +if (x_14 == 0) +{ +uint8_t x_15; +lean_dec(x_6); +x_15 = 0; +return x_15; } else { -lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; -lean_dec(x_66); -lean_dec(x_58); -lean_dec(x_9); -lean_dec(x_7); +lean_object* x_16; lean_object* x_17; +x_16 = lean_unsigned_to_nat(1u); +x_17 = lean_nat_add(x_6, x_16); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_2); -x_90 = lean_ctor_get(x_67, 0); -lean_inc(x_90); -x_91 = lean_ctor_get(x_67, 1); -lean_inc(x_91); -if (lean_is_exclusive(x_67)) { - lean_ctor_release(x_67, 0); - lean_ctor_release(x_67, 1); - x_92 = x_67; -} else { - lean_dec_ref(x_67); - x_92 = lean_box(0); +x_3 = lean_box(0); +x_6 = x_17; +goto _start; } -if (lean_is_scalar(x_92)) { - x_93 = lean_alloc_ctor(1, 2, 0); -} else { - x_93 = x_92; } -lean_ctor_set(x_93, 0, x_90); -lean_ctor_set(x_93, 1, x_91); -return x_93; } } -else +LEAN_EXPORT uint8_t l_Array_isEqvAux___at_Lean_Meta_synthInstance_x3f___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_94; -lean_dec(x_58); -lean_dec(x_9); +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); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_94 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_94, 0, x_1); -lean_ctor_set(x_94, 1, x_54); -return x_94; -} -} +x_9 = 1; +return x_9; } +else +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_10 = lean_array_fget(x_4, x_6); +x_11 = lean_array_fget(x_5, x_6); +x_12 = lean_ctor_get(x_10, 1); +lean_inc(x_12); +lean_dec(x_10); +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +lean_dec(x_11); +x_14 = lean_expr_eqv(x_12, x_13); +lean_dec(x_13); +lean_dec(x_12); +if (x_14 == 0) +{ +uint8_t x_15; +lean_dec(x_6); +x_15 = 0; +return x_15; } else { -lean_object* x_95; -lean_dec(x_9); -lean_dec(x_7); +lean_object* x_16; lean_object* x_17; +x_16 = lean_unsigned_to_nat(1u); +x_17 = lean_nat_add(x_6, x_16); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_95 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_95, 0, x_1); -lean_ctor_set(x_95, 1, x_8); -return x_95; +x_3 = lean_box(0); +x_6 = x_17; +goto _start; } } } -LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessOutParam(lean_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_PersistentHashMap_findAtAux___at_Lean_Meta_synthInstance_x3f___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_7; lean_object* x_8; -lean_inc(x_1); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessOutParam___lambda__1), 8, 1); -lean_closure_set(x_7, 0, x_1); -x_8 = l_Lean_Meta_forallTelescope___at___private_Lean_Meta_InferType_0__Lean_Meta_inferForallType___spec__2___rarg(x_1, x_7, x_2, x_3, x_4, x_5, x_6); +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); +x_8 = lean_box(0); 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; uint8_t x_16; +x_9 = lean_array_fget(x_1, x_4); +x_10 = lean_ctor_get(x_5, 0); +x_11 = lean_ctor_get(x_5, 1); +x_12 = lean_ctor_get(x_9, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_9, 1); +lean_inc(x_13); +lean_dec(x_9); +x_14 = lean_array_get_size(x_10); +x_15 = lean_array_get_size(x_12); +x_16 = lean_nat_dec_eq(x_14, x_15); +lean_dec(x_15); +lean_dec(x_14); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; +lean_dec(x_13); +lean_dec(x_12); +x_17 = lean_unsigned_to_nat(1u); +x_18 = lean_nat_add(x_4, x_17); +lean_dec(x_4); +x_3 = lean_box(0); +x_4 = x_18; +goto _start; } -LEAN_EXPORT lean_object* l_Lean_Meta_synthInstance_x3f_assignOutParams___lambda__1(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) { -_start: +else { -lean_object* x_8; lean_object* x_9; -x_8 = lean_box(x_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); -return x_9; +lean_object* x_20; uint8_t x_21; +x_20 = lean_unsigned_to_nat(0u); +x_21 = l_Array_isEqvAux___at_Lean_Meta_synthInstance_x3f___spec__5(x_10, x_12, lean_box(0), x_10, x_12, x_20); +lean_dec(x_12); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; +lean_dec(x_13); +x_22 = lean_unsigned_to_nat(1u); +x_23 = lean_nat_add(x_4, x_22); +lean_dec(x_4); +x_3 = lean_box(0); +x_4 = x_23; +goto _start; } +else +{ +uint8_t x_25; +x_25 = lean_expr_eqv(x_11, x_13); +lean_dec(x_13); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; +x_26 = lean_unsigned_to_nat(1u); +x_27 = lean_nat_add(x_4, x_26); +lean_dec(x_4); +x_3 = lean_box(0); +x_4 = x_27; +goto _start; } -static lean_object* _init_l_Lean_Meta_synthInstance_x3f_assignOutParams___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes(" result type", 12); -return x_1; +lean_object* x_29; lean_object* x_30; +x_29 = lean_array_fget(x_2, x_4); +lean_dec(x_4); +x_30 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_30, 0, x_29); +return x_30; } } -static lean_object* _init_l_Lean_Meta_synthInstance_x3f_assignOutParams___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_synthInstance_x3f_assignOutParams___closed__1; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; } } -static lean_object* _init_l_Lean_Meta_synthInstance_x3f_assignOutParams___closed__3() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__4; -x_2 = l_Lean_Meta_synthInstance_x3f_assignOutParams___closed__2; -x_3 = lean_alloc_ctor(7, 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_Meta_synthInstance_x3f_assignOutParams___closed__4() { +static size_t _init_l_Lean_PersistentHashMap_findAux___at_Lean_Meta_synthInstance_x3f___spec__2___closed__1() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("\nis not definitionally equal to", 31); -return x_1; +size_t x_1; size_t x_2; size_t x_3; +x_1 = 1; +x_2 = 5; +x_3 = lean_usize_shift_left(x_1, x_2); +return x_3; } } -static lean_object* _init_l_Lean_Meta_synthInstance_x3f_assignOutParams___closed__5() { +static size_t _init_l_Lean_PersistentHashMap_findAux___at_Lean_Meta_synthInstance_x3f___spec__2___closed__2() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_synthInstance_x3f_assignOutParams___closed__4; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +size_t x_1; size_t x_2; size_t x_3; +x_1 = 1; +x_2 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_synthInstance_x3f___spec__2___closed__1; +x_3 = lean_usize_sub(x_2, x_1); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Meta_synthInstance_x3f_assignOutParams(lean_object* x_1, lean_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_PersistentHashMap_findAux___at_Lean_Meta_synthInstance_x3f___spec__2(lean_object* x_1, size_t x_2, lean_object* x_3) { _start: { -lean_object* x_8; -lean_inc(x_6); -lean_inc(x_5); +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_4; size_t x_5; size_t x_6; size_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_4 = lean_ctor_get(x_1, 0); lean_inc(x_4); -lean_inc(x_3); -x_8 = lean_infer_type(x_2, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_8) == 0) +lean_dec(x_1); +x_5 = 5; +x_6 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_synthInstance_x3f___spec__2___closed__2; +x_7 = lean_usize_land(x_2, x_6); +x_8 = lean_usize_to_nat(x_7); +x_9 = lean_box(2); +x_10 = lean_array_get(x_9, x_4, x_8); +lean_dec(x_8); +lean_dec(x_4); +switch (lean_obj_tag(x_10)) { +case 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; uint8_t x_17; -x_9 = lean_ctor_get(x_3, 0); -lean_inc(x_9); -x_10 = lean_ctor_get(x_8, 0); -lean_inc(x_10); -x_11 = lean_ctor_get(x_8, 1); +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; +x_11 = lean_ctor_get(x_10, 0); lean_inc(x_11); -lean_dec(x_8); -x_12 = lean_ctor_get(x_3, 1); +x_12 = lean_ctor_get(x_10, 1); lean_inc(x_12); -x_13 = lean_ctor_get(x_3, 2); -lean_inc(x_13); -x_14 = lean_ctor_get(x_3, 3); -lean_inc(x_14); -x_15 = lean_ctor_get(x_3, 4); +lean_dec(x_10); +x_13 = lean_ctor_get(x_3, 0); +x_14 = lean_ctor_get(x_3, 1); +x_15 = lean_ctor_get(x_11, 0); lean_inc(x_15); -x_16 = lean_ctor_get(x_3, 5); +x_16 = lean_ctor_get(x_11, 1); lean_inc(x_16); -x_17 = !lean_is_exclusive(x_9); -if (x_17 == 0) -{ -uint8_t x_18; uint8_t x_19; lean_object* x_20; lean_object* x_21; -x_18 = 1; -x_19 = 1; -lean_ctor_set_uint8(x_9, 5, x_18); -lean_ctor_set_uint8(x_9, 10, x_19); -x_20 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_20, 0, x_9); -lean_ctor_set(x_20, 1, x_12); -lean_ctor_set(x_20, 2, x_13); -lean_ctor_set(x_20, 3, x_14); -lean_ctor_set(x_20, 4, x_15); -lean_ctor_set(x_20, 5, x_16); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_10); -lean_inc(x_1); -x_21 = l_Lean_Meta_isExprDefEq(x_1, x_10, x_20, x_4, x_5, x_6, x_11); -if (lean_obj_tag(x_21) == 0) -{ -lean_object* x_22; uint8_t x_23; -x_22 = lean_ctor_get(x_21, 0); -lean_inc(x_22); -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; uint8_t x_28; -x_24 = lean_ctor_get(x_21, 1); -lean_inc(x_24); -lean_dec(x_21); -x_25 = l_Lean_Meta_SynthInstance_newSubgoal___closed__1; -x_26 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_25, x_3, x_4, x_5, x_6, 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) -{ -uint8_t x_29; -lean_dec(x_10); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -x_29 = !lean_is_exclusive(x_26); -if (x_29 == 0) +lean_dec(x_11); +x_17 = lean_array_get_size(x_13); +x_18 = lean_array_get_size(x_15); +x_19 = lean_nat_dec_eq(x_17, x_18); +lean_dec(x_18); +lean_dec(x_17); +if (x_19 == 0) { -lean_object* x_30; -x_30 = lean_ctor_get(x_26, 0); -lean_dec(x_30); -lean_ctor_set(x_26, 0, x_22); -return x_26; +lean_object* x_20; +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_12); +x_20 = lean_box(0); +return x_20; } else { -lean_object* x_31; lean_object* x_32; -x_31 = lean_ctor_get(x_26, 1); -lean_inc(x_31); -lean_dec(x_26); -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_22); -lean_ctor_set(x_32, 1, x_31); -return x_32; -} +lean_object* x_21; uint8_t x_22; +x_21 = lean_unsigned_to_nat(0u); +x_22 = l_Array_isEqvAux___at_Lean_Meta_synthInstance_x3f___spec__3(x_13, x_15, lean_box(0), x_13, x_15, x_21); +lean_dec(x_15); +if (x_22 == 0) +{ +lean_object* x_23; +lean_dec(x_16); +lean_dec(x_12); +x_23 = lean_box(0); +return x_23; } 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; uint8_t x_44; -x_33 = lean_ctor_get(x_26, 1); -lean_inc(x_33); -lean_dec(x_26); -x_34 = l_Lean_indentExpr(x_10); -x_35 = l_Lean_Meta_synthInstance_x3f_assignOutParams___closed__3; -x_36 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_36, 0, x_35); -lean_ctor_set(x_36, 1, x_34); -x_37 = l_Lean_Meta_synthInstance_x3f_assignOutParams___closed__5; -x_38 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_38, 0, x_36); -lean_ctor_set(x_38, 1, x_37); -x_39 = l_Lean_indentExpr(x_1); -x_40 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set(x_40, 1, x_39); -x_41 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; -x_42 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_42, 0, x_40); -lean_ctor_set(x_42, 1, x_41); -x_43 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_25, x_42, x_3, x_4, x_5, x_6, x_33); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -x_44 = !lean_is_exclusive(x_43); -if (x_44 == 0) +uint8_t x_24; +x_24 = lean_expr_eqv(x_14, x_16); +lean_dec(x_16); +if (x_24 == 0) { -lean_object* x_45; -x_45 = lean_ctor_get(x_43, 0); -lean_dec(x_45); -lean_ctor_set(x_43, 0, x_22); -return x_43; +lean_object* x_25; +lean_dec(x_12); +x_25 = lean_box(0); +return x_25; } else { -lean_object* x_46; lean_object* x_47; -x_46 = lean_ctor_get(x_43, 1); -lean_inc(x_46); -lean_dec(x_43); -x_47 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_47, 0, x_22); -lean_ctor_set(x_47, 1, x_46); -return x_47; +lean_object* x_26; +x_26 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_26, 0, x_12); +return x_26; } } } -else +} +case 1: { -uint8_t x_48; +lean_object* x_27; size_t x_28; +x_27 = lean_ctor_get(x_10, 0); +lean_inc(x_27); lean_dec(x_10); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -x_48 = !lean_is_exclusive(x_21); -if (x_48 == 0) +x_28 = lean_usize_shift_right(x_2, x_5); +x_1 = x_27; +x_2 = x_28; +goto _start; +} +default: { -lean_object* x_49; -x_49 = lean_ctor_get(x_21, 0); -lean_dec(x_49); -return x_21; +lean_object* x_30; +x_30 = lean_box(0); +return x_30; +} +} } else { -lean_object* x_50; lean_object* x_51; -x_50 = lean_ctor_get(x_21, 1); -lean_inc(x_50); -lean_dec(x_21); -x_51 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_51, 0, x_22); -lean_ctor_set(x_51, 1, x_50); -return x_51; +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_31 = lean_ctor_get(x_1, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_1, 1); +lean_inc(x_32); +lean_dec(x_1); +x_33 = lean_unsigned_to_nat(0u); +x_34 = l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_synthInstance_x3f___spec__4(x_31, x_32, lean_box(0), x_33, x_3); +lean_dec(x_32); +lean_dec(x_31); +return x_34; } } } -else +LEAN_EXPORT uint64_t l_Array_foldlMUnsafe_fold___at_Lean_Meta_synthInstance_x3f___spec__6(lean_object* x_1, size_t x_2, size_t x_3, uint64_t x_4) { +_start: { -uint8_t x_52; -lean_dec(x_10); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -x_52 = !lean_is_exclusive(x_21); -if (x_52 == 0) +uint8_t x_5; +x_5 = lean_usize_dec_eq(x_2, x_3); +if (x_5 == 0) { -return x_21; +lean_object* x_6; lean_object* x_7; uint64_t x_8; uint64_t x_9; size_t x_10; size_t x_11; +x_6 = lean_array_uget(x_1, x_2); +x_7 = lean_ctor_get(x_6, 1); +lean_inc(x_7); +lean_dec(x_6); +x_8 = l_Lean_Expr_hash(x_7); +lean_dec(x_7); +x_9 = lean_uint64_mix_hash(x_4, x_8); +x_10 = 1; +x_11 = lean_usize_add(x_2, x_10); +x_2 = x_11; +x_4 = x_9; +goto _start; } else { -lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_53 = lean_ctor_get(x_21, 0); -x_54 = lean_ctor_get(x_21, 1); -lean_inc(x_54); -lean_inc(x_53); -lean_dec(x_21); -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; +return x_4; } } } -else +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_synthInstance_x3f___spec__1(lean_object* x_1, lean_object* x_2) { +_start: { -uint8_t x_56; uint8_t x_57; uint8_t x_58; uint8_t x_59; uint8_t x_60; uint8_t x_61; uint8_t x_62; uint8_t x_63; uint8_t x_64; uint8_t x_65; uint8_t x_66; uint8_t x_67; uint8_t x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; -x_56 = lean_ctor_get_uint8(x_9, 0); -x_57 = lean_ctor_get_uint8(x_9, 1); -x_58 = lean_ctor_get_uint8(x_9, 2); -x_59 = lean_ctor_get_uint8(x_9, 3); -x_60 = lean_ctor_get_uint8(x_9, 4); -x_61 = lean_ctor_get_uint8(x_9, 6); -x_62 = lean_ctor_get_uint8(x_9, 7); -x_63 = lean_ctor_get_uint8(x_9, 8); -x_64 = lean_ctor_get_uint8(x_9, 9); -x_65 = lean_ctor_get_uint8(x_9, 11); -x_66 = lean_ctor_get_uint8(x_9, 12); -lean_dec(x_9); -x_67 = 1; -x_68 = 1; -x_69 = lean_alloc_ctor(0, 0, 13); -lean_ctor_set_uint8(x_69, 0, x_56); -lean_ctor_set_uint8(x_69, 1, x_57); -lean_ctor_set_uint8(x_69, 2, x_58); -lean_ctor_set_uint8(x_69, 3, x_59); -lean_ctor_set_uint8(x_69, 4, x_60); -lean_ctor_set_uint8(x_69, 5, x_67); -lean_ctor_set_uint8(x_69, 6, x_61); -lean_ctor_set_uint8(x_69, 7, x_62); -lean_ctor_set_uint8(x_69, 8, x_63); -lean_ctor_set_uint8(x_69, 9, x_64); -lean_ctor_set_uint8(x_69, 10, x_68); -lean_ctor_set_uint8(x_69, 11, x_65); -lean_ctor_set_uint8(x_69, 12, x_66); -x_70 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_70, 0, x_69); -lean_ctor_set(x_70, 1, x_12); -lean_ctor_set(x_70, 2, x_13); -lean_ctor_set(x_70, 3, x_14); -lean_ctor_set(x_70, 4, x_15); -lean_ctor_set(x_70, 5, x_16); -lean_inc(x_6); -lean_inc(x_5); +lean_object* x_3; lean_object* x_4; lean_object* x_5; uint64_t x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; uint64_t x_10; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +lean_dec(x_1); +x_4 = lean_ctor_get(x_2, 0); lean_inc(x_4); -lean_inc(x_10); -lean_inc(x_1); -x_71 = l_Lean_Meta_isExprDefEq(x_1, x_10, x_70, x_4, x_5, x_6, x_11); -if (lean_obj_tag(x_71) == 0) -{ -lean_object* x_72; uint8_t x_73; -x_72 = lean_ctor_get(x_71, 0); -lean_inc(x_72); -x_73 = lean_unbox(x_72); -if (x_73 == 0) -{ -lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; uint8_t x_78; -x_74 = lean_ctor_get(x_71, 1); -lean_inc(x_74); -lean_dec(x_71); -x_75 = l_Lean_Meta_SynthInstance_newSubgoal___closed__1; -x_76 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_75, x_3, x_4, x_5, x_6, x_74); -x_77 = lean_ctor_get(x_76, 0); -lean_inc(x_77); -x_78 = lean_unbox(x_77); -lean_dec(x_77); -if (x_78 == 0) -{ -lean_object* x_79; lean_object* x_80; lean_object* x_81; -lean_dec(x_10); -lean_dec(x_6); +x_5 = lean_ctor_get(x_2, 1); +lean_inc(x_5); +x_6 = 7; +x_7 = lean_array_get_size(x_4); +x_8 = lean_unsigned_to_nat(0u); +x_9 = lean_nat_dec_lt(x_8, x_7); +x_10 = l_Lean_Expr_hash(x_5); lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -x_79 = lean_ctor_get(x_76, 1); -lean_inc(x_79); -if (lean_is_exclusive(x_76)) { - lean_ctor_release(x_76, 0); - lean_ctor_release(x_76, 1); - x_80 = x_76; -} else { - lean_dec_ref(x_76); - x_80 = lean_box(0); -} -if (lean_is_scalar(x_80)) { - x_81 = lean_alloc_ctor(0, 2, 0); -} else { - x_81 = x_80; -} -lean_ctor_set(x_81, 0, x_72); -lean_ctor_set(x_81, 1, x_79); -return x_81; -} -else +if (x_9 == 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; -x_82 = lean_ctor_get(x_76, 1); -lean_inc(x_82); -lean_dec(x_76); -x_83 = l_Lean_indentExpr(x_10); -x_84 = l_Lean_Meta_synthInstance_x3f_assignOutParams___closed__3; -x_85 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_85, 0, x_84); -lean_ctor_set(x_85, 1, x_83); -x_86 = l_Lean_Meta_synthInstance_x3f_assignOutParams___closed__5; -x_87 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_87, 0, x_85); -lean_ctor_set(x_87, 1, x_86); -x_88 = l_Lean_indentExpr(x_1); -x_89 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_89, 0, x_87); -lean_ctor_set(x_89, 1, x_88); -x_90 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; -x_91 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_91, 0, x_89); -lean_ctor_set(x_91, 1, x_90); -x_92 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_75, x_91, x_3, x_4, x_5, x_6, x_82); -lean_dec(x_6); -lean_dec(x_5); +uint64_t x_11; size_t x_12; lean_object* x_13; +lean_dec(x_7); lean_dec(x_4); -lean_dec(x_3); -x_93 = lean_ctor_get(x_92, 1); -lean_inc(x_93); -if (lean_is_exclusive(x_92)) { - lean_ctor_release(x_92, 0); - lean_ctor_release(x_92, 1); - x_94 = x_92; -} else { - lean_dec_ref(x_92); - x_94 = lean_box(0); -} -if (lean_is_scalar(x_94)) { - x_95 = lean_alloc_ctor(0, 2, 0); -} else { - x_95 = x_94; -} -lean_ctor_set(x_95, 0, x_72); -lean_ctor_set(x_95, 1, x_93); -return x_95; -} +x_11 = lean_uint64_mix_hash(x_6, x_10); +x_12 = lean_uint64_to_usize(x_11); +x_13 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_synthInstance_x3f___spec__2(x_3, x_12, x_2); +lean_dec(x_2); +return x_13; } else { -lean_object* x_96; lean_object* x_97; lean_object* x_98; -lean_dec(x_10); -lean_dec(x_6); -lean_dec(x_5); +uint8_t x_14; +x_14 = lean_nat_dec_le(x_7, x_7); +if (x_14 == 0) +{ +uint64_t x_15; size_t x_16; lean_object* x_17; +lean_dec(x_7); lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -x_96 = lean_ctor_get(x_71, 1); -lean_inc(x_96); -if (lean_is_exclusive(x_71)) { - lean_ctor_release(x_71, 0); - lean_ctor_release(x_71, 1); - x_97 = x_71; -} else { - lean_dec_ref(x_71); - x_97 = lean_box(0); -} -if (lean_is_scalar(x_97)) { - x_98 = lean_alloc_ctor(0, 2, 0); -} else { - x_98 = x_97; -} -lean_ctor_set(x_98, 0, x_72); -lean_ctor_set(x_98, 1, x_96); -return x_98; -} +x_15 = lean_uint64_mix_hash(x_6, x_10); +x_16 = lean_uint64_to_usize(x_15); +x_17 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_synthInstance_x3f___spec__2(x_3, x_16, x_2); +lean_dec(x_2); +return x_17; } else { -lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; -lean_dec(x_10); -lean_dec(x_6); -lean_dec(x_5); +size_t x_18; size_t x_19; uint64_t x_20; uint64_t x_21; size_t x_22; lean_object* x_23; +x_18 = 0; +x_19 = lean_usize_of_nat(x_7); +lean_dec(x_7); +x_20 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_synthInstance_x3f___spec__6(x_4, x_18, x_19, x_6); lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -x_99 = lean_ctor_get(x_71, 0); -lean_inc(x_99); -x_100 = lean_ctor_get(x_71, 1); -lean_inc(x_100); -if (lean_is_exclusive(x_71)) { - lean_ctor_release(x_71, 0); - lean_ctor_release(x_71, 1); - x_101 = x_71; -} else { - lean_dec_ref(x_71); - x_101 = lean_box(0); +x_21 = lean_uint64_mix_hash(x_20, x_10); +x_22 = lean_uint64_to_usize(x_21); +x_23 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_synthInstance_x3f___spec__2(x_3, x_22, x_2); +lean_dec(x_2); +return x_23; } -if (lean_is_scalar(x_101)) { - x_102 = lean_alloc_ctor(1, 2, 0); -} else { - x_102 = x_101; } -lean_ctor_set(x_102, 0, x_99); -lean_ctor_set(x_102, 1, x_100); -return x_102; } } +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_synthInstance_x3f___spec__9(size_t 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_2); +x_8 = lean_nat_dec_lt(x_5, x_7); +lean_dec(x_7); +if (x_8 == 0) +{ +lean_dec(x_5); +return x_6; } else { -uint8_t x_103; -lean_dec(x_6); +lean_object* x_9; lean_object* x_10; size_t x_11; size_t 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; uint64_t x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; uint64_t x_23; +x_9 = lean_array_fget(x_2, x_5); +x_10 = lean_array_fget(x_3, x_5); +x_11 = 1; +x_12 = lean_usize_sub(x_1, x_11); +x_13 = 5; +x_14 = lean_usize_mul(x_13, x_12); +x_15 = lean_unsigned_to_nat(1u); +x_16 = lean_nat_add(x_5, x_15); lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -x_103 = !lean_is_exclusive(x_8); -if (x_103 == 0) +x_17 = lean_ctor_get(x_9, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_9, 1); +lean_inc(x_18); +x_19 = 7; +x_20 = lean_array_get_size(x_17); +x_21 = lean_unsigned_to_nat(0u); +x_22 = lean_nat_dec_lt(x_21, x_20); +x_23 = l_Lean_Expr_hash(x_18); +lean_dec(x_18); +if (x_22 == 0) { -return x_8; +uint64_t x_24; size_t x_25; size_t x_26; lean_object* x_27; +lean_dec(x_20); +lean_dec(x_17); +x_24 = lean_uint64_mix_hash(x_19, x_23); +x_25 = lean_uint64_to_usize(x_24); +x_26 = lean_usize_shift_right(x_25, x_14); +x_27 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_synthInstance_x3f___spec__8(x_6, x_26, x_1, x_9, x_10); +x_4 = lean_box(0); +x_5 = x_16; +x_6 = x_27; +goto _start; } else { -lean_object* x_104; lean_object* x_105; lean_object* x_106; -x_104 = lean_ctor_get(x_8, 0); -x_105 = lean_ctor_get(x_8, 1); -lean_inc(x_105); -lean_inc(x_104); -lean_dec(x_8); -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; +uint8_t x_29; +x_29 = lean_nat_dec_le(x_20, x_20); +if (x_29 == 0) +{ +uint64_t x_30; size_t x_31; size_t x_32; lean_object* x_33; +lean_dec(x_20); +lean_dec(x_17); +x_30 = lean_uint64_mix_hash(x_19, x_23); +x_31 = lean_uint64_to_usize(x_30); +x_32 = lean_usize_shift_right(x_31, x_14); +x_33 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_synthInstance_x3f___spec__8(x_6, x_32, x_1, x_9, x_10); +x_4 = lean_box(0); +x_5 = x_16; +x_6 = x_33; +goto _start; } +else +{ +size_t x_35; size_t x_36; uint64_t x_37; uint64_t x_38; size_t x_39; size_t x_40; lean_object* x_41; +x_35 = 0; +x_36 = lean_usize_of_nat(x_20); +lean_dec(x_20); +x_37 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_synthInstance_x3f___spec__6(x_17, x_35, x_36, x_19); +lean_dec(x_17); +x_38 = lean_uint64_mix_hash(x_37, x_23); +x_39 = lean_uint64_to_usize(x_38); +x_40 = lean_usize_shift_right(x_39, x_14); +x_41 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_synthInstance_x3f___spec__8(x_6, x_40, x_1, x_9, x_10); +x_4 = lean_box(0); +x_5 = x_16; +x_6 = x_41; +goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_synthInstance_x3f_assignOutParams___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: -{ -uint8_t x_8; lean_object* x_9; -x_8 = lean_unbox(x_1); -lean_dec(x_1); -x_9 = l_Lean_Meta_synthInstance_x3f_assignOutParams___lambda__1(x_8, x_2, x_3, x_4, x_5, x_6, 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 uint8_t l_Array_isEqvAux___at_Lean_Meta_synthInstance_x3f___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_EXPORT uint8_t l_Array_isEqvAux___at_Lean_Meta_synthInstance_x3f___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) { _start: { lean_object* x_7; uint8_t x_8; @@ -20778,7 +21848,7 @@ goto _start; } } } -LEAN_EXPORT uint8_t l_Array_isEqvAux___at_Lean_Meta_synthInstance_x3f___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_EXPORT uint8_t l_Array_isEqvAux___at_Lean_Meta_synthInstance_x3f___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) { _start: { lean_object* x_7; uint8_t x_8; @@ -20826,2931 +21896,3160 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_synthInstance_x3f___spec__4(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_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_synthInstance_x3f___spec__11(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _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_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; +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_array_get_size(x_5); +x_8 = lean_nat_dec_lt(x_2, x_7); +lean_dec(x_7); +if (x_8 == 0) { -lean_object* x_8; -lean_dec(x_4); -x_8 = lean_box(0); -return x_8; +uint8_t x_9; +lean_dec(x_2); +x_9 = !lean_is_exclusive(x_1); +if (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_1, 1); +lean_dec(x_10); +x_11 = lean_ctor_get(x_1, 0); +lean_dec(x_11); +x_12 = lean_array_push(x_5, x_3); +x_13 = lean_array_push(x_6, x_4); +lean_ctor_set(x_1, 1, x_13); +lean_ctor_set(x_1, 0, x_12); +return x_1; } 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; uint8_t x_16; -x_9 = lean_array_fget(x_1, x_4); -x_10 = lean_ctor_get(x_5, 0); -x_11 = lean_ctor_get(x_5, 1); -x_12 = lean_ctor_get(x_9, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_9, 1); -lean_inc(x_13); -lean_dec(x_9); -x_14 = lean_array_get_size(x_10); -x_15 = lean_array_get_size(x_12); -x_16 = lean_nat_dec_eq(x_14, x_15); -lean_dec(x_15); -lean_dec(x_14); -if (x_16 == 0) +lean_object* x_14; lean_object* x_15; lean_object* x_16; +lean_dec(x_1); +x_14 = lean_array_push(x_5, x_3); +x_15 = lean_array_push(x_6, x_4); +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; +} +} +else { -lean_object* x_17; lean_object* x_18; -lean_dec(x_13); -lean_dec(x_12); -x_17 = lean_unsigned_to_nat(1u); -x_18 = lean_nat_add(x_4, x_17); -lean_dec(x_4); -x_3 = lean_box(0); -x_4 = 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; lean_object* x_23; uint8_t x_24; +x_17 = lean_array_fget(x_5, x_2); +x_18 = lean_ctor_get(x_3, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_3, 1); +lean_inc(x_19); +x_20 = lean_ctor_get(x_17, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_17, 1); +lean_inc(x_21); +lean_dec(x_17); +x_22 = lean_array_get_size(x_18); +x_23 = lean_array_get_size(x_20); +x_24 = lean_nat_dec_eq(x_22, x_23); +lean_dec(x_23); +lean_dec(x_22); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_6); +lean_dec(x_5); +x_25 = lean_unsigned_to_nat(1u); +x_26 = lean_nat_add(x_2, x_25); +lean_dec(x_2); +x_2 = x_26; goto _start; } else { -lean_object* x_20; uint8_t x_21; -x_20 = lean_unsigned_to_nat(0u); -x_21 = l_Array_isEqvAux___at_Lean_Meta_synthInstance_x3f___spec__5(x_10, x_12, lean_box(0), x_10, x_12, x_20); -lean_dec(x_12); -if (x_21 == 0) +lean_object* x_28; uint8_t x_29; +x_28 = lean_unsigned_to_nat(0u); +x_29 = l_Array_isEqvAux___at_Lean_Meta_synthInstance_x3f___spec__12(x_18, x_20, lean_box(0), x_18, x_20, x_28); +lean_dec(x_20); +lean_dec(x_18); +if (x_29 == 0) { -lean_object* x_22; lean_object* x_23; -lean_dec(x_13); -x_22 = lean_unsigned_to_nat(1u); -x_23 = lean_nat_add(x_4, x_22); -lean_dec(x_4); -x_3 = lean_box(0); -x_4 = x_23; +lean_object* x_30; lean_object* x_31; +lean_dec(x_21); +lean_dec(x_19); +lean_dec(x_6); +lean_dec(x_5); +x_30 = lean_unsigned_to_nat(1u); +x_31 = lean_nat_add(x_2, x_30); +lean_dec(x_2); +x_2 = x_31; goto _start; } else { -uint8_t x_25; -x_25 = lean_expr_eqv(x_11, x_13); -lean_dec(x_13); -if (x_25 == 0) +uint8_t x_33; +x_33 = lean_expr_eqv(x_19, x_21); +lean_dec(x_21); +lean_dec(x_19); +if (x_33 == 0) { -lean_object* x_26; lean_object* x_27; -x_26 = lean_unsigned_to_nat(1u); -x_27 = lean_nat_add(x_4, x_26); -lean_dec(x_4); -x_3 = lean_box(0); -x_4 = x_27; +lean_object* x_34; lean_object* x_35; +lean_dec(x_6); +lean_dec(x_5); +x_34 = lean_unsigned_to_nat(1u); +x_35 = lean_nat_add(x_2, x_34); +lean_dec(x_2); +x_2 = x_35; goto _start; } else { -lean_object* x_29; lean_object* x_30; -x_29 = lean_array_fget(x_2, x_4); -lean_dec(x_4); -x_30 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_30, 0, x_29); -return x_30; +uint8_t x_37; +x_37 = !lean_is_exclusive(x_1); +if (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_1, 1); +lean_dec(x_38); +x_39 = lean_ctor_get(x_1, 0); +lean_dec(x_39); +x_40 = lean_array_fset(x_5, x_2, x_3); +x_41 = lean_array_fset(x_6, x_2, x_4); +lean_dec(x_2); +lean_ctor_set(x_1, 1, x_41); +lean_ctor_set(x_1, 0, x_40); +return x_1; } +else +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; +lean_dec(x_1); +x_42 = lean_array_fset(x_5, x_2, x_3); +x_43 = lean_array_fset(x_6, x_2, x_4); +lean_dec(x_2); +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; } } } } } -static size_t _init_l_Lean_PersistentHashMap_findAux___at_Lean_Meta_synthInstance_x3f___spec__2___closed__1() { -_start: -{ -size_t x_1; size_t x_2; size_t x_3; -x_1 = 1; -x_2 = 5; -x_3 = lean_usize_shift_left(x_1, x_2); -return x_3; } } -static size_t _init_l_Lean_PersistentHashMap_findAux___at_Lean_Meta_synthInstance_x3f___spec__2___closed__2() { +static lean_object* _init_l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_synthInstance_x3f___spec__8___closed__1() { _start: { -size_t x_1; size_t x_2; size_t x_3; -x_1 = 1; -x_2 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_synthInstance_x3f___spec__2___closed__1; -x_3 = lean_usize_sub(x_2, x_1); -return x_3; +lean_object* x_1; +x_1 = l_Lean_PersistentHashMap_mkEmptyEntries(lean_box(0), lean_box(0)); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_synthInstance_x3f___spec__2(lean_object* x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_synthInstance_x3f___spec__8(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5) { _start: { if (lean_obj_tag(x_1) == 0) { -lean_object* x_4; size_t x_5; size_t x_6; size_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_4 = lean_ctor_get(x_1, 0); -lean_inc(x_4); -lean_dec(x_1); -x_5 = 5; -x_6 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_synthInstance_x3f___spec__2___closed__2; -x_7 = lean_usize_land(x_2, x_6); -x_8 = lean_usize_to_nat(x_7); -x_9 = lean_box(2); -x_10 = lean_array_get(x_9, x_4, x_8); -lean_dec(x_8); -lean_dec(x_4); -switch (lean_obj_tag(x_10)) { -case 0: +uint8_t x_6; +x_6 = !lean_is_exclusive(x_1); +if (x_6 == 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; uint8_t x_19; -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_3, 0); -x_14 = lean_ctor_get(x_3, 1); -x_15 = lean_ctor_get(x_11, 0); -lean_inc(x_15); -x_16 = lean_ctor_get(x_11, 1); -lean_inc(x_16); -lean_dec(x_11); -x_17 = lean_array_get_size(x_13); -x_18 = lean_array_get_size(x_15); -x_19 = lean_nat_dec_eq(x_17, x_18); -lean_dec(x_18); -lean_dec(x_17); -if (x_19 == 0) +lean_object* x_7; size_t x_8; size_t x_9; size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_7 = lean_ctor_get(x_1, 0); +x_8 = 1; +x_9 = 5; +x_10 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_synthInstance_x3f___spec__2___closed__2; +x_11 = lean_usize_land(x_2, x_10); +x_12 = lean_usize_to_nat(x_11); +x_13 = lean_array_get_size(x_7); +x_14 = lean_nat_dec_lt(x_12, x_13); +lean_dec(x_13); +if (x_14 == 0) { -lean_object* x_20; -lean_dec(x_16); -lean_dec(x_15); lean_dec(x_12); -x_20 = lean_box(0); -return x_20; +lean_dec(x_5); +lean_dec(x_4); +return x_1; } else { -lean_object* x_21; uint8_t x_22; -x_21 = lean_unsigned_to_nat(0u); -x_22 = l_Array_isEqvAux___at_Lean_Meta_synthInstance_x3f___spec__3(x_13, x_15, lean_box(0), x_13, x_15, x_21); -lean_dec(x_15); -if (x_22 == 0) +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_array_fget(x_7, x_12); +x_16 = lean_box(0); +x_17 = lean_array_fset(x_7, x_12, x_16); +switch (lean_obj_tag(x_15)) { +case 0: { -lean_object* x_23; -lean_dec(x_16); +uint8_t x_18; +x_18 = !lean_is_exclusive(x_15); +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; uint8_t x_27; +x_19 = lean_ctor_get(x_15, 0); +x_20 = lean_ctor_get(x_15, 1); +x_21 = lean_ctor_get(x_4, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_4, 1); +lean_inc(x_22); +x_23 = lean_ctor_get(x_19, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_19, 1); +lean_inc(x_24); +x_25 = lean_array_get_size(x_21); +x_26 = lean_array_get_size(x_23); +x_27 = lean_nat_dec_eq(x_25, x_26); +lean_dec(x_26); +lean_dec(x_25); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; +lean_dec(x_24); +lean_dec(x_23); +lean_dec(x_22); +lean_dec(x_21); +lean_free_object(x_15); +x_28 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_19, x_20, x_4, x_5); +x_29 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_29, 0, x_28); +x_30 = lean_array_fset(x_17, x_12, x_29); lean_dec(x_12); -x_23 = lean_box(0); -return x_23; +lean_ctor_set(x_1, 0, x_30); +return x_1; } else { -uint8_t x_24; -x_24 = lean_expr_eqv(x_14, x_16); -lean_dec(x_16); -if (x_24 == 0) +lean_object* x_31; uint8_t x_32; +x_31 = lean_unsigned_to_nat(0u); +x_32 = l_Array_isEqvAux___at_Lean_Meta_synthInstance_x3f___spec__10(x_21, x_23, lean_box(0), x_21, x_23, x_31); +lean_dec(x_23); +lean_dec(x_21); +if (x_32 == 0) { -lean_object* x_25; +lean_object* x_33; lean_object* x_34; lean_object* x_35; +lean_dec(x_24); +lean_dec(x_22); +lean_free_object(x_15); +x_33 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_19, x_20, x_4, x_5); +x_34 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_34, 0, x_33); +x_35 = lean_array_fset(x_17, x_12, x_34); lean_dec(x_12); -x_25 = lean_box(0); -return x_25; +lean_ctor_set(x_1, 0, x_35); +return x_1; } else { -lean_object* x_26; -x_26 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_26, 0, x_12); -return x_26; -} -} -} -} -case 1: -{ -lean_object* x_27; size_t x_28; -x_27 = lean_ctor_get(x_10, 0); -lean_inc(x_27); -lean_dec(x_10); -x_28 = lean_usize_shift_right(x_2, x_5); -x_1 = x_27; -x_2 = x_28; -goto _start; -} -default: +uint8_t x_36; +x_36 = lean_expr_eqv(x_22, x_24); +lean_dec(x_24); +lean_dec(x_22); +if (x_36 == 0) { -lean_object* x_30; -x_30 = lean_box(0); -return x_30; -} -} +lean_object* x_37; lean_object* x_38; lean_object* x_39; +lean_free_object(x_15); +x_37 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_19, x_20, x_4, x_5); +x_38 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_38, 0, x_37); +x_39 = lean_array_fset(x_17, x_12, x_38); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_39); +return x_1; } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_31 = lean_ctor_get(x_1, 0); -lean_inc(x_31); -x_32 = lean_ctor_get(x_1, 1); -lean_inc(x_32); -lean_dec(x_1); -x_33 = lean_unsigned_to_nat(0u); -x_34 = l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_synthInstance_x3f___spec__4(x_31, x_32, lean_box(0), x_33, x_3); -lean_dec(x_32); -lean_dec(x_31); -return x_34; +lean_object* x_40; +lean_dec(x_20); +lean_dec(x_19); +lean_ctor_set(x_15, 1, x_5); +lean_ctor_set(x_15, 0, x_4); +x_40 = lean_array_fset(x_17, x_12, x_15); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_40); +return x_1; } } } -LEAN_EXPORT uint64_t l_Array_foldlMUnsafe_fold___at_Lean_Meta_synthInstance_x3f___spec__6(lean_object* x_1, size_t x_2, size_t x_3, uint64_t x_4) { -_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; uint64_t x_8; uint64_t x_9; size_t x_10; size_t x_11; -x_6 = lean_array_uget(x_1, x_2); -x_7 = lean_ctor_get(x_6, 1); -lean_inc(x_7); -lean_dec(x_6); -x_8 = l_Lean_Expr_hash(x_7); -lean_dec(x_7); -x_9 = lean_uint64_mix_hash(x_4, x_8); -x_10 = 1; -x_11 = lean_usize_add(x_2, x_10); -x_2 = x_11; -x_4 = x_9; -goto _start; } else { -return x_4; -} -} +lean_object* x_41; 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; uint8_t x_49; +x_41 = lean_ctor_get(x_15, 0); +x_42 = lean_ctor_get(x_15, 1); +lean_inc(x_42); +lean_inc(x_41); +lean_dec(x_15); +x_43 = lean_ctor_get(x_4, 0); +lean_inc(x_43); +x_44 = lean_ctor_get(x_4, 1); +lean_inc(x_44); +x_45 = lean_ctor_get(x_41, 0); +lean_inc(x_45); +x_46 = lean_ctor_get(x_41, 1); +lean_inc(x_46); +x_47 = lean_array_get_size(x_43); +x_48 = lean_array_get_size(x_45); +x_49 = lean_nat_dec_eq(x_47, x_48); +lean_dec(x_48); +lean_dec(x_47); +if (x_49 == 0) +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; +lean_dec(x_46); +lean_dec(x_45); +lean_dec(x_44); +lean_dec(x_43); +x_50 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_41, x_42, x_4, x_5); +x_51 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_51, 0, x_50); +x_52 = lean_array_fset(x_17, x_12, x_51); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_52); +return x_1; } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_synthInstance_x3f___spec__1(lean_object* x_1, lean_object* x_2) { -_start: +else { -lean_object* x_3; lean_object* x_4; lean_object* x_5; uint64_t x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; uint64_t x_10; -x_3 = lean_ctor_get(x_1, 0); -lean_inc(x_3); -lean_dec(x_1); -x_4 = lean_ctor_get(x_2, 0); -lean_inc(x_4); -x_5 = lean_ctor_get(x_2, 1); -lean_inc(x_5); -x_6 = 7; -x_7 = lean_array_get_size(x_4); -x_8 = lean_unsigned_to_nat(0u); -x_9 = lean_nat_dec_lt(x_8, x_7); -x_10 = l_Lean_Expr_hash(x_5); -lean_dec(x_5); -if (x_9 == 0) +lean_object* x_53; uint8_t x_54; +x_53 = lean_unsigned_to_nat(0u); +x_54 = l_Array_isEqvAux___at_Lean_Meta_synthInstance_x3f___spec__10(x_43, x_45, lean_box(0), x_43, x_45, x_53); +lean_dec(x_45); +lean_dec(x_43); +if (x_54 == 0) { -uint64_t x_11; size_t x_12; lean_object* x_13; -lean_dec(x_7); -lean_dec(x_4); -x_11 = lean_uint64_mix_hash(x_6, x_10); -x_12 = lean_uint64_to_usize(x_11); -x_13 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_synthInstance_x3f___spec__2(x_3, x_12, x_2); -lean_dec(x_2); -return x_13; +lean_object* x_55; lean_object* x_56; lean_object* x_57; +lean_dec(x_46); +lean_dec(x_44); +x_55 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_41, x_42, x_4, x_5); +x_56 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_56, 0, x_55); +x_57 = lean_array_fset(x_17, x_12, x_56); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_57); +return x_1; } else { -uint8_t x_14; -x_14 = lean_nat_dec_le(x_7, x_7); -if (x_14 == 0) +uint8_t x_58; +x_58 = lean_expr_eqv(x_44, x_46); +lean_dec(x_46); +lean_dec(x_44); +if (x_58 == 0) { -uint64_t x_15; size_t x_16; lean_object* x_17; -lean_dec(x_7); -lean_dec(x_4); -x_15 = lean_uint64_mix_hash(x_6, x_10); -x_16 = lean_uint64_to_usize(x_15); -x_17 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_synthInstance_x3f___spec__2(x_3, x_16, x_2); -lean_dec(x_2); -return x_17; +lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_59 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_41, x_42, x_4, x_5); +x_60 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_60, 0, x_59); +x_61 = lean_array_fset(x_17, x_12, x_60); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_61); +return x_1; } else { -size_t x_18; size_t x_19; uint64_t x_20; uint64_t x_21; size_t x_22; lean_object* x_23; -x_18 = 0; -x_19 = lean_usize_of_nat(x_7); -lean_dec(x_7); -x_20 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_synthInstance_x3f___spec__6(x_4, x_18, x_19, x_6); -lean_dec(x_4); -x_21 = lean_uint64_mix_hash(x_20, x_10); -x_22 = lean_uint64_to_usize(x_21); -x_23 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_synthInstance_x3f___spec__2(x_3, x_22, x_2); -lean_dec(x_2); -return x_23; +lean_object* x_62; lean_object* x_63; +lean_dec(x_42); +lean_dec(x_41); +x_62 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_62, 0, x_4); +lean_ctor_set(x_62, 1, x_5); +x_63 = lean_array_fset(x_17, x_12, x_62); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_63); +return x_1; } } } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_synthInstance_x3f___spec__9(size_t 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_2); -x_8 = lean_nat_dec_lt(x_5, x_7); -lean_dec(x_7); -if (x_8 == 0) -{ -lean_dec(x_5); -return x_6; } -else +case 1: { -lean_object* x_9; lean_object* x_10; size_t x_11; size_t 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; uint64_t x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; uint64_t x_23; -x_9 = lean_array_fget(x_2, x_5); -x_10 = lean_array_fget(x_3, x_5); -x_11 = 1; -x_12 = lean_usize_sub(x_1, x_11); -x_13 = 5; -x_14 = lean_usize_mul(x_13, x_12); -x_15 = lean_unsigned_to_nat(1u); -x_16 = lean_nat_add(x_5, x_15); -lean_dec(x_5); -x_17 = lean_ctor_get(x_9, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_9, 1); -lean_inc(x_18); -x_19 = 7; -x_20 = lean_array_get_size(x_17); -x_21 = lean_unsigned_to_nat(0u); -x_22 = lean_nat_dec_lt(x_21, x_20); -x_23 = l_Lean_Expr_hash(x_18); -lean_dec(x_18); -if (x_22 == 0) +uint8_t x_64; +x_64 = !lean_is_exclusive(x_15); +if (x_64 == 0) { -uint64_t x_24; size_t x_25; size_t x_26; lean_object* x_27; -lean_dec(x_20); -lean_dec(x_17); -x_24 = lean_uint64_mix_hash(x_19, x_23); -x_25 = lean_uint64_to_usize(x_24); -x_26 = lean_usize_shift_right(x_25, x_14); -x_27 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_synthInstance_x3f___spec__8(x_6, x_26, x_1, x_9, x_10); -x_4 = lean_box(0); -x_5 = x_16; -x_6 = x_27; -goto _start; +lean_object* x_65; size_t x_66; size_t x_67; lean_object* x_68; lean_object* x_69; +x_65 = lean_ctor_get(x_15, 0); +x_66 = lean_usize_shift_right(x_2, x_9); +x_67 = lean_usize_add(x_3, x_8); +x_68 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_synthInstance_x3f___spec__8(x_65, x_66, x_67, x_4, x_5); +lean_ctor_set(x_15, 0, x_68); +x_69 = lean_array_fset(x_17, x_12, x_15); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_69); +return x_1; } else { -uint8_t x_29; -x_29 = lean_nat_dec_le(x_20, x_20); -if (x_29 == 0) -{ -uint64_t x_30; size_t x_31; size_t x_32; lean_object* x_33; -lean_dec(x_20); -lean_dec(x_17); -x_30 = lean_uint64_mix_hash(x_19, x_23); -x_31 = lean_uint64_to_usize(x_30); -x_32 = lean_usize_shift_right(x_31, x_14); -x_33 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_synthInstance_x3f___spec__8(x_6, x_32, x_1, x_9, x_10); -x_4 = lean_box(0); -x_5 = x_16; -x_6 = x_33; -goto _start; +lean_object* x_70; size_t x_71; size_t x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; +x_70 = lean_ctor_get(x_15, 0); +lean_inc(x_70); +lean_dec(x_15); +x_71 = lean_usize_shift_right(x_2, x_9); +x_72 = lean_usize_add(x_3, x_8); +x_73 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_synthInstance_x3f___spec__8(x_70, x_71, x_72, x_4, x_5); +x_74 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_74, 0, x_73); +x_75 = lean_array_fset(x_17, x_12, x_74); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_75); +return x_1; } -else -{ -size_t x_35; size_t x_36; uint64_t x_37; uint64_t x_38; size_t x_39; size_t x_40; lean_object* x_41; -x_35 = 0; -x_36 = lean_usize_of_nat(x_20); -lean_dec(x_20); -x_37 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_synthInstance_x3f___spec__6(x_17, x_35, x_36, x_19); -lean_dec(x_17); -x_38 = lean_uint64_mix_hash(x_37, x_23); -x_39 = lean_uint64_to_usize(x_38); -x_40 = lean_usize_shift_right(x_39, x_14); -x_41 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_synthInstance_x3f___spec__8(x_6, x_40, x_1, x_9, x_10); -x_4 = lean_box(0); -x_5 = x_16; -x_6 = x_41; -goto _start; } +default: +{ +lean_object* x_76; lean_object* x_77; +x_76 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_76, 0, x_4); +lean_ctor_set(x_76, 1, x_5); +x_77 = lean_array_fset(x_17, x_12, x_76); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_77); +return x_1; } } } } -LEAN_EXPORT uint8_t l_Array_isEqvAux___at_Lean_Meta_synthInstance_x3f___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) { -_start: +else { -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) +lean_object* x_78; size_t x_79; size_t x_80; size_t x_81; size_t x_82; lean_object* x_83; lean_object* x_84; uint8_t x_85; +x_78 = lean_ctor_get(x_1, 0); +lean_inc(x_78); +lean_dec(x_1); +x_79 = 1; +x_80 = 5; +x_81 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_synthInstance_x3f___spec__2___closed__2; +x_82 = lean_usize_land(x_2, x_81); +x_83 = lean_usize_to_nat(x_82); +x_84 = lean_array_get_size(x_78); +x_85 = lean_nat_dec_lt(x_83, x_84); +lean_dec(x_84); +if (x_85 == 0) { -uint8_t x_9; -lean_dec(x_6); -x_9 = 1; -return x_9; +lean_object* x_86; +lean_dec(x_83); +lean_dec(x_5); +lean_dec(x_4); +x_86 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_86, 0, x_78); +return x_86; } else { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_10 = lean_array_fget(x_4, x_6); -x_11 = lean_array_fget(x_5, x_6); -x_12 = lean_ctor_get(x_10, 1); -lean_inc(x_12); -lean_dec(x_10); -x_13 = lean_ctor_get(x_11, 1); -lean_inc(x_13); -lean_dec(x_11); -x_14 = lean_expr_eqv(x_12, x_13); -lean_dec(x_13); -lean_dec(x_12); -if (x_14 == 0) +lean_object* x_87; lean_object* x_88; lean_object* x_89; +x_87 = lean_array_fget(x_78, x_83); +x_88 = lean_box(0); +x_89 = lean_array_fset(x_78, x_83, x_88); +switch (lean_obj_tag(x_87)) { +case 0: { -uint8_t x_15; -lean_dec(x_6); -x_15 = 0; -return x_15; +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_90 = lean_ctor_get(x_87, 0); +lean_inc(x_90); +x_91 = lean_ctor_get(x_87, 1); +lean_inc(x_91); +if (lean_is_exclusive(x_87)) { + lean_ctor_release(x_87, 0); + lean_ctor_release(x_87, 1); + x_92 = x_87; +} else { + lean_dec_ref(x_87); + x_92 = lean_box(0); } -else +x_93 = lean_ctor_get(x_4, 0); +lean_inc(x_93); +x_94 = lean_ctor_get(x_4, 1); +lean_inc(x_94); +x_95 = lean_ctor_get(x_90, 0); +lean_inc(x_95); +x_96 = lean_ctor_get(x_90, 1); +lean_inc(x_96); +x_97 = lean_array_get_size(x_93); +x_98 = lean_array_get_size(x_95); +x_99 = lean_nat_dec_eq(x_97, x_98); +lean_dec(x_98); +lean_dec(x_97); +if (x_99 == 0) { -lean_object* x_16; lean_object* x_17; -x_16 = lean_unsigned_to_nat(1u); -x_17 = lean_nat_add(x_6, x_16); -lean_dec(x_6); -x_3 = lean_box(0); -x_6 = x_17; -goto _start; -} -} -} +lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; +lean_dec(x_96); +lean_dec(x_95); +lean_dec(x_94); +lean_dec(x_93); +lean_dec(x_92); +x_100 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_90, x_91, x_4, x_5); +x_101 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_101, 0, x_100); +x_102 = lean_array_fset(x_89, x_83, x_101); +lean_dec(x_83); +x_103 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_103, 0, x_102); +return x_103; } -LEAN_EXPORT uint8_t l_Array_isEqvAux___at_Lean_Meta_synthInstance_x3f___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) { -_start: +else { -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) +lean_object* x_104; uint8_t x_105; +x_104 = lean_unsigned_to_nat(0u); +x_105 = l_Array_isEqvAux___at_Lean_Meta_synthInstance_x3f___spec__10(x_93, x_95, lean_box(0), x_93, x_95, x_104); +lean_dec(x_95); +lean_dec(x_93); +if (x_105 == 0) { -uint8_t x_9; -lean_dec(x_6); -x_9 = 1; -return x_9; +lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; +lean_dec(x_96); +lean_dec(x_94); +lean_dec(x_92); +x_106 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_90, x_91, x_4, x_5); +x_107 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_107, 0, x_106); +x_108 = lean_array_fset(x_89, x_83, x_107); +lean_dec(x_83); +x_109 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_109, 0, x_108); +return x_109; } else { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_10 = lean_array_fget(x_4, x_6); -x_11 = lean_array_fget(x_5, x_6); -x_12 = lean_ctor_get(x_10, 1); -lean_inc(x_12); -lean_dec(x_10); -x_13 = lean_ctor_get(x_11, 1); -lean_inc(x_13); -lean_dec(x_11); -x_14 = lean_expr_eqv(x_12, x_13); -lean_dec(x_13); -lean_dec(x_12); -if (x_14 == 0) +uint8_t x_110; +x_110 = lean_expr_eqv(x_94, x_96); +lean_dec(x_96); +lean_dec(x_94); +if (x_110 == 0) { -uint8_t x_15; -lean_dec(x_6); -x_15 = 0; -return x_15; +lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; +lean_dec(x_92); +x_111 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_90, x_91, x_4, x_5); +x_112 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_112, 0, x_111); +x_113 = lean_array_fset(x_89, x_83, x_112); +lean_dec(x_83); +x_114 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_114, 0, x_113); +return x_114; } else { -lean_object* x_16; lean_object* x_17; -x_16 = lean_unsigned_to_nat(1u); -x_17 = lean_nat_add(x_6, x_16); -lean_dec(x_6); -x_3 = lean_box(0); -x_6 = x_17; -goto _start; +lean_object* x_115; lean_object* x_116; lean_object* x_117; +lean_dec(x_91); +lean_dec(x_90); +if (lean_is_scalar(x_92)) { + x_115 = lean_alloc_ctor(0, 2, 0); +} else { + x_115 = x_92; +} +lean_ctor_set(x_115, 0, x_4); +lean_ctor_set(x_115, 1, x_5); +x_116 = lean_array_fset(x_89, x_83, x_115); +lean_dec(x_83); +x_117 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_117, 0, x_116); +return x_117; } } } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_synthInstance_x3f___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; uint8_t x_8; -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_array_get_size(x_5); -x_8 = lean_nat_dec_lt(x_2, x_7); -lean_dec(x_7); -if (x_8 == 0) +case 1: { -uint8_t x_9; -lean_dec(x_2); -x_9 = !lean_is_exclusive(x_1); -if (x_9 == 0) +lean_object* x_118; lean_object* x_119; size_t x_120; size_t x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; +x_118 = lean_ctor_get(x_87, 0); +lean_inc(x_118); +if (lean_is_exclusive(x_87)) { + lean_ctor_release(x_87, 0); + x_119 = x_87; +} else { + lean_dec_ref(x_87); + x_119 = lean_box(0); +} +x_120 = lean_usize_shift_right(x_2, x_80); +x_121 = lean_usize_add(x_3, x_79); +x_122 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_synthInstance_x3f___spec__8(x_118, x_120, x_121, x_4, x_5); +if (lean_is_scalar(x_119)) { + x_123 = lean_alloc_ctor(1, 1, 0); +} else { + x_123 = x_119; +} +lean_ctor_set(x_123, 0, x_122); +x_124 = lean_array_fset(x_89, x_83, x_123); +lean_dec(x_83); +x_125 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_125, 0, x_124); +return x_125; +} +default: { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_10 = lean_ctor_get(x_1, 1); -lean_dec(x_10); -x_11 = lean_ctor_get(x_1, 0); -lean_dec(x_11); -x_12 = lean_array_push(x_5, x_3); -x_13 = lean_array_push(x_6, x_4); -lean_ctor_set(x_1, 1, x_13); -lean_ctor_set(x_1, 0, x_12); -return x_1; +lean_object* x_126; lean_object* x_127; lean_object* x_128; +x_126 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_126, 0, x_4); +lean_ctor_set(x_126, 1, x_5); +x_127 = lean_array_fset(x_89, x_83, x_126); +lean_dec(x_83); +x_128 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_128, 0, x_127); +return x_128; +} +} } -else -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; -lean_dec(x_1); -x_14 = lean_array_push(x_5, x_3); -x_15 = lean_array_push(x_6, x_4); -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; } } 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; uint8_t x_24; -x_17 = lean_array_fget(x_5, x_2); -x_18 = lean_ctor_get(x_3, 0); -lean_inc(x_18); -x_19 = lean_ctor_get(x_3, 1); -lean_inc(x_19); -x_20 = lean_ctor_get(x_17, 0); -lean_inc(x_20); -x_21 = lean_ctor_get(x_17, 1); -lean_inc(x_21); -lean_dec(x_17); -x_22 = lean_array_get_size(x_18); -x_23 = lean_array_get_size(x_20); -x_24 = lean_nat_dec_eq(x_22, x_23); -lean_dec(x_23); -lean_dec(x_22); -if (x_24 == 0) +uint8_t x_129; +x_129 = !lean_is_exclusive(x_1); +if (x_129 == 0) { -lean_object* x_25; lean_object* x_26; -lean_dec(x_21); -lean_dec(x_20); -lean_dec(x_19); -lean_dec(x_18); -lean_dec(x_6); -lean_dec(x_5); -x_25 = lean_unsigned_to_nat(1u); -x_26 = lean_nat_add(x_2, x_25); -lean_dec(x_2); -x_2 = x_26; -goto _start; -} -else +lean_object* x_130; lean_object* x_131; size_t x_132; uint8_t x_133; +x_130 = lean_unsigned_to_nat(0u); +x_131 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_synthInstance_x3f___spec__11(x_1, x_130, x_4, x_5); +x_132 = 7; +x_133 = lean_usize_dec_le(x_132, x_3); +if (x_133 == 0) { -lean_object* x_28; uint8_t x_29; -x_28 = lean_unsigned_to_nat(0u); -x_29 = l_Array_isEqvAux___at_Lean_Meta_synthInstance_x3f___spec__12(x_18, x_20, lean_box(0), x_18, x_20, x_28); -lean_dec(x_20); -lean_dec(x_18); -if (x_29 == 0) +lean_object* x_134; lean_object* x_135; uint8_t x_136; +x_134 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_131); +x_135 = lean_unsigned_to_nat(4u); +x_136 = lean_nat_dec_lt(x_134, x_135); +lean_dec(x_134); +if (x_136 == 0) { -lean_object* x_30; lean_object* x_31; -lean_dec(x_21); -lean_dec(x_19); -lean_dec(x_6); -lean_dec(x_5); -x_30 = lean_unsigned_to_nat(1u); -x_31 = lean_nat_add(x_2, x_30); -lean_dec(x_2); -x_2 = x_31; -goto _start; +lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; +x_137 = lean_ctor_get(x_131, 0); +lean_inc(x_137); +x_138 = lean_ctor_get(x_131, 1); +lean_inc(x_138); +lean_dec(x_131); +x_139 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_synthInstance_x3f___spec__8___closed__1; +x_140 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_synthInstance_x3f___spec__9(x_3, x_137, x_138, lean_box(0), x_130, x_139); +lean_dec(x_138); +lean_dec(x_137); +return x_140; } else { -uint8_t x_33; -x_33 = lean_expr_eqv(x_19, x_21); -lean_dec(x_21); -lean_dec(x_19); -if (x_33 == 0) +return x_131; +} +} +else { -lean_object* x_34; lean_object* x_35; -lean_dec(x_6); -lean_dec(x_5); -x_34 = lean_unsigned_to_nat(1u); -x_35 = lean_nat_add(x_2, x_34); -lean_dec(x_2); -x_2 = x_35; -goto _start; +return x_131; +} } else { -uint8_t x_37; -x_37 = !lean_is_exclusive(x_1); -if (x_37 == 0) +lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; size_t x_146; uint8_t x_147; +x_141 = lean_ctor_get(x_1, 0); +x_142 = lean_ctor_get(x_1, 1); +lean_inc(x_142); +lean_inc(x_141); +lean_dec(x_1); +x_143 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_143, 0, x_141); +lean_ctor_set(x_143, 1, x_142); +x_144 = lean_unsigned_to_nat(0u); +x_145 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_synthInstance_x3f___spec__11(x_143, x_144, x_4, x_5); +x_146 = 7; +x_147 = lean_usize_dec_le(x_146, x_3); +if (x_147 == 0) { -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_38 = lean_ctor_get(x_1, 1); -lean_dec(x_38); -x_39 = lean_ctor_get(x_1, 0); -lean_dec(x_39); -x_40 = lean_array_fset(x_5, x_2, x_3); -x_41 = lean_array_fset(x_6, x_2, x_4); -lean_dec(x_2); -lean_ctor_set(x_1, 1, x_41); -lean_ctor_set(x_1, 0, x_40); -return x_1; +lean_object* x_148; lean_object* x_149; uint8_t x_150; +x_148 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_145); +x_149 = lean_unsigned_to_nat(4u); +x_150 = lean_nat_dec_lt(x_148, x_149); +lean_dec(x_148); +if (x_150 == 0) +{ +lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; +x_151 = lean_ctor_get(x_145, 0); +lean_inc(x_151); +x_152 = lean_ctor_get(x_145, 1); +lean_inc(x_152); +lean_dec(x_145); +x_153 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_synthInstance_x3f___spec__8___closed__1; +x_154 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_synthInstance_x3f___spec__9(x_3, x_151, x_152, lean_box(0), x_144, x_153); +lean_dec(x_152); +lean_dec(x_151); +return x_154; } else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; -lean_dec(x_1); -x_42 = lean_array_fset(x_5, x_2, x_3); -x_43 = lean_array_fset(x_6, x_2, x_4); -lean_dec(x_2); -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; +return x_145; } } +else +{ +return x_145; } } } } } -static lean_object* _init_l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_synthInstance_x3f___spec__8___closed__1() { +static lean_object* _init_l_Lean_PersistentHashMap_insert___at_Lean_Meta_synthInstance_x3f___spec__7___closed__1() { _start: { lean_object* x_1; -x_1 = l_Lean_PersistentHashMap_mkEmptyEntries(lean_box(0), lean_box(0)); +x_1 = lean_alloc_closure((void*)(l_Lean_instBEqLocalInstance___boxed), 2, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_synthInstance_x3f___spec__8(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5) { +static lean_object* _init_l_Lean_PersistentHashMap_insert___at_Lean_Meta_synthInstance_x3f___spec__7___closed__2() { _start: { -if (lean_obj_tag(x_1) == 0) -{ -uint8_t x_6; -x_6 = !lean_is_exclusive(x_1); -if (x_6 == 0) -{ -lean_object* x_7; size_t x_8; size_t x_9; size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_7 = lean_ctor_get(x_1, 0); -x_8 = 1; -x_9 = 5; -x_10 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_synthInstance_x3f___spec__2___closed__2; -x_11 = lean_usize_land(x_2, x_10); -x_12 = lean_usize_to_nat(x_11); -x_13 = lean_array_get_size(x_7); -x_14 = lean_nat_dec_lt(x_12, x_13); -lean_dec(x_13); -if (x_14 == 0) -{ -lean_dec(x_12); -lean_dec(x_5); -lean_dec(x_4); -return x_1; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_synthInstance_x3f___spec__7___closed__1; +x_2 = lean_alloc_closure((void*)(l_Array_instBEqArray___rarg___boxed), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; } -else -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_array_fget(x_7, x_12); -x_16 = lean_box(0); -x_17 = lean_array_fset(x_7, x_12, x_16); -switch (lean_obj_tag(x_15)) { -case 0: -{ -uint8_t x_18; -x_18 = !lean_is_exclusive(x_15); -if (x_18 == 0) +} +static lean_object* _init_l_Lean_PersistentHashMap_insert___at_Lean_Meta_synthInstance_x3f___spec__7___closed__3() { +_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; uint8_t x_27; -x_19 = lean_ctor_get(x_15, 0); -x_20 = lean_ctor_get(x_15, 1); -x_21 = lean_ctor_get(x_4, 0); -lean_inc(x_21); -x_22 = lean_ctor_get(x_4, 1); -lean_inc(x_22); -x_23 = lean_ctor_get(x_19, 0); -lean_inc(x_23); -x_24 = lean_ctor_get(x_19, 1); -lean_inc(x_24); -x_25 = lean_array_get_size(x_21); -x_26 = lean_array_get_size(x_23); -x_27 = lean_nat_dec_eq(x_25, x_26); -lean_dec(x_26); -lean_dec(x_25); -if (x_27 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_synthInstance_x3f___spec__7___closed__2; +x_2 = l_Lean_Expr_instBEqExpr; +x_3 = lean_alloc_closure((void*)(l_instBEqProd___rarg), 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_PersistentHashMap_insert___at_Lean_Meta_synthInstance_x3f___spec__7___closed__4() { +_start: { -lean_object* x_28; lean_object* x_29; lean_object* x_30; -lean_dec(x_24); -lean_dec(x_23); -lean_dec(x_22); -lean_dec(x_21); -lean_free_object(x_15); -x_28 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_19, x_20, x_4, x_5); -x_29 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_29, 0, x_28); -x_30 = lean_array_fset(x_17, x_12, x_29); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_30); +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_instHashableLocalInstance___boxed), 1, 0); return x_1; } -else +} +static lean_object* _init_l_Lean_PersistentHashMap_insert___at_Lean_Meta_synthInstance_x3f___spec__7___closed__5() { +_start: { -lean_object* x_31; uint8_t x_32; -x_31 = lean_unsigned_to_nat(0u); -x_32 = l_Array_isEqvAux___at_Lean_Meta_synthInstance_x3f___spec__10(x_21, x_23, lean_box(0), x_21, x_23, x_31); -lean_dec(x_23); -lean_dec(x_21); -if (x_32 == 0) +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_synthInstance_x3f___spec__7___closed__4; +x_2 = lean_alloc_closure((void*)(l_instHashableArray___rarg___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_PersistentHashMap_insert___at_Lean_Meta_synthInstance_x3f___spec__7___closed__6() { +_start: { -lean_object* x_33; lean_object* x_34; lean_object* x_35; -lean_dec(x_24); -lean_dec(x_22); -lean_free_object(x_15); -x_33 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_19, x_20, x_4, x_5); -x_34 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_34, 0, x_33); -x_35 = lean_array_fset(x_17, x_12, x_34); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_35); +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_synthInstance_x3f___spec__7___closed__5; +x_2 = l_Lean_Expr_instHashableExpr; +x_3 = lean_alloc_closure((void*)(l_instHashableProd___rarg___boxed), 3, 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_PersistentHashMap_insert___at_Lean_Meta_synthInstance_x3f___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_1); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; size_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint64_t x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; uint64_t x_16; +x_5 = lean_ctor_get(x_1, 0); +x_6 = lean_ctor_get(x_1, 1); +x_7 = 1; +x_8 = lean_unsigned_to_nat(1u); +x_9 = lean_nat_add(x_6, x_8); +lean_dec(x_6); +x_10 = lean_ctor_get(x_2, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_2, 1); +lean_inc(x_11); +x_12 = 7; +x_13 = lean_array_get_size(x_10); +x_14 = lean_unsigned_to_nat(0u); +x_15 = lean_nat_dec_lt(x_14, x_13); +x_16 = l_Lean_Expr_hash(x_11); +lean_dec(x_11); +if (x_15 == 0) +{ +uint64_t x_17; size_t x_18; lean_object* x_19; +lean_dec(x_13); +lean_dec(x_10); +x_17 = lean_uint64_mix_hash(x_12, x_16); +x_18 = lean_uint64_to_usize(x_17); +x_19 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_synthInstance_x3f___spec__8(x_5, x_18, x_7, x_2, x_3); +lean_ctor_set(x_1, 1, x_9); +lean_ctor_set(x_1, 0, x_19); return x_1; } else { -uint8_t x_36; -x_36 = lean_expr_eqv(x_22, x_24); -lean_dec(x_24); -lean_dec(x_22); -if (x_36 == 0) +uint8_t x_20; +x_20 = lean_nat_dec_le(x_13, x_13); +if (x_20 == 0) { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -lean_free_object(x_15); -x_37 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_19, x_20, x_4, x_5); -x_38 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_38, 0, x_37); -x_39 = lean_array_fset(x_17, x_12, x_38); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_39); +uint64_t x_21; size_t x_22; lean_object* x_23; +lean_dec(x_13); +lean_dec(x_10); +x_21 = lean_uint64_mix_hash(x_12, x_16); +x_22 = lean_uint64_to_usize(x_21); +x_23 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_synthInstance_x3f___spec__8(x_5, x_22, x_7, x_2, x_3); +lean_ctor_set(x_1, 1, x_9); +lean_ctor_set(x_1, 0, x_23); return x_1; } else { -lean_object* x_40; -lean_dec(x_20); -lean_dec(x_19); -lean_ctor_set(x_15, 1, x_5); -lean_ctor_set(x_15, 0, x_4); -x_40 = lean_array_fset(x_17, x_12, x_15); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_40); +size_t x_24; size_t x_25; uint64_t x_26; uint64_t x_27; size_t x_28; lean_object* x_29; +x_24 = 0; +x_25 = lean_usize_of_nat(x_13); +lean_dec(x_13); +x_26 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_synthInstance_x3f___spec__6(x_10, x_24, x_25, x_12); +lean_dec(x_10); +x_27 = lean_uint64_mix_hash(x_26, x_16); +x_28 = lean_uint64_to_usize(x_27); +x_29 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_synthInstance_x3f___spec__8(x_5, x_28, x_7, x_2, x_3); +lean_ctor_set(x_1, 1, x_9); +lean_ctor_set(x_1, 0, x_29); return x_1; } } } -} 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; lean_object* x_47; lean_object* x_48; uint8_t x_49; -x_41 = lean_ctor_get(x_15, 0); -x_42 = lean_ctor_get(x_15, 1); -lean_inc(x_42); -lean_inc(x_41); -lean_dec(x_15); -x_43 = lean_ctor_get(x_4, 0); -lean_inc(x_43); -x_44 = lean_ctor_get(x_4, 1); -lean_inc(x_44); -x_45 = lean_ctor_get(x_41, 0); -lean_inc(x_45); -x_46 = lean_ctor_get(x_41, 1); -lean_inc(x_46); -x_47 = lean_array_get_size(x_43); -x_48 = lean_array_get_size(x_45); -x_49 = lean_nat_dec_eq(x_47, x_48); -lean_dec(x_48); -lean_dec(x_47); -if (x_49 == 0) +lean_object* x_30; lean_object* x_31; size_t x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint64_t x_37; lean_object* x_38; lean_object* x_39; uint8_t x_40; uint64_t x_41; +x_30 = lean_ctor_get(x_1, 0); +x_31 = lean_ctor_get(x_1, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_1); +x_32 = 1; +x_33 = lean_unsigned_to_nat(1u); +x_34 = lean_nat_add(x_31, x_33); +lean_dec(x_31); +x_35 = lean_ctor_get(x_2, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_2, 1); +lean_inc(x_36); +x_37 = 7; +x_38 = lean_array_get_size(x_35); +x_39 = lean_unsigned_to_nat(0u); +x_40 = lean_nat_dec_lt(x_39, x_38); +x_41 = l_Lean_Expr_hash(x_36); +lean_dec(x_36); +if (x_40 == 0) { -lean_object* x_50; lean_object* x_51; lean_object* x_52; -lean_dec(x_46); -lean_dec(x_45); -lean_dec(x_44); -lean_dec(x_43); -x_50 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_41, x_42, x_4, x_5); -x_51 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_51, 0, x_50); -x_52 = lean_array_fset(x_17, x_12, x_51); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_52); -return x_1; +uint64_t x_42; size_t x_43; lean_object* x_44; lean_object* x_45; +lean_dec(x_38); +lean_dec(x_35); +x_42 = lean_uint64_mix_hash(x_37, x_41); +x_43 = lean_uint64_to_usize(x_42); +x_44 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_synthInstance_x3f___spec__8(x_30, x_43, x_32, x_2, x_3); +x_45 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_45, 0, x_44); +lean_ctor_set(x_45, 1, x_34); +return x_45; } else { -lean_object* x_53; uint8_t x_54; -x_53 = lean_unsigned_to_nat(0u); -x_54 = l_Array_isEqvAux___at_Lean_Meta_synthInstance_x3f___spec__10(x_43, x_45, lean_box(0), x_43, x_45, x_53); -lean_dec(x_45); -lean_dec(x_43); -if (x_54 == 0) +uint8_t x_46; +x_46 = lean_nat_dec_le(x_38, x_38); +if (x_46 == 0) { -lean_object* x_55; lean_object* x_56; lean_object* x_57; -lean_dec(x_46); -lean_dec(x_44); -x_55 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_41, x_42, x_4, x_5); -x_56 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_56, 0, x_55); -x_57 = lean_array_fset(x_17, x_12, x_56); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_57); -return x_1; +uint64_t x_47; size_t x_48; lean_object* x_49; lean_object* x_50; +lean_dec(x_38); +lean_dec(x_35); +x_47 = lean_uint64_mix_hash(x_37, x_41); +x_48 = lean_uint64_to_usize(x_47); +x_49 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_synthInstance_x3f___spec__8(x_30, x_48, x_32, x_2, x_3); +x_50 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_50, 0, x_49); +lean_ctor_set(x_50, 1, x_34); +return x_50; } else { -uint8_t x_58; -x_58 = lean_expr_eqv(x_44, x_46); -lean_dec(x_46); -lean_dec(x_44); -if (x_58 == 0) +size_t x_51; size_t x_52; uint64_t x_53; uint64_t x_54; size_t x_55; lean_object* x_56; lean_object* x_57; +x_51 = 0; +x_52 = lean_usize_of_nat(x_38); +lean_dec(x_38); +x_53 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_synthInstance_x3f___spec__6(x_35, x_51, x_52, x_37); +lean_dec(x_35); +x_54 = lean_uint64_mix_hash(x_53, x_41); +x_55 = lean_uint64_to_usize(x_54); +x_56 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_synthInstance_x3f___spec__8(x_30, x_55, x_32, x_2, x_3); +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; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_withSeconds___at_Lean_Meta_synthInstance_x3f___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) { +_start: { -lean_object* x_59; lean_object* x_60; lean_object* x_61; -x_59 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_41, x_42, x_4, x_5); -x_60 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_60, 0, x_59); -x_61 = lean_array_fset(x_17, x_12, x_60); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_61); -return x_1; +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_7 = lean_io_mono_nanos_now(x_6); +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 = lean_apply_5(x_1, x_2, x_3, x_4, x_5, x_9); +if (lean_obj_tag(x_10) == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +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_io_mono_nanos_now(x_12); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; double x_19; double x_20; double x_21; lean_object* x_22; lean_object* x_23; +x_15 = lean_ctor_get(x_13, 0); +x_16 = lean_nat_sub(x_15, x_8); +lean_dec(x_8); +lean_dec(x_15); +x_17 = 0; +x_18 = lean_unsigned_to_nat(0u); +x_19 = l_Float_ofScientific(x_16, x_17, x_18); +lean_dec(x_16); +x_20 = l_Lean_withSeconds___at_Lean_Meta_SynthInstance_newSubgoal___spec__12___closed__1; +x_21 = lean_float_div(x_19, x_20); +x_22 = lean_box_float(x_21); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_11); +lean_ctor_set(x_23, 1, x_22); +lean_ctor_set(x_13, 0, x_23); +return x_13; } else { -lean_object* x_62; lean_object* x_63; -lean_dec(x_42); -lean_dec(x_41); -x_62 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_62, 0, x_4); -lean_ctor_set(x_62, 1, x_5); -x_63 = lean_array_fset(x_17, x_12, x_62); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_63); -return x_1; -} -} -} +lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; lean_object* x_28; double x_29; double x_30; double x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_24 = lean_ctor_get(x_13, 0); +x_25 = lean_ctor_get(x_13, 1); +lean_inc(x_25); +lean_inc(x_24); +lean_dec(x_13); +x_26 = lean_nat_sub(x_24, x_8); +lean_dec(x_8); +lean_dec(x_24); +x_27 = 0; +x_28 = lean_unsigned_to_nat(0u); +x_29 = l_Float_ofScientific(x_26, x_27, x_28); +lean_dec(x_26); +x_30 = l_Lean_withSeconds___at_Lean_Meta_SynthInstance_newSubgoal___spec__12___closed__1; +x_31 = lean_float_div(x_29, x_30); +x_32 = lean_box_float(x_31); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_11); +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_25); +return x_34; } } -case 1: +else { -uint8_t x_64; -x_64 = !lean_is_exclusive(x_15); -if (x_64 == 0) +uint8_t x_35; +lean_dec(x_8); +x_35 = !lean_is_exclusive(x_10); +if (x_35 == 0) { -lean_object* x_65; size_t x_66; size_t x_67; lean_object* x_68; lean_object* x_69; -x_65 = lean_ctor_get(x_15, 0); -x_66 = lean_usize_shift_right(x_2, x_9); -x_67 = lean_usize_add(x_3, x_8); -x_68 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_synthInstance_x3f___spec__8(x_65, x_66, x_67, x_4, x_5); -lean_ctor_set(x_15, 0, x_68); -x_69 = lean_array_fset(x_17, x_12, x_15); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_69); -return x_1; +return x_10; } else { -lean_object* x_70; size_t x_71; size_t x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; -x_70 = lean_ctor_get(x_15, 0); -lean_inc(x_70); -lean_dec(x_15); -x_71 = lean_usize_shift_right(x_2, x_9); -x_72 = lean_usize_add(x_3, x_8); -x_73 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_synthInstance_x3f___spec__8(x_70, x_71, x_72, x_4, x_5); -x_74 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_74, 0, x_73); -x_75 = lean_array_fset(x_17, x_12, x_74); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_75); -return x_1; -} -} -default: -{ -lean_object* x_76; lean_object* x_77; -x_76 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_76, 0, x_4); -lean_ctor_set(x_76, 1, x_5); -x_77 = lean_array_fset(x_17, x_12, x_76); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_77); -return x_1; +lean_object* x_36; lean_object* x_37; lean_object* x_38; +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 = 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_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_synthInstance_x3f___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) { +_start: { -lean_object* x_78; size_t x_79; size_t x_80; size_t x_81; size_t x_82; lean_object* x_83; lean_object* x_84; uint8_t x_85; -x_78 = lean_ctor_get(x_1, 0); -lean_inc(x_78); -lean_dec(x_1); -x_79 = 1; -x_80 = 5; -x_81 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_synthInstance_x3f___spec__2___closed__2; -x_82 = lean_usize_land(x_2, x_81); -x_83 = lean_usize_to_nat(x_82); -x_84 = lean_array_get_size(x_78); -x_85 = lean_nat_dec_lt(x_83, x_84); -lean_dec(x_84); -if (x_85 == 0) +if (lean_obj_tag(x_1) == 0) { -lean_object* x_86; -lean_dec(x_83); -lean_dec(x_5); -lean_dec(x_4); -x_86 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_86, 0, x_78); -return x_86; +lean_object* x_7; lean_object* x_8; +x_7 = lean_ctor_get(x_1, 0); +lean_inc(x_7); +x_8 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_8, 0, x_7); +lean_ctor_set(x_8, 1, x_6); +return x_8; } else { -lean_object* x_87; lean_object* x_88; lean_object* x_89; -x_87 = lean_array_fget(x_78, x_83); -x_88 = lean_box(0); -x_89 = lean_array_fset(x_78, x_83, x_88); -switch (lean_obj_tag(x_87)) { -case 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_object* x_97; lean_object* x_98; uint8_t x_99; -x_90 = lean_ctor_get(x_87, 0); -lean_inc(x_90); -x_91 = lean_ctor_get(x_87, 1); -lean_inc(x_91); -if (lean_is_exclusive(x_87)) { - lean_ctor_release(x_87, 0); - lean_ctor_release(x_87, 1); - x_92 = x_87; -} else { - lean_dec_ref(x_87); - x_92 = lean_box(0); +lean_object* x_9; lean_object* x_10; +x_9 = lean_ctor_get(x_1, 0); +lean_inc(x_9); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_6); +return x_10; } -x_93 = lean_ctor_get(x_4, 0); -lean_inc(x_93); -x_94 = lean_ctor_get(x_4, 1); -lean_inc(x_94); -x_95 = lean_ctor_get(x_90, 0); -lean_inc(x_95); -x_96 = lean_ctor_get(x_90, 1); -lean_inc(x_96); -x_97 = lean_array_get_size(x_93); -x_98 = lean_array_get_size(x_95); -x_99 = lean_nat_dec_eq(x_97, x_98); -lean_dec(x_98); -lean_dec(x_97); -if (x_99 == 0) -{ -lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; -lean_dec(x_96); -lean_dec(x_95); -lean_dec(x_94); -lean_dec(x_93); -lean_dec(x_92); -x_100 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_90, x_91, x_4, x_5); -x_101 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_101, 0, x_100); -x_102 = lean_array_fset(x_89, x_83, x_101); -lean_dec(x_83); -x_103 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_103, 0, x_102); -return x_103; } -else -{ -lean_object* x_104; uint8_t x_105; -x_104 = lean_unsigned_to_nat(0u); -x_105 = l_Array_isEqvAux___at_Lean_Meta_synthInstance_x3f___spec__10(x_93, x_95, lean_box(0), x_93, x_95, x_104); -lean_dec(x_95); -lean_dec(x_93); -if (x_105 == 0) +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_synthInstance_x3f___spec__13___lambda__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) { +_start: { -lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; -lean_dec(x_96); -lean_dec(x_94); -lean_dec(x_92); -x_106 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_90, x_91, x_4, x_5); -x_107 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_107, 0, x_106); -x_108 = lean_array_fset(x_89, x_83, x_107); -lean_dec(x_83); -x_109 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_109, 0, x_108); -return x_109; +lean_object* x_13; lean_object* x_14; lean_object* x_15; +lean_inc(x_10); +x_13 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_processPostponed___spec__4(x_1, x_2, x_3, x_6, x_4, x_8, x_9, x_10, x_11, x_12); +x_14 = lean_ctor_get(x_13, 1); +lean_inc(x_14); +lean_dec(x_13); +x_15 = l_MonadExcept_ofExcept___at_Lean_Meta_synthInstance_x3f___spec__15(x_5, x_8, x_9, x_10, x_11, x_14); +lean_dec(x_10); +return x_15; } -else +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_synthInstance_x3f___spec__13___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, uint8_t x_7, double 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_110; -x_110 = lean_expr_eqv(x_94, x_96); -lean_dec(x_96); -lean_dec(x_94); -if (x_110 == 0) +lean_object* x_15; lean_object* x_16; +lean_dec(x_9); +x_15 = lean_ctor_get(x_12, 5); +lean_inc(x_15); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_2); +x_16 = lean_apply_6(x_1, x_2, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; -lean_dec(x_92); -x_111 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_90, x_91, x_4, x_5); -x_112 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_112, 0, x_111); -x_113 = lean_array_fset(x_89, x_83, x_112); -lean_dec(x_83); -x_114 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_114, 0, x_113); -return x_114; +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); +x_19 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__1; +x_20 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_6, x_19); +lean_dec(x_6); +if (x_20 == 0) +{ +if (x_7 == 0) +{ +lean_object* x_21; lean_object* x_22; +x_21 = lean_box(0); +x_22 = l_Lean_withTraceNode___at_Lean_Meta_synthInstance_x3f___spec__13___lambda__1(x_3, x_4, x_15, x_5, x_2, x_17, x_21, x_10, x_11, x_12, x_13, x_18); +lean_dec(x_13); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_2); +return x_22; } else { -lean_object* x_115; lean_object* x_116; lean_object* x_117; -lean_dec(x_91); -lean_dec(x_90); -if (lean_is_scalar(x_92)) { - x_115 = lean_alloc_ctor(0, 2, 0); -} else { - x_115 = x_92; -} -lean_ctor_set(x_115, 0, x_4); -lean_ctor_set(x_115, 1, x_5); -x_116 = lean_array_fset(x_89, x_83, x_115); -lean_dec(x_83); -x_117 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_117, 0, x_116); -return x_117; -} -} +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; 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_23 = lean_float_to_string(x_8); +x_24 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_24, 0, x_23); +x_25 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_25, 0, x_24); +x_26 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__3; +x_27 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_25); +x_28 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__5; +x_29 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +x_30 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_17); +x_31 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; +x_32 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +x_33 = lean_box(0); +x_34 = l_Lean_withTraceNode___at_Lean_Meta_synthInstance_x3f___spec__13___lambda__1(x_3, x_4, x_15, x_5, x_2, x_32, x_33, x_10, x_11, x_12, x_13, x_18); +lean_dec(x_13); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_2); +return x_34; } } -case 1: +else { -lean_object* x_118; lean_object* x_119; size_t x_120; size_t x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; -x_118 = lean_ctor_get(x_87, 0); -lean_inc(x_118); -if (lean_is_exclusive(x_87)) { - lean_ctor_release(x_87, 0); - x_119 = x_87; -} else { - lean_dec_ref(x_87); - x_119 = lean_box(0); -} -x_120 = lean_usize_shift_right(x_2, x_80); -x_121 = lean_usize_add(x_3, x_79); -x_122 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_synthInstance_x3f___spec__8(x_118, x_120, x_121, x_4, x_5); -if (lean_is_scalar(x_119)) { - x_123 = lean_alloc_ctor(1, 1, 0); -} else { - x_123 = x_119; +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; 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_float_to_string(x_8); +x_36 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_36, 0, x_35); +x_37 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_37, 0, x_36); +x_38 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__3; +x_39 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_37); +x_40 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__5; +x_41 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_40); +x_42 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_17); +x_43 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; +x_44 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_44, 0, x_42); +lean_ctor_set(x_44, 1, x_43); +x_45 = lean_box(0); +x_46 = l_Lean_withTraceNode___at_Lean_Meta_synthInstance_x3f___spec__13___lambda__1(x_3, x_4, x_15, x_5, x_2, x_44, x_45, x_10, x_11, x_12, x_13, x_18); +lean_dec(x_13); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_2); +return x_46; } -lean_ctor_set(x_123, 0, x_122); -x_124 = lean_array_fset(x_89, x_83, x_123); -lean_dec(x_83); -x_125 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_125, 0, x_124); -return x_125; } -default: +else { -lean_object* x_126; lean_object* x_127; lean_object* x_128; -x_126 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_126, 0, x_4); -lean_ctor_set(x_126, 1, x_5); -x_127 = lean_array_fset(x_89, x_83, x_126); -lean_dec(x_83); -x_128 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_128, 0, x_127); -return x_128; +uint8_t x_47; +lean_dec(x_15); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_47 = !lean_is_exclusive(x_16); +if (x_47 == 0) +{ +return x_16; } +else +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_16, 0); +x_49 = lean_ctor_get(x_16, 1); +lean_inc(x_49); +lean_inc(x_48); +lean_dec(x_16); +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; } } } } -else -{ -uint8_t x_129; -x_129 = !lean_is_exclusive(x_1); -if (x_129 == 0) +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_synthInstance_x3f___spec__13___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t 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, lean_object* x_12) { +_start: { -lean_object* x_130; lean_object* x_131; size_t x_132; uint8_t x_133; -x_130 = lean_unsigned_to_nat(0u); -x_131 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_synthInstance_x3f___spec__11(x_1, x_130, x_4, x_5); -x_132 = 7; -x_133 = lean_usize_dec_le(x_132, x_3); -if (x_133 == 0) +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +lean_dec(x_7); +x_13 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_processPostponed___spec__2___rarg(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_alloc_closure((void*)(l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_tryResolve___spec__2___lambda__1), 6, 1); +lean_closure_set(x_16, 0, x_1); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_17 = l_Lean_withSeconds___at_Lean_Meta_synthInstance_x3f___spec__14(x_16, x_8, x_9, x_10, x_11, x_15); +if (lean_obj_tag(x_17) == 0) { -lean_object* x_134; lean_object* x_135; uint8_t x_136; -x_134 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_131); -x_135 = lean_unsigned_to_nat(4u); -x_136 = lean_nat_dec_lt(x_134, x_135); -lean_dec(x_134); -if (x_136 == 0) +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; lean_object* x_67; uint8_t x_68; +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_ctor_get(x_18, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_18, 1); +lean_inc(x_21); +lean_dec(x_18); +x_67 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__4___closed__1; +x_68 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_5, x_67); +if (x_68 == 0) { -lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; -x_137 = lean_ctor_get(x_131, 0); -lean_inc(x_137); -x_138 = lean_ctor_get(x_131, 1); -lean_inc(x_138); -lean_dec(x_131); -x_139 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_synthInstance_x3f___spec__8___closed__1; -x_140 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_synthInstance_x3f___spec__9(x_3, x_137, x_138, lean_box(0), x_130, x_139); -lean_dec(x_138); -lean_dec(x_137); -return x_140; +uint8_t x_69; +x_69 = 0; +x_22 = x_69; +goto block_66; } else { -return x_131; -} +double x_70; double x_71; uint8_t x_72; +x_70 = l_Lean_trace_profiler_threshold_getSecs(x_5); +x_71 = lean_unbox_float(x_21); +x_72 = lean_float_decLt(x_70, x_71); +x_22 = x_72; +goto block_66; } -else +block_66: { -return x_131; -} -} -else +if (x_6 == 0) { -lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; size_t x_146; uint8_t x_147; -x_141 = lean_ctor_get(x_1, 0); -x_142 = lean_ctor_get(x_1, 1); -lean_inc(x_142); -lean_inc(x_141); -lean_dec(x_1); -x_143 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_143, 0, x_141); -lean_ctor_set(x_143, 1, x_142); -x_144 = lean_unsigned_to_nat(0u); -x_145 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_synthInstance_x3f___spec__11(x_143, x_144, x_4, x_5); -x_146 = 7; -x_147 = lean_usize_dec_le(x_146, x_3); -if (x_147 == 0) +if (x_22 == 0) { -lean_object* x_148; lean_object* x_149; uint8_t x_150; -x_148 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_145); -x_149 = lean_unsigned_to_nat(4u); -x_150 = lean_nat_dec_lt(x_148, x_149); -lean_dec(x_148); -if (x_150 == 0) +lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; +lean_dec(x_21); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +x_23 = lean_st_ref_take(x_11, 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_is_exclusive(x_24); +if (x_26 == 0) { -lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; -x_151 = lean_ctor_get(x_145, 0); -lean_inc(x_151); -x_152 = lean_ctor_get(x_145, 1); -lean_inc(x_152); -lean_dec(x_145); -x_153 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_synthInstance_x3f___spec__8___closed__1; -x_154 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_synthInstance_x3f___spec__9(x_3, x_151, x_152, lean_box(0), x_144, x_153); -lean_dec(x_152); -lean_dec(x_151); -return x_154; +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_24, 3); +x_28 = l_Lean_PersistentArray_append___rarg(x_14, x_27); +lean_ctor_set(x_24, 3, x_28); +x_29 = lean_st_ref_set(x_11, x_24, x_25); +x_30 = lean_ctor_get(x_29, 1); +lean_inc(x_30); +lean_dec(x_29); +x_31 = l_MonadExcept_ofExcept___at_Lean_Meta_synthInstance_x3f___spec__15(x_20, x_8, x_9, x_10, x_11, x_30); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_20); +if (lean_obj_tag(x_31) == 0) +{ +uint8_t x_32; +x_32 = !lean_is_exclusive(x_31); +if (x_32 == 0) +{ +return x_31; +} +else +{ +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(0, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +return x_35; +} } else { -return x_145; -} +uint8_t x_36; +x_36 = !lean_is_exclusive(x_31); +if (x_36 == 0) +{ +return x_31; } else { -return x_145; +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_31, 0); +x_38 = lean_ctor_get(x_31, 1); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_31); +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_40; lean_object* x_41; 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_40 = lean_ctor_get(x_24, 0); +x_41 = lean_ctor_get(x_24, 1); +x_42 = lean_ctor_get(x_24, 2); +x_43 = lean_ctor_get(x_24, 3); +x_44 = lean_ctor_get(x_24, 4); +x_45 = lean_ctor_get(x_24, 5); +x_46 = lean_ctor_get(x_24, 6); +lean_inc(x_46); +lean_inc(x_45); +lean_inc(x_44); +lean_inc(x_43); +lean_inc(x_42); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_24); +x_47 = l_Lean_PersistentArray_append___rarg(x_14, x_43); +x_48 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_48, 0, x_40); +lean_ctor_set(x_48, 1, x_41); +lean_ctor_set(x_48, 2, x_42); +lean_ctor_set(x_48, 3, x_47); +lean_ctor_set(x_48, 4, x_44); +lean_ctor_set(x_48, 5, x_45); +lean_ctor_set(x_48, 6, x_46); +x_49 = lean_st_ref_set(x_11, x_48, x_25); +x_50 = lean_ctor_get(x_49, 1); +lean_inc(x_50); +lean_dec(x_49); +x_51 = l_MonadExcept_ofExcept___at_Lean_Meta_synthInstance_x3f___spec__15(x_20, x_8, x_9, x_10, x_11, x_50); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_20); +if (lean_obj_tag(x_51) == 0) +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_52 = lean_ctor_get(x_51, 0); +lean_inc(x_52); +x_53 = lean_ctor_get(x_51, 1); +lean_inc(x_53); +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_51); + x_54 = lean_box(0); } +if (lean_is_scalar(x_54)) { + x_55 = lean_alloc_ctor(0, 2, 0); +} else { + x_55 = x_54; } +lean_ctor_set(x_55, 0, x_52); +lean_ctor_set(x_55, 1, x_53); +return x_55; } -static lean_object* _init_l_Lean_PersistentHashMap_insert___at_Lean_Meta_synthInstance_x3f___spec__7___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_instBEqLocalInstance___boxed), 2, 0); -return x_1; +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_56 = lean_ctor_get(x_51, 0); +lean_inc(x_56); +x_57 = lean_ctor_get(x_51, 1); +lean_inc(x_57); +if (lean_is_exclusive(x_51)) { + lean_ctor_release(x_51, 0); + lean_ctor_release(x_51, 1); + x_58 = x_51; +} else { + lean_dec_ref(x_51); + x_58 = lean_box(0); +} +if (lean_is_scalar(x_58)) { + x_59 = lean_alloc_ctor(1, 2, 0); +} else { + x_59 = x_58; } +lean_ctor_set(x_59, 0, x_56); +lean_ctor_set(x_59, 1, x_57); +return x_59; } -static lean_object* _init_l_Lean_PersistentHashMap_insert___at_Lean_Meta_synthInstance_x3f___spec__7___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_synthInstance_x3f___spec__7___closed__1; -x_2 = lean_alloc_closure((void*)(l_Array_instBEqArray___rarg___boxed), 3, 1); -lean_closure_set(x_2, 0, x_1); -return x_2; } } -static lean_object* _init_l_Lean_PersistentHashMap_insert___at_Lean_Meta_synthInstance_x3f___spec__7___closed__3() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_synthInstance_x3f___spec__7___closed__2; -x_2 = l_Lean_Expr_instBEqExpr; -x_3 = lean_alloc_closure((void*)(l_instBEqProd___rarg), 4, 2); -lean_closure_set(x_3, 0, x_1); -lean_closure_set(x_3, 1, x_2); -return x_3; +lean_object* x_60; double x_61; lean_object* x_62; +x_60 = lean_box(0); +x_61 = lean_unbox_float(x_21); +lean_dec(x_21); +x_62 = l_Lean_withTraceNode___at_Lean_Meta_synthInstance_x3f___spec__13___lambda__2(x_2, x_20, x_14, x_3, x_4, x_5, x_22, x_61, x_60, x_8, x_9, x_10, x_11, x_19); +return x_62; } } -static lean_object* _init_l_Lean_PersistentHashMap_insert___at_Lean_Meta_synthInstance_x3f___spec__7___closed__4() { -_start: +else { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_instHashableLocalInstance___boxed), 1, 0); -return x_1; +lean_object* x_63; double x_64; lean_object* x_65; +x_63 = lean_box(0); +x_64 = lean_unbox_float(x_21); +lean_dec(x_21); +x_65 = l_Lean_withTraceNode___at_Lean_Meta_synthInstance_x3f___spec__13___lambda__2(x_2, x_20, x_14, x_3, x_4, x_5, x_22, x_64, x_63, x_8, x_9, x_10, x_11, x_19); +return x_65; } } -static lean_object* _init_l_Lean_PersistentHashMap_insert___at_Lean_Meta_synthInstance_x3f___spec__7___closed__5() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_synthInstance_x3f___spec__7___closed__4; -x_2 = lean_alloc_closure((void*)(l_instHashableArray___rarg___boxed), 2, 1); -lean_closure_set(x_2, 0, x_1); -return x_2; } +else +{ +uint8_t x_73; +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +x_73 = !lean_is_exclusive(x_17); +if (x_73 == 0) +{ +return x_17; } -static lean_object* _init_l_Lean_PersistentHashMap_insert___at_Lean_Meta_synthInstance_x3f___spec__7___closed__6() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_synthInstance_x3f___spec__7___closed__5; -x_2 = l_Lean_Expr_instHashableExpr; -x_3 = lean_alloc_closure((void*)(l_instHashableProd___rarg___boxed), 3, 2); -lean_closure_set(x_3, 0, x_1); -lean_closure_set(x_3, 1, x_2); -return x_3; +lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_74 = lean_ctor_get(x_17, 0); +x_75 = lean_ctor_get(x_17, 1); +lean_inc(x_75); +lean_inc(x_74); +lean_dec(x_17); +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_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_synthInstance_x3f___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_synthInstance_x3f___spec__13(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) { _start: { -uint8_t x_4; -x_4 = !lean_is_exclusive(x_1); -if (x_4 == 0) -{ -lean_object* x_5; lean_object* x_6; size_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint64_t x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; uint64_t x_16; -x_5 = lean_ctor_get(x_1, 0); -x_6 = lean_ctor_get(x_1, 1); -x_7 = 1; -x_8 = lean_unsigned_to_nat(1u); -x_9 = lean_nat_add(x_6, x_8); -lean_dec(x_6); -x_10 = lean_ctor_get(x_2, 0); +lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_10 = lean_ctor_get(x_7, 2); lean_inc(x_10); -x_11 = lean_ctor_get(x_2, 1); -lean_inc(x_11); -x_12 = 7; -x_13 = lean_array_get_size(x_10); -x_14 = lean_unsigned_to_nat(0u); -x_15 = lean_nat_dec_lt(x_14, x_13); -x_16 = l_Lean_Expr_hash(x_11); +lean_inc(x_1); +x_11 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___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_unbox(x_12); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); lean_dec(x_11); -if (x_15 == 0) +x_15 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__4___closed__1; +x_16 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_10, x_15); +if (x_16 == 0) { -uint64_t x_17; size_t x_18; lean_object* x_19; -lean_dec(x_13); +lean_object* x_17; +lean_dec(x_12); lean_dec(x_10); -x_17 = lean_uint64_mix_hash(x_12, x_16); -x_18 = lean_uint64_to_usize(x_17); -x_19 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_synthInstance_x3f___spec__8(x_5, x_18, x_7, x_2, x_3); -lean_ctor_set(x_1, 1, x_9); -lean_ctor_set(x_1, 0, x_19); -return x_1; -} -else -{ -uint8_t x_20; -x_20 = lean_nat_dec_le(x_13, x_13); -if (x_20 == 0) +lean_dec(x_2); +lean_dec(x_1); +x_17 = lean_apply_5(x_3, x_5, x_6, x_7, x_8, x_14); +if (lean_obj_tag(x_17) == 0) { -uint64_t x_21; size_t x_22; lean_object* x_23; -lean_dec(x_13); -lean_dec(x_10); -x_21 = lean_uint64_mix_hash(x_12, x_16); -x_22 = lean_uint64_to_usize(x_21); -x_23 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_synthInstance_x3f___spec__8(x_5, x_22, x_7, x_2, x_3); -lean_ctor_set(x_1, 1, x_9); -lean_ctor_set(x_1, 0, x_23); -return x_1; -} -else +uint8_t x_18; +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) { -size_t x_24; size_t x_25; uint64_t x_26; uint64_t x_27; size_t x_28; lean_object* x_29; -x_24 = 0; -x_25 = lean_usize_of_nat(x_13); -lean_dec(x_13); -x_26 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_synthInstance_x3f___spec__6(x_10, x_24, x_25, x_12); -lean_dec(x_10); -x_27 = lean_uint64_mix_hash(x_26, x_16); -x_28 = lean_uint64_to_usize(x_27); -x_29 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_synthInstance_x3f___spec__8(x_5, x_28, x_7, x_2, x_3); -lean_ctor_set(x_1, 1, x_9); -lean_ctor_set(x_1, 0, x_29); -return x_1; -} -} +return x_17; } else { -lean_object* x_30; lean_object* x_31; size_t x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint64_t x_37; lean_object* x_38; lean_object* x_39; uint8_t x_40; uint64_t x_41; -x_30 = lean_ctor_get(x_1, 0); -x_31 = lean_ctor_get(x_1, 1); -lean_inc(x_31); -lean_inc(x_30); -lean_dec(x_1); -x_32 = 1; -x_33 = lean_unsigned_to_nat(1u); -x_34 = lean_nat_add(x_31, x_33); -lean_dec(x_31); -x_35 = lean_ctor_get(x_2, 0); -lean_inc(x_35); -x_36 = lean_ctor_get(x_2, 1); -lean_inc(x_36); -x_37 = 7; -x_38 = lean_array_get_size(x_35); -x_39 = lean_unsigned_to_nat(0u); -x_40 = lean_nat_dec_lt(x_39, x_38); -x_41 = l_Lean_Expr_hash(x_36); -lean_dec(x_36); -if (x_40 == 0) -{ -uint64_t x_42; size_t x_43; lean_object* x_44; lean_object* x_45; -lean_dec(x_38); -lean_dec(x_35); -x_42 = lean_uint64_mix_hash(x_37, x_41); -x_43 = lean_uint64_to_usize(x_42); -x_44 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_synthInstance_x3f___spec__8(x_30, x_43, x_32, x_2, x_3); -x_45 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_45, 0, x_44); -lean_ctor_set(x_45, 1, x_34); -return x_45; +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_17, 0); +x_20 = lean_ctor_get(x_17, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_17); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); +return x_21; +} } else { -uint8_t x_46; -x_46 = lean_nat_dec_le(x_38, x_38); -if (x_46 == 0) +uint8_t x_22; +x_22 = !lean_is_exclusive(x_17); +if (x_22 == 0) { -uint64_t x_47; size_t x_48; lean_object* x_49; lean_object* x_50; -lean_dec(x_38); -lean_dec(x_35); -x_47 = lean_uint64_mix_hash(x_37, x_41); -x_48 = lean_uint64_to_usize(x_47); -x_49 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_synthInstance_x3f___spec__8(x_30, x_48, x_32, x_2, x_3); -x_50 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_50, 0, x_49); -lean_ctor_set(x_50, 1, x_34); -return x_50; +return x_17; } else { -size_t x_51; size_t x_52; uint64_t x_53; uint64_t x_54; size_t x_55; lean_object* x_56; lean_object* x_57; -x_51 = 0; -x_52 = lean_usize_of_nat(x_38); -lean_dec(x_38); -x_53 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_synthInstance_x3f___spec__6(x_35, x_51, x_52, x_37); -lean_dec(x_35); -x_54 = lean_uint64_mix_hash(x_53, x_41); -x_55 = lean_uint64_to_usize(x_54); -x_56 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_synthInstance_x3f___spec__8(x_30, x_55, x_32, x_2, x_3); -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; +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_17, 0); +x_24 = lean_ctor_get(x_17, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_17); +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_26; uint8_t x_27; lean_object* x_28; +x_26 = lean_box(0); +x_27 = lean_unbox(x_12); +lean_dec(x_12); +x_28 = l_Lean_withTraceNode___at_Lean_Meta_synthInstance_x3f___spec__13___lambda__3(x_3, x_2, x_1, x_4, x_10, x_27, x_26, x_5, x_6, x_7, x_8, x_14); +return x_28; +} +} +else +{ +lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; +x_29 = lean_ctor_get(x_11, 1); +lean_inc(x_29); +lean_dec(x_11); +x_30 = lean_box(0); +x_31 = lean_unbox(x_12); +lean_dec(x_12); +x_32 = l_Lean_withTraceNode___at_Lean_Meta_synthInstance_x3f___spec__13___lambda__3(x_3, x_2, x_1, x_4, x_10, x_31, x_30, x_5, x_6, x_7, x_8, x_29); +return x_32; } } } -LEAN_EXPORT lean_object* l_Lean_withSeconds___at_Lean_Meta_synthInstance_x3f___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_EXPORT lean_object* l_Lean_profileitM___at_Lean_Meta_synthInstance_x3f___spec__16___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) { _start: { -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_7 = lean_io_mono_nanos_now(x_6); -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 = lean_apply_5(x_1, x_2, x_3, x_4, x_5, x_9); -if (lean_obj_tag(x_10) == 0) +lean_object* x_10; lean_object* x_11; +x_10 = lean_apply_4(x_3, x_5, x_6, x_7, x_8); +x_11 = l_Lean_profileitIOUnsafe___rarg(x_1, x_2, x_10, x_4, x_9); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Meta_synthInstance_x3f___spec__16(lean_object* x_1) { +_start: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -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_io_mono_nanos_now(x_12); -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_profileitM___at_Lean_Meta_synthInstance_x3f___spec__16___rarg___boxed), 9, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_synthInstance_x3f___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_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; double x_19; double x_20; double x_21; lean_object* x_22; lean_object* x_23; -x_15 = lean_ctor_get(x_13, 0); -x_16 = lean_nat_sub(x_15, x_8); -lean_dec(x_8); -lean_dec(x_15); -x_17 = 0; -x_18 = lean_unsigned_to_nat(0u); -x_19 = l_Float_ofScientific(x_16, x_17, x_18); -lean_dec(x_16); -x_20 = l_Lean_withSeconds___at_Lean_Meta_SynthInstance_newSubgoal___spec__14___closed__1; -x_21 = lean_float_div(x_19, x_20); -x_22 = lean_box_float(x_21); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_11); -lean_ctor_set(x_23, 1, x_22); -lean_ctor_set(x_13, 0, x_23); -return x_13; +lean_object* x_8; uint8_t x_9; +x_8 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_1, x_3, x_4, x_5, x_6, x_7); +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; +x_10 = lean_ctor_get(x_8, 0); +x_11 = l_Lean_exceptOptionEmoji___rarg(x_2); +x_12 = l_Lean_stringToMessageData(x_11); +lean_dec(x_11); +x_13 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; +x_14 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_12); +x_15 = l_Lean_Meta_SynthInstance_tryResolve___lambda__1___closed__2; +x_16 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_16, 0, x_14); +lean_ctor_set(x_16, 1, x_15); +x_17 = l_Lean_MessageData_ofExpr(x_10); +x_18 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_17); +x_19 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_13); +lean_ctor_set(x_8, 0, x_19); +return x_8; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; lean_object* x_28; double x_29; double x_30; double x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_24 = lean_ctor_get(x_13, 0); -x_25 = lean_ctor_get(x_13, 1); -lean_inc(x_25); -lean_inc(x_24); -lean_dec(x_13); -x_26 = lean_nat_sub(x_24, x_8); +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_8, 0); +x_21 = lean_ctor_get(x_8, 1); +lean_inc(x_21); +lean_inc(x_20); lean_dec(x_8); -lean_dec(x_24); -x_27 = 0; -x_28 = lean_unsigned_to_nat(0u); -x_29 = l_Float_ofScientific(x_26, x_27, x_28); -lean_dec(x_26); -x_30 = l_Lean_withSeconds___at_Lean_Meta_SynthInstance_newSubgoal___spec__14___closed__1; -x_31 = lean_float_div(x_29, x_30); -x_32 = lean_box_float(x_31); -x_33 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_33, 0, x_11); -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_25); -return x_34; +x_22 = l_Lean_exceptOptionEmoji___rarg(x_2); +x_23 = l_Lean_stringToMessageData(x_22); +lean_dec(x_22); +x_24 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; +x_25 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_23); +x_26 = l_Lean_Meta_SynthInstance_tryResolve___lambda__1___closed__2; +x_27 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = l_Lean_MessageData_ofExpr(x_20); +x_29 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +x_30 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_24); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_21); +return x_31; } } -else +} +LEAN_EXPORT lean_object* l_Lean_Meta_synthInstance_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) { +_start: { -uint8_t x_35; +lean_object* x_8; +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_8 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessOutParam(x_1, x_3, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; +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_dec(x_8); -x_35 = !lean_is_exclusive(x_10); -if (x_35 == 0) +x_11 = l_Lean_Meta_SynthInstance_main(x_9, x_2, x_3, x_4, x_5, x_6, x_10); +return x_11; +} +else { -return x_10; +uint8_t x_12; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_12 = !lean_is_exclusive(x_8); +if (x_12 == 0) +{ +return x_8; } else { -lean_object* x_36; lean_object* x_37; lean_object* x_38; -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 = 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; +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_ctor_get(x_8, 0); +x_14 = lean_ctor_get(x_8, 1); +lean_inc(x_14); +lean_inc(x_13); +lean_dec(x_8); +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_Lean_withOptProfile___at_Lean_Meta_synthInstance_x3f___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_EXPORT lean_object* l_Lean_Meta_synthInstance_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) { _start: { -lean_object* x_7; uint8_t x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_7 = l_Lean_withOptProfile___at_Lean_Meta_SynthInstance_newSubgoal___spec__12___closed__2; -x_8 = 0; -x_9 = l_Lean_getBoolOption___at_Lean_Meta_processPostponed___spec__4(x_7, x_8, x_2, x_3, x_4, x_5, x_6); -x_10 = lean_ctor_get(x_9, 0); +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_8 = lean_st_ref_take(x_4, x_7); +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_9, 1); lean_inc(x_10); -x_11 = lean_unbox(x_10); -lean_dec(x_10); -if (x_11 == 0) -{ -lean_object* x_12; lean_object* x_13; -x_12 = lean_ctor_get(x_9, 1); -lean_inc(x_12); -lean_dec(x_9); -x_13 = lean_apply_5(x_1, x_2, x_3, x_4, x_5, x_12); -if (lean_obj_tag(x_13) == 0) +x_11 = lean_ctor_get(x_8, 1); +lean_inc(x_11); +lean_dec(x_8); +x_12 = !lean_is_exclusive(x_9); +if (x_12 == 0) { -uint8_t x_14; -x_14 = !lean_is_exclusive(x_13); +lean_object* x_13; uint8_t x_14; +x_13 = lean_ctor_get(x_9, 1); +lean_dec(x_13); +x_14 = !lean_is_exclusive(x_10); if (x_14 == 0) { -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_13, 0); -x_16 = lean_box(0); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_15); -lean_ctor_set(x_17, 1, x_16); -lean_ctor_set(x_13, 0, x_17); -return x_13; +lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_15 = lean_ctor_get(x_10, 2); +lean_inc(x_2); +x_16 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_synthInstance_x3f___spec__7(x_15, x_1, x_2); +lean_ctor_set(x_10, 2, x_16); +x_17 = lean_st_ref_set(x_4, x_9, x_11); +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) +{ +lean_object* x_19; +x_19 = lean_ctor_get(x_17, 0); +lean_dec(x_19); +lean_ctor_set(x_17, 0, x_2); +return x_17; } else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_18 = lean_ctor_get(x_13, 0); -x_19 = lean_ctor_get(x_13, 1); -lean_inc(x_19); -lean_inc(x_18); -lean_dec(x_13); -x_20 = lean_box(0); +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_17, 1); +lean_inc(x_20); +lean_dec(x_17); x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_18); +lean_ctor_set(x_21, 0, x_2); lean_ctor_set(x_21, 1, 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_19); -return x_22; -} +return x_21; } -else -{ -uint8_t x_23; -x_23 = !lean_is_exclusive(x_13); -if (x_23 == 0) -{ -return x_13; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_13, 0); -x_25 = lean_ctor_get(x_13, 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; +x_22 = lean_ctor_get(x_10, 0); +x_23 = lean_ctor_get(x_10, 1); +x_24 = lean_ctor_get(x_10, 3); +x_25 = lean_ctor_get(x_10, 4); +x_26 = lean_ctor_get(x_10, 5); +x_27 = lean_ctor_get(x_10, 2); +lean_inc(x_26); lean_inc(x_25); lean_inc(x_24); -lean_dec(x_13); -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; +lean_inc(x_27); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_10); +lean_inc(x_2); +x_28 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_synthInstance_x3f___spec__7(x_27, x_1, x_2); +x_29 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_29, 0, x_22); +lean_ctor_set(x_29, 1, x_23); +lean_ctor_set(x_29, 2, x_28); +lean_ctor_set(x_29, 3, x_24); +lean_ctor_set(x_29, 4, x_25); +lean_ctor_set(x_29, 5, x_26); +lean_ctor_set(x_9, 1, x_29); +x_30 = lean_st_ref_set(x_4, x_9, x_11); +x_31 = lean_ctor_get(x_30, 1); +lean_inc(x_31); +if (lean_is_exclusive(x_30)) { + lean_ctor_release(x_30, 0); + lean_ctor_release(x_30, 1); + x_32 = x_30; +} else { + lean_dec_ref(x_30); + x_32 = lean_box(0); } +if (lean_is_scalar(x_32)) { + x_33 = lean_alloc_ctor(0, 2, 0); +} else { + x_33 = x_32; } +lean_ctor_set(x_33, 0, x_2); +lean_ctor_set(x_33, 1, x_31); +return x_33; } -else -{ -lean_object* x_27; lean_object* x_28; -x_27 = lean_ctor_get(x_9, 1); -lean_inc(x_27); -lean_dec(x_9); -x_28 = l_Lean_withSeconds___at_Lean_Meta_synthInstance_x3f___spec__15(x_1, x_2, x_3, x_4, x_5, x_27); -if (lean_obj_tag(x_28) == 0) -{ -uint8_t x_29; -x_29 = !lean_is_exclusive(x_28); -if (x_29 == 0) -{ -lean_object* x_30; uint8_t x_31; -x_30 = lean_ctor_get(x_28, 0); -x_31 = !lean_is_exclusive(x_30); -if (x_31 == 0) -{ -lean_object* x_32; lean_object* x_33; -x_32 = lean_ctor_get(x_30, 1); -x_33 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_33, 0, x_32); -lean_ctor_set(x_30, 1, x_33); -return x_28; } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_34 = lean_ctor_get(x_30, 0); -x_35 = lean_ctor_get(x_30, 1); +lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; 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_34 = lean_ctor_get(x_9, 0); +x_35 = lean_ctor_get(x_9, 2); +x_36 = lean_ctor_get(x_9, 3); +lean_inc(x_36); lean_inc(x_35); lean_inc(x_34); -lean_dec(x_30); -x_36 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_36, 0, x_35); -x_37 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_37, 0, x_34); -lean_ctor_set(x_37, 1, x_36); -lean_ctor_set(x_28, 0, x_37); -return x_28; -} -} -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; -x_38 = lean_ctor_get(x_28, 0); -x_39 = lean_ctor_get(x_28, 1); -lean_inc(x_39); +lean_dec(x_9); +x_37 = lean_ctor_get(x_10, 0); +lean_inc(x_37); +x_38 = lean_ctor_get(x_10, 1); lean_inc(x_38); -lean_dec(x_28); -x_40 = lean_ctor_get(x_38, 0); +x_39 = lean_ctor_get(x_10, 3); +lean_inc(x_39); +x_40 = lean_ctor_get(x_10, 4); lean_inc(x_40); -x_41 = lean_ctor_get(x_38, 1); +x_41 = lean_ctor_get(x_10, 5); lean_inc(x_41); -if (lean_is_exclusive(x_38)) { - lean_ctor_release(x_38, 0); - lean_ctor_release(x_38, 1); - x_42 = x_38; +x_42 = lean_ctor_get(x_10, 2); +lean_inc(x_42); +if (lean_is_exclusive(x_10)) { + lean_ctor_release(x_10, 0); + lean_ctor_release(x_10, 1); + lean_ctor_release(x_10, 2); + lean_ctor_release(x_10, 3); + lean_ctor_release(x_10, 4); + lean_ctor_release(x_10, 5); + x_43 = x_10; } else { - lean_dec_ref(x_38); - x_42 = lean_box(0); + lean_dec_ref(x_10); + x_43 = lean_box(0); } -x_43 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_43, 0, x_41); -if (lean_is_scalar(x_42)) { - x_44 = lean_alloc_ctor(0, 2, 0); +lean_inc(x_2); +x_44 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_synthInstance_x3f___spec__7(x_42, x_1, x_2); +if (lean_is_scalar(x_43)) { + x_45 = lean_alloc_ctor(0, 6, 0); } else { - x_44 = x_42; -} -lean_ctor_set(x_44, 0, x_40); -lean_ctor_set(x_44, 1, 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_39); -return x_45; -} -} -else -{ -uint8_t x_46; -x_46 = !lean_is_exclusive(x_28); -if (x_46 == 0) -{ -return x_28; + x_45 = x_43; } -else -{ -lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_47 = lean_ctor_get(x_28, 0); -x_48 = lean_ctor_get(x_28, 1); +lean_ctor_set(x_45, 0, x_37); +lean_ctor_set(x_45, 1, x_38); +lean_ctor_set(x_45, 2, x_44); +lean_ctor_set(x_45, 3, x_39); +lean_ctor_set(x_45, 4, x_40); +lean_ctor_set(x_45, 5, x_41); +x_46 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_46, 0, x_34); +lean_ctor_set(x_46, 1, x_45); +lean_ctor_set(x_46, 2, x_35); +lean_ctor_set(x_46, 3, x_36); +x_47 = lean_st_ref_set(x_4, x_46, x_11); +x_48 = lean_ctor_get(x_47, 1); lean_inc(x_48); -lean_inc(x_47); -lean_dec(x_28); -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; +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); } +if (lean_is_scalar(x_49)) { + x_50 = lean_alloc_ctor(0, 2, 0); +} else { + x_50 = x_49; } +lean_ctor_set(x_50, 0, x_2); +lean_ctor_set(x_50, 1, x_48); +return x_50; } } } -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_synthInstance_x3f___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_EXPORT lean_object* l_Lean_Meta_synthInstance_x3f___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: { -if (lean_obj_tag(x_1) == 0) -{ -lean_object* x_7; lean_object* x_8; -x_7 = lean_ctor_get(x_1, 0); +lean_object* x_10; +lean_inc(x_8); lean_inc(x_7); -x_8 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_8, 0, x_7); -lean_ctor_set(x_8, 1, x_6); -return x_8; -} -else -{ -lean_object* x_9; lean_object* x_10; -x_9 = lean_ctor_get(x_1, 0); -lean_inc(x_9); -x_10 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_10, 0, x_9); -lean_ctor_set(x_10, 1, x_6); -return x_10; -} -} -} -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_synthInstance_x3f___spec__13___lambda__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) { -_start: -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; -lean_inc(x_10); -x_13 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_processPostponed___spec__6(x_1, x_2, x_3, x_6, x_4, x_8, x_9, x_10, x_11, x_12); -x_14 = lean_ctor_get(x_13, 1); -lean_inc(x_14); -lean_dec(x_13); -x_15 = l_MonadExcept_ofExcept___at_Lean_Meta_synthInstance_x3f___spec__16(x_5, x_8, x_9, x_10, x_11, x_14); -lean_dec(x_10); -return x_15; -} -} -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_synthInstance_x3f___spec__13(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) { -_start: +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_2); +x_10 = l_Lean_Meta_synthInstance_x3f_assignOutParams(x_1, x_2, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_10) == 0) { -lean_object* x_10; lean_object* x_11; uint8_t x_12; -lean_inc(x_1); -x_10 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_1, x_5, x_6, x_7, x_8, x_9); +lean_object* x_11; uint8_t x_12; x_11 = lean_ctor_get(x_10, 0); lean_inc(x_11); x_12 = lean_unbox(x_11); lean_dec(x_11); if (x_12 == 0) { -lean_object* x_13; lean_object* x_14; +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_dec(x_2); -lean_dec(x_1); x_13 = lean_ctor_get(x_10, 1); lean_inc(x_13); lean_dec(x_10); -x_14 = lean_apply_5(x_3, x_5, x_6, x_7, x_8, x_13); -return x_14; +x_14 = lean_box(0); +x_15 = lean_apply_6(x_3, x_14, x_5, x_6, x_7, x_8, x_13); +return x_15; } -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; -x_15 = lean_ctor_get(x_10, 1); -lean_inc(x_15); -lean_dec(x_10); -x_16 = lean_ctor_get(x_7, 5); +else +{ +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_10, 1); lean_inc(x_16); -x_17 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_processPostponed___spec__2___rarg(x_8, x_15); +lean_dec(x_10); +x_17 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_2, x_5, x_6, x_7, x_8, 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_closure((void*)(l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_tryResolve___spec__2___lambda__1), 6, 1); -lean_closure_set(x_20, 0, x_3); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_21 = l_Lean_withOptProfile___at_Lean_Meta_synthInstance_x3f___spec__14(x_20, x_5, x_6, x_7, x_8, 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; -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_22, 0); -lean_inc(x_24); -x_25 = lean_ctor_get(x_22, 1); -lean_inc(x_25); -lean_dec(x_22); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_24); -x_26 = lean_apply_6(x_2, x_24, x_5, x_6, x_7, x_8, x_23); -if (lean_obj_tag(x_26) == 0) -{ -if (lean_obj_tag(x_25) == 0) -{ -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; -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_box(0); -x_30 = l_Lean_withTraceNode___at_Lean_Meta_synthInstance_x3f___spec__13___lambda__1(x_18, x_1, x_16, x_4, x_24, x_27, x_29, x_5, x_6, x_7, x_8, x_28); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_24); -return x_30; -} -else +lean_inc(x_18); +x_20 = l_Lean_Meta_check(x_18, x_5, x_6, x_7, x_8, x_19); +if (lean_obj_tag(x_20) == 0) { -lean_object* x_31; lean_object* x_32; lean_object* x_33; double x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; 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_31 = lean_ctor_get(x_26, 0); -lean_inc(x_31); -x_32 = lean_ctor_get(x_26, 1); -lean_inc(x_32); -lean_dec(x_26); -x_33 = lean_ctor_get(x_25, 0); -lean_inc(x_33); -lean_dec(x_25); -x_34 = lean_unbox_float(x_33); -lean_dec(x_33); -x_35 = lean_float_to_string(x_34); -x_36 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_36, 0, x_35); -x_37 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_37, 0, x_36); -x_38 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___closed__2; -x_39 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_39, 1, x_37); -x_40 = l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___closed__4; -x_41 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_41, 0, x_39); -lean_ctor_set(x_41, 1, x_40); -x_42 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_31); -x_43 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; -x_44 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_44, 0, x_42); -lean_ctor_set(x_44, 1, x_43); -x_45 = lean_box(0); -x_46 = l_Lean_withTraceNode___at_Lean_Meta_synthInstance_x3f___spec__13___lambda__1(x_18, x_1, x_16, x_4, x_24, x_44, x_45, x_5, x_6, x_7, x_8, x_32); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_24); -return x_46; -} +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 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_22, 0, x_18); +x_23 = lean_apply_6(x_3, x_22, x_5, x_6, x_7, x_8, x_21); +return x_23; } else { -uint8_t x_47; -lean_dec(x_25); -lean_dec(x_24); +uint8_t x_24; lean_dec(x_18); -lean_dec(x_16); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_1); -x_47 = !lean_is_exclusive(x_26); -if (x_47 == 0) +lean_dec(x_3); +x_24 = !lean_is_exclusive(x_20); +if (x_24 == 0) { -return x_26; +return x_20; } else { -lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_48 = lean_ctor_get(x_26, 0); -x_49 = lean_ctor_get(x_26, 1); -lean_inc(x_49); -lean_inc(x_48); -lean_dec(x_26); -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; +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_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; +} } } } else { -uint8_t x_51; -lean_dec(x_18); -lean_dec(x_16); +uint8_t x_28; lean_dec(x_8); 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_51 = !lean_is_exclusive(x_21); -if (x_51 == 0) +x_28 = !lean_is_exclusive(x_10); +if (x_28 == 0) { -return x_21; +return x_10; } else { -lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_52 = lean_ctor_get(x_21, 0); -x_53 = lean_ctor_get(x_21, 1); -lean_inc(x_53); -lean_inc(x_52); -lean_dec(x_21); -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_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_10, 0); +x_30 = lean_ctor_get(x_10, 1); +lean_inc(x_30); +lean_inc(x_29); +lean_dec(x_10); +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_profileitM___at_Lean_Meta_synthInstance_x3f___spec__17___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_EXPORT lean_object* l_Lean_Meta_synthInstance_x3f___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) { _start: { -lean_object* x_10; lean_object* x_11; -x_10 = lean_apply_4(x_3, x_5, x_6, x_7, x_8); -x_11 = l_Lean_profileitIOUnsafe___rarg(x_1, x_2, x_10, x_4, x_9); -return x_11; -} -} -LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Meta_synthInstance_x3f___spec__17(lean_object* x_1) { -_start: +if (lean_obj_tag(x_1) == 0) { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_profileitM___at_Lean_Meta_synthInstance_x3f___spec__17___rarg___boxed), 9, 0); -return x_2; -} +lean_object* x_9; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_1); +lean_ctor_set(x_9, 1, x_8); +return x_9; } -LEAN_EXPORT lean_object* l_Lean_Meta_synthInstance_x3f___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: +else { -lean_object* x_8; uint8_t x_9; -x_8 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_1, x_3, x_4, x_5, x_6, x_7); -x_9 = !lean_is_exclusive(x_8); -if (x_9 == 0) +lean_object* x_10; lean_object* x_11; +x_10 = lean_ctor_get(x_1, 0); +lean_inc(x_10); +x_11 = l_Lean_Meta_synthInstance_x3f_assignOutParams(x_2, x_10, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_11) == 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; -x_10 = lean_ctor_get(x_8, 0); -x_11 = l_Lean_exceptOptionEmoji___rarg(x_2); -x_12 = l_Lean_stringToMessageData(x_11); +lean_object* x_12; uint8_t x_13; +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_unbox(x_12); +lean_dec(x_12); +if (x_13 == 0) +{ +uint8_t x_14; +lean_dec(x_1); +x_14 = !lean_is_exclusive(x_11); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_11, 0); +lean_dec(x_15); +x_16 = lean_box(0); +lean_ctor_set(x_11, 0, x_16); +return x_11; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_11, 1); +lean_inc(x_17); lean_dec(x_11); -x_13 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; -x_14 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_14, 0, x_13); -lean_ctor_set(x_14, 1, x_12); -x_15 = l_Lean_Meta_SynthInstance_tryResolve___lambda__1___closed__2; -x_16 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_16, 0, x_14); -lean_ctor_set(x_16, 1, x_15); -x_17 = l_Lean_MessageData_ofExpr(x_10); -x_18 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_18, 0, x_16); -lean_ctor_set(x_18, 1, x_17); -x_19 = lean_alloc_ctor(7, 2, 0); +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_19, 0, x_18); -lean_ctor_set(x_19, 1, x_13); -lean_ctor_set(x_8, 0, x_19); -return x_8; +lean_ctor_set(x_19, 1, x_17); +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; -x_20 = lean_ctor_get(x_8, 0); -x_21 = lean_ctor_get(x_8, 1); -lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_8); -x_22 = l_Lean_exceptOptionEmoji___rarg(x_2); -x_23 = l_Lean_stringToMessageData(x_22); -lean_dec(x_22); -x_24 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; -x_25 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_25, 0, x_24); -lean_ctor_set(x_25, 1, x_23); -x_26 = l_Lean_Meta_SynthInstance_tryResolve___lambda__1___closed__2; -x_27 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_27, 0, x_25); -lean_ctor_set(x_27, 1, x_26); -x_28 = l_Lean_MessageData_ofExpr(x_20); -x_29 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_29, 0, x_27); -lean_ctor_set(x_29, 1, x_28); -x_30 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_30, 0, x_29); -lean_ctor_set(x_30, 1, x_24); -x_31 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_21); -return x_31; +uint8_t x_20; +x_20 = !lean_is_exclusive(x_11); +if (x_20 == 0) +{ +lean_object* x_21; +x_21 = lean_ctor_get(x_11, 0); +lean_dec(x_21); +lean_ctor_set(x_11, 0, x_1); +return x_11; } +else +{ +lean_object* x_22; lean_object* x_23; +x_22 = lean_ctor_get(x_11, 1); +lean_inc(x_22); +lean_dec(x_11); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_1); +lean_ctor_set(x_23, 1, x_22); +return x_23; } } -LEAN_EXPORT lean_object* l_Lean_Meta_synthInstance_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) { -_start: -{ -lean_object* x_8; -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -x_8 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessOutParam(x_1, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_8) == 0) -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; -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_dec(x_8); -x_11 = l_Lean_Meta_SynthInstance_main(x_9, x_2, x_3, x_4, x_5, x_6, x_10); -return x_11; } else { -uint8_t x_12; -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_12 = !lean_is_exclusive(x_8); -if (x_12 == 0) +uint8_t x_24; +lean_dec(x_1); +x_24 = !lean_is_exclusive(x_11); +if (x_24 == 0) { -return x_8; +return x_11; } else { -lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_13 = lean_ctor_get(x_8, 0); -x_14 = lean_ctor_get(x_8, 1); -lean_inc(x_14); -lean_inc(x_13); -lean_dec(x_8); -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_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_11, 0); +x_26 = lean_ctor_get(x_11, 1); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_11); +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_EXPORT lean_object* l_Lean_Meta_synthInstance_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) { +} +static lean_object* _init_l_Lean_Meta_synthInstance_x3f___lambda__6___closed__1() { _start: { -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_8 = lean_st_ref_take(x_4, x_7); -x_9 = lean_ctor_get(x_8, 0); -lean_inc(x_9); -x_10 = lean_ctor_get(x_9, 1); -lean_inc(x_10); -x_11 = lean_ctor_get(x_8, 1); -lean_inc(x_11); -lean_dec(x_8); -x_12 = !lean_is_exclusive(x_9); -if (x_12 == 0) -{ -lean_object* x_13; uint8_t x_14; -x_13 = lean_ctor_get(x_9, 1); -lean_dec(x_13); -x_14 = !lean_is_exclusive(x_10); -if (x_14 == 0) -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_15 = lean_ctor_get(x_10, 2); -lean_inc(x_2); -x_16 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_synthInstance_x3f___spec__7(x_15, x_1, x_2); -lean_ctor_set(x_10, 2, x_16); -x_17 = lean_st_ref_set(x_4, x_9, x_11); -x_18 = !lean_is_exclusive(x_17); -if (x_18 == 0) +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("result ", 7); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_synthInstance_x3f___lambda__6___closed__2() { +_start: { -lean_object* x_19; -x_19 = lean_ctor_get(x_17, 0); -lean_dec(x_19); -lean_ctor_set(x_17, 0, x_2); -return x_17; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_synthInstance_x3f___lambda__6___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } -else +} +static lean_object* _init_l_Lean_Meta_synthInstance_x3f___lambda__6___closed__3() { +_start: { -lean_object* x_20; lean_object* x_21; -x_20 = lean_ctor_get(x_17, 1); -lean_inc(x_20); -lean_dec(x_17); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_2); -lean_ctor_set(x_21, 1, x_20); -return x_21; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(" (cached)", 9); +return x_1; } } -else +static lean_object* _init_l_Lean_Meta_synthInstance_x3f___lambda__6___closed__4() { +_start: { -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_22 = lean_ctor_get(x_10, 0); -x_23 = lean_ctor_get(x_10, 1); -x_24 = lean_ctor_get(x_10, 3); -x_25 = lean_ctor_get(x_10, 4); -x_26 = lean_ctor_get(x_10, 5); -x_27 = lean_ctor_get(x_10, 2); -lean_inc(x_26); -lean_inc(x_25); -lean_inc(x_24); -lean_inc(x_27); -lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_10); -lean_inc(x_2); -x_28 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_synthInstance_x3f___spec__7(x_27, x_1, x_2); -x_29 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_29, 0, x_22); -lean_ctor_set(x_29, 1, x_23); -lean_ctor_set(x_29, 2, x_28); -lean_ctor_set(x_29, 3, x_24); -lean_ctor_set(x_29, 4, x_25); -lean_ctor_set(x_29, 5, x_26); -lean_ctor_set(x_9, 1, x_29); -x_30 = lean_st_ref_set(x_4, x_9, x_11); -x_31 = lean_ctor_get(x_30, 1); -lean_inc(x_31); -if (lean_is_exclusive(x_30)) { - lean_ctor_release(x_30, 0); - lean_ctor_release(x_30, 1); - x_32 = x_30; -} else { - lean_dec_ref(x_30); - x_32 = lean_box(0); +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_synthInstance_x3f___lambda__6___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } -if (lean_is_scalar(x_32)) { - x_33 = lean_alloc_ctor(0, 2, 0); -} else { - x_33 = x_32; } -lean_ctor_set(x_33, 0, x_2); -lean_ctor_set(x_33, 1, x_31); -return x_33; +static lean_object* _init_l_Lean_Meta_synthInstance_x3f___lambda__6___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("", 15); +return x_1; } } -else +static lean_object* _init_l_Lean_Meta_synthInstance_x3f___lambda__6___closed__6() { +_start: { -lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; 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_34 = lean_ctor_get(x_9, 0); -x_35 = lean_ctor_get(x_9, 2); -x_36 = lean_ctor_get(x_9, 3); -lean_inc(x_36); -lean_inc(x_35); -lean_inc(x_34); -lean_dec(x_9); -x_37 = lean_ctor_get(x_10, 0); -lean_inc(x_37); -x_38 = lean_ctor_get(x_10, 1); -lean_inc(x_38); -x_39 = lean_ctor_get(x_10, 3); -lean_inc(x_39); -x_40 = lean_ctor_get(x_10, 4); -lean_inc(x_40); -x_41 = lean_ctor_get(x_10, 5); -lean_inc(x_41); -x_42 = lean_ctor_get(x_10, 2); -lean_inc(x_42); -if (lean_is_exclusive(x_10)) { - lean_ctor_release(x_10, 0); - lean_ctor_release(x_10, 1); - lean_ctor_release(x_10, 2); - lean_ctor_release(x_10, 3); - lean_ctor_release(x_10, 4); - lean_ctor_release(x_10, 5); - x_43 = x_10; -} else { - lean_dec_ref(x_10); - x_43 = lean_box(0); +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_synthInstance_x3f___lambda__6___closed__5; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; } -lean_inc(x_2); -x_44 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_synthInstance_x3f___spec__7(x_42, x_1, x_2); -if (lean_is_scalar(x_43)) { - x_45 = lean_alloc_ctor(0, 6, 0); -} else { - x_45 = x_43; } -lean_ctor_set(x_45, 0, x_37); -lean_ctor_set(x_45, 1, x_38); -lean_ctor_set(x_45, 2, x_44); -lean_ctor_set(x_45, 3, x_39); -lean_ctor_set(x_45, 4, x_40); -lean_ctor_set(x_45, 5, x_41); -x_46 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_46, 0, x_34); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_46, 2, x_35); -lean_ctor_set(x_46, 3, x_36); -x_47 = lean_st_ref_set(x_4, x_46, x_11); -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); +static lean_object* _init_l_Lean_Meta_synthInstance_x3f___lambda__6___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_synthInstance_x3f___lambda__6___closed__6; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; } -if (lean_is_scalar(x_49)) { - x_50 = lean_alloc_ctor(0, 2, 0); -} else { - x_50 = x_49; } -lean_ctor_set(x_50, 0, x_2); -lean_ctor_set(x_50, 1, x_48); -return x_50; +static lean_object* _init_l_Lean_Meta_synthInstance_x3f___lambda__6___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_synthInstance_x3f___lambda__6___closed__2; +x_2 = l_Lean_Meta_synthInstance_x3f___lambda__6___closed__7; +x_3 = lean_alloc_ctor(7, 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_Meta_synthInstance_x3f___lambda__6___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_synthInstance_x3f___lambda__6___closed__8; +x_2 = l_Lean_Meta_synthInstance_x3f___lambda__6___closed__4; +x_3 = lean_alloc_ctor(7, 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_synthInstance_x3f___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_EXPORT lean_object* l_Lean_Meta_synthInstance_x3f___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) { _start: { -lean_object* x_10; -lean_inc(x_8); +lean_object* x_9; uint8_t x_19; +x_19 = !lean_is_exclusive(x_4); +if (x_19 == 0) +{ +lean_object* x_20; uint8_t x_21; +x_20 = lean_ctor_get(x_4, 0); +x_21 = !lean_is_exclusive(x_20); +if (x_21 == 0) +{ +uint8_t x_22; uint8_t 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; +x_22 = 1; +x_23 = 0; +x_24 = 3; +lean_ctor_set_uint8(x_20, 0, x_22); +lean_ctor_set_uint8(x_20, 1, x_22); +lean_ctor_set_uint8(x_20, 3, x_23); +lean_ctor_set_uint8(x_20, 4, x_22); +lean_ctor_set_uint8(x_20, 5, x_24); +x_25 = l_Lean_Meta_getLocalInstances(x_4, x_5, x_6, x_7, x_8); +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_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_1, x_4, x_5, x_6, x_7, 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 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocess___closed__1; lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_2); -x_10 = l_Lean_Meta_synthInstance_x3f_assignOutParams(x_1, x_2, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_10) == 0) -{ -lean_object* x_11; uint8_t x_12; -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -x_12 = lean_unbox(x_11); -lean_dec(x_11); -if (x_12 == 0) +lean_inc(x_4); +x_32 = l_Lean_Meta_forallTelescopeReducing___at_Lean_Meta_getParamNames___spec__2___rarg(x_29, x_31, x_4, x_5, x_6, x_7, x_30); +if (lean_obj_tag(x_32) == 0) { -lean_object* x_13; lean_object* x_14; lean_object* x_15; -lean_dec(x_2); -x_13 = lean_ctor_get(x_10, 1); -lean_inc(x_13); -lean_dec(x_10); -x_14 = lean_box(0); -x_15 = lean_apply_6(x_3, x_14, x_5, x_6, x_7, x_8, x_13); -return x_15; -} -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; +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_st_ref_get(x_5, 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); +lean_dec(x_35); +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +lean_dec(x_36); +x_39 = lean_ctor_get(x_38, 2); +lean_inc(x_39); +lean_dec(x_38); +lean_inc(x_33); +x_40 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_40, 0, x_26); +lean_ctor_set(x_40, 1, x_33); +lean_inc(x_40); +x_41 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_synthInstance_x3f___spec__1(x_39, x_40); +if (lean_obj_tag(x_41) == 0) { -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_10, 1); -lean_inc(x_16); -lean_dec(x_10); -x_17 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_2, x_5, x_6, x_7, x_8, 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); -lean_inc(x_8); +lean_object* x_42; lean_object* x_43; +lean_inc(x_33); +x_42 = lean_alloc_closure((void*)(l_Lean_Meta_synthInstance_x3f___lambda__2), 7, 2); +lean_closure_set(x_42, 0, x_33); +lean_closure_set(x_42, 1, x_2); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_18); -x_20 = l_Lean_Meta_check(x_18, x_5, x_6, x_7, x_8, x_19); -if (lean_obj_tag(x_20) == 0) +lean_inc(x_4); +x_43 = l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Meta_WHNF_0__Lean_Meta_toCtorWhenK___spec__1___rarg(x_42, x_22, x_4, x_5, x_6, x_7, x_37); +if (lean_obj_tag(x_43) == 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 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_22, 0, x_18); -x_23 = lean_apply_6(x_3, x_22, x_5, x_6, x_7, x_8, x_21); -return x_23; -} -else +lean_object* x_44; +x_44 = lean_ctor_get(x_43, 0); +lean_inc(x_44); +if (lean_obj_tag(x_44) == 0) { -uint8_t x_24; -lean_dec(x_18); -lean_dec(x_8); +lean_object* x_45; lean_object* x_46; lean_object* x_47; +lean_dec(x_33); +lean_dec(x_3); +x_45 = lean_ctor_get(x_43, 1); +lean_inc(x_45); +lean_dec(x_43); +x_46 = lean_box(0); +x_47 = l_Lean_Meta_synthInstance_x3f___lambda__3(x_40, x_46, x_4, x_5, x_6, x_7, x_45); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_3); -x_24 = !lean_is_exclusive(x_20); -if (x_24 == 0) -{ -return x_20; +lean_dec(x_4); +x_9 = x_47; +goto block_18; } 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_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_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; uint8_t x_58; +x_48 = lean_ctor_get(x_43, 1); +lean_inc(x_48); +lean_dec(x_43); +x_49 = lean_ctor_get(x_44, 0); +lean_inc(x_49); +lean_dec(x_44); +x_50 = lean_alloc_closure((void*)(l_Lean_Meta_synthInstance_x3f___lambda__3___boxed), 7, 1); +lean_closure_set(x_50, 0, x_40); +lean_inc(x_4); +x_51 = l_Lean_Meta_openAbstractMVarsResult(x_49, x_4, x_5, x_6, x_7, x_48); +x_52 = lean_ctor_get(x_51, 0); +lean_inc(x_52); +x_53 = lean_ctor_get(x_52, 1); +lean_inc(x_53); +lean_dec(x_52); +x_54 = lean_ctor_get(x_51, 1); +lean_inc(x_54); +lean_dec(x_51); +x_55 = lean_ctor_get(x_53, 1); +lean_inc(x_55); +lean_dec(x_53); +lean_inc(x_3); +x_56 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_3, x_4, x_5, x_6, x_7, x_54); +x_57 = lean_ctor_get(x_56, 0); +lean_inc(x_57); +x_58 = lean_unbox(x_57); +lean_dec(x_57); +if (x_58 == 0) +{ +lean_object* x_59; lean_object* x_60; lean_object* x_61; +lean_dec(x_3); +x_59 = lean_ctor_get(x_56, 1); +lean_inc(x_59); +lean_dec(x_56); +x_60 = lean_box(0); +x_61 = l_Lean_Meta_synthInstance_x3f___lambda__4(x_33, x_55, x_50, x_60, x_4, x_5, x_6, x_7, x_59); +x_9 = x_61; +goto block_18; } +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_56, 1); +lean_inc(x_62); +lean_dec(x_56); +lean_inc(x_55); +x_63 = l_Lean_MessageData_ofExpr(x_55); +x_64 = l_Lean_Meta_synthInstance_x3f___lambda__6___closed__2; +x_65 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_65, 0, x_64); +lean_ctor_set(x_65, 1, x_63); +x_66 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; +x_67 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_67, 0, x_65); +lean_ctor_set(x_67, 1, x_66); +x_68 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_3, x_67, x_4, x_5, x_6, x_7, x_62); +x_69 = lean_ctor_get(x_68, 0); +lean_inc(x_69); +x_70 = lean_ctor_get(x_68, 1); +lean_inc(x_70); +lean_dec(x_68); +x_71 = l_Lean_Meta_synthInstance_x3f___lambda__4(x_33, x_55, x_50, x_69, x_4, x_5, x_6, x_7, x_70); +lean_dec(x_69); +x_9 = x_71; +goto block_18; } } } else { -uint8_t x_28; -lean_dec(x_8); +uint8_t x_72; +lean_dec(x_40); +lean_dec(x_33); +lean_dec(x_4); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -lean_dec(x_2); -x_28 = !lean_is_exclusive(x_10); -if (x_28 == 0) +x_72 = !lean_is_exclusive(x_43); +if (x_72 == 0) { -return x_10; +x_9 = x_43; +goto block_18; } else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = lean_ctor_get(x_10, 0); -x_30 = lean_ctor_get(x_10, 1); -lean_inc(x_30); -lean_inc(x_29); -lean_dec(x_10); -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_object* x_73; lean_object* x_74; lean_object* x_75; +x_73 = lean_ctor_get(x_43, 0); +x_74 = lean_ctor_get(x_43, 1); +lean_inc(x_74); +lean_inc(x_73); +lean_dec(x_43); +x_75 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_75, 0, x_73); +lean_ctor_set(x_75, 1, x_74); +x_9 = x_75; +goto block_18; } -LEAN_EXPORT lean_object* l_Lean_Meta_synthInstance_x3f___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) { -_start: -{ -if (lean_obj_tag(x_1) == 0) -{ -lean_object* x_9; -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_2); -x_9 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_9, 0, x_1); -lean_ctor_set(x_9, 1, x_8); -return x_9; } -else -{ -lean_object* x_10; lean_object* x_11; -x_10 = lean_ctor_get(x_1, 0); -lean_inc(x_10); -x_11 = l_Lean_Meta_synthInstance_x3f_assignOutParams(x_2, x_10, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_11) == 0) -{ -lean_object* x_12; uint8_t x_13; -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_unbox(x_12); -lean_dec(x_12); -if (x_13 == 0) -{ -uint8_t x_14; -lean_dec(x_1); -x_14 = !lean_is_exclusive(x_11); -if (x_14 == 0) -{ -lean_object* x_15; lean_object* x_16; -x_15 = lean_ctor_get(x_11, 0); -lean_dec(x_15); -x_16 = lean_box(0); -lean_ctor_set(x_11, 0, x_16); -return x_11; } else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_17 = lean_ctor_get(x_11, 1); -lean_inc(x_17); -lean_dec(x_11); -x_18 = lean_box(0); -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_18); -lean_ctor_set(x_19, 1, x_17); -return x_19; -} +lean_object* x_76; lean_object* x_77; lean_object* x_78; uint8_t x_79; +lean_dec(x_40); +lean_dec(x_2); +x_76 = lean_ctor_get(x_41, 0); +lean_inc(x_76); +lean_dec(x_41); +lean_inc(x_3); +x_77 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_3, x_4, x_5, x_6, x_7, x_37); +x_78 = lean_ctor_get(x_77, 0); +lean_inc(x_78); +x_79 = lean_unbox(x_78); +lean_dec(x_78); +if (x_79 == 0) +{ +lean_object* x_80; lean_object* x_81; lean_object* x_82; +lean_dec(x_3); +x_80 = lean_ctor_get(x_77, 1); +lean_inc(x_80); +lean_dec(x_77); +x_81 = lean_box(0); +x_82 = l_Lean_Meta_synthInstance_x3f___lambda__5(x_76, x_33, x_81, x_4, x_5, x_6, x_7, x_80); +x_9 = x_82; +goto block_18; } else { -uint8_t x_20; -x_20 = !lean_is_exclusive(x_11); -if (x_20 == 0) +if (lean_obj_tag(x_76) == 0) { -lean_object* x_21; -x_21 = lean_ctor_get(x_11, 0); -lean_dec(x_21); -lean_ctor_set(x_11, 0, x_1); -return x_11; +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_83 = lean_ctor_get(x_77, 1); +lean_inc(x_83); +lean_dec(x_77); +x_84 = l_Lean_Meta_synthInstance_x3f___lambda__6___closed__9; +x_85 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_3, x_84, x_4, x_5, x_6, x_7, x_83); +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 = l_Lean_Meta_synthInstance_x3f___lambda__5(x_76, x_33, x_86, x_4, x_5, x_6, x_7, x_87); +lean_dec(x_86); +x_9 = x_88; +goto block_18; } else { -lean_object* x_22; lean_object* x_23; -x_22 = lean_ctor_get(x_11, 1); -lean_inc(x_22); -lean_dec(x_11); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_1); -lean_ctor_set(x_23, 1, x_22); -return x_23; +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_89 = lean_ctor_get(x_77, 1); +lean_inc(x_89); +lean_dec(x_77); +x_90 = lean_ctor_get(x_76, 0); +lean_inc(x_90); +x_91 = l_Lean_MessageData_ofExpr(x_90); +x_92 = l_Lean_Meta_synthInstance_x3f___lambda__6___closed__2; +x_93 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_93, 0, x_92); +lean_ctor_set(x_93, 1, x_91); +x_94 = l_Lean_Meta_synthInstance_x3f___lambda__6___closed__4; +x_95 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_95, 0, x_93); +lean_ctor_set(x_95, 1, x_94); +x_96 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_3, x_95, x_4, x_5, x_6, x_7, x_89); +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_Meta_synthInstance_x3f___lambda__5(x_76, x_33, x_97, x_4, x_5, x_6, x_7, x_98); +lean_dec(x_97); +x_9 = x_99; +goto block_18; +} } } } else { -uint8_t x_24; -lean_dec(x_1); -x_24 = !lean_is_exclusive(x_11); -if (x_24 == 0) +uint8_t x_100; +lean_dec(x_26); +lean_dec(x_4); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +x_100 = !lean_is_exclusive(x_32); +if (x_100 == 0) { -return x_11; +return x_32; } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_25 = lean_ctor_get(x_11, 0); -x_26 = lean_ctor_get(x_11, 1); -lean_inc(x_26); -lean_inc(x_25); -lean_dec(x_11); -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_object* x_101; lean_object* x_102; lean_object* x_103; +x_101 = lean_ctor_get(x_32, 0); +x_102 = lean_ctor_get(x_32, 1); +lean_inc(x_102); +lean_inc(x_101); +lean_dec(x_32); +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; } -static lean_object* _init_l_Lean_Meta_synthInstance_x3f___lambda__6___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_Meta_synthInstance_etaExperiment; -return x_1; } } -static lean_object* _init_l_Lean_Meta_synthInstance_x3f___lambda__6___closed__2() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("result ", 7); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_synthInstance_x3f___lambda__6___closed__3() { -_start: +uint8_t x_104; uint8_t x_105; uint8_t x_106; uint8_t x_107; uint8_t x_108; uint8_t x_109; uint8_t x_110; uint8_t x_111; uint8_t x_112; uint8_t x_113; uint8_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; +x_104 = lean_ctor_get_uint8(x_20, 2); +x_105 = lean_ctor_get_uint8(x_20, 6); +x_106 = lean_ctor_get_uint8(x_20, 7); +x_107 = lean_ctor_get_uint8(x_20, 8); +x_108 = lean_ctor_get_uint8(x_20, 9); +x_109 = lean_ctor_get_uint8(x_20, 10); +x_110 = lean_ctor_get_uint8(x_20, 11); +x_111 = lean_ctor_get_uint8(x_20, 12); +lean_dec(x_20); +x_112 = 1; +x_113 = 0; +x_114 = 3; +x_115 = lean_alloc_ctor(0, 0, 13); +lean_ctor_set_uint8(x_115, 0, x_112); +lean_ctor_set_uint8(x_115, 1, x_112); +lean_ctor_set_uint8(x_115, 2, x_104); +lean_ctor_set_uint8(x_115, 3, x_113); +lean_ctor_set_uint8(x_115, 4, x_112); +lean_ctor_set_uint8(x_115, 5, x_114); +lean_ctor_set_uint8(x_115, 6, x_105); +lean_ctor_set_uint8(x_115, 7, x_106); +lean_ctor_set_uint8(x_115, 8, x_107); +lean_ctor_set_uint8(x_115, 9, x_108); +lean_ctor_set_uint8(x_115, 10, x_109); +lean_ctor_set_uint8(x_115, 11, x_110); +lean_ctor_set_uint8(x_115, 12, x_111); +lean_ctor_set(x_4, 0, x_115); +x_116 = l_Lean_Meta_getLocalInstances(x_4, x_5, x_6, x_7, x_8); +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 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_1, x_4, x_5, x_6, x_7, 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 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocess___closed__1; +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_123 = l_Lean_Meta_forallTelescopeReducing___at_Lean_Meta_getParamNames___spec__2___rarg(x_120, x_122, x_4, x_5, x_6, x_7, x_121); +if (lean_obj_tag(x_123) == 0) { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_synthInstance_x3f___lambda__6___closed__2; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Meta_synthInstance_x3f___lambda__6___closed__4() { -_start: +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_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_st_ref_get(x_5, 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_ctor_get(x_127, 1); +lean_inc(x_129); +lean_dec(x_127); +x_130 = lean_ctor_get(x_129, 2); +lean_inc(x_130); +lean_dec(x_129); +lean_inc(x_124); +x_131 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_131, 0, x_117); +lean_ctor_set(x_131, 1, x_124); +lean_inc(x_131); +x_132 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_synthInstance_x3f___spec__1(x_130, x_131); +if (lean_obj_tag(x_132) == 0) { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes(" (cached)", 9); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_synthInstance_x3f___lambda__6___closed__5() { -_start: +lean_object* x_133; lean_object* x_134; +lean_inc(x_124); +x_133 = lean_alloc_closure((void*)(l_Lean_Meta_synthInstance_x3f___lambda__2), 7, 2); +lean_closure_set(x_133, 0, x_124); +lean_closure_set(x_133, 1, x_2); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_134 = l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Meta_WHNF_0__Lean_Meta_toCtorWhenK___spec__1___rarg(x_133, x_112, x_4, x_5, x_6, x_7, x_128); +if (lean_obj_tag(x_134) == 0) { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_synthInstance_x3f___lambda__6___closed__4; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Meta_synthInstance_x3f___lambda__6___closed__6() { -_start: +lean_object* x_135; +x_135 = lean_ctor_get(x_134, 0); +lean_inc(x_135); +if (lean_obj_tag(x_135) == 0) { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("", 15); -return x_1; -} +lean_object* x_136; lean_object* x_137; lean_object* x_138; +lean_dec(x_124); +lean_dec(x_3); +x_136 = lean_ctor_get(x_134, 1); +lean_inc(x_136); +lean_dec(x_134); +x_137 = lean_box(0); +x_138 = l_Lean_Meta_synthInstance_x3f___lambda__3(x_131, x_137, x_4, x_5, x_6, x_7, x_136); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_9 = x_138; +goto block_18; } -static lean_object* _init_l_Lean_Meta_synthInstance_x3f___lambda__6___closed__7() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_synthInstance_x3f___lambda__6___closed__6; -x_2 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Meta_synthInstance_x3f___lambda__6___closed__8() { -_start: +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; uint8_t x_149; +x_139 = lean_ctor_get(x_134, 1); +lean_inc(x_139); +lean_dec(x_134); +x_140 = lean_ctor_get(x_135, 0); +lean_inc(x_140); +lean_dec(x_135); +x_141 = lean_alloc_closure((void*)(l_Lean_Meta_synthInstance_x3f___lambda__3___boxed), 7, 1); +lean_closure_set(x_141, 0, x_131); +lean_inc(x_4); +x_142 = l_Lean_Meta_openAbstractMVarsResult(x_140, x_4, x_5, x_6, x_7, x_139); +x_143 = lean_ctor_get(x_142, 0); +lean_inc(x_143); +x_144 = lean_ctor_get(x_143, 1); +lean_inc(x_144); +lean_dec(x_143); +x_145 = lean_ctor_get(x_142, 1); +lean_inc(x_145); +lean_dec(x_142); +x_146 = lean_ctor_get(x_144, 1); +lean_inc(x_146); +lean_dec(x_144); +lean_inc(x_3); +x_147 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_3, x_4, x_5, x_6, x_7, x_145); +x_148 = lean_ctor_get(x_147, 0); +lean_inc(x_148); +x_149 = lean_unbox(x_148); +lean_dec(x_148); +if (x_149 == 0) { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_synthInstance_x3f___lambda__6___closed__7; -x_2 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} +lean_object* x_150; lean_object* x_151; lean_object* x_152; +lean_dec(x_3); +x_150 = lean_ctor_get(x_147, 1); +lean_inc(x_150); +lean_dec(x_147); +x_151 = lean_box(0); +x_152 = l_Lean_Meta_synthInstance_x3f___lambda__4(x_124, x_146, x_141, x_151, x_4, x_5, x_6, x_7, x_150); +x_9 = x_152; +goto block_18; } -static lean_object* _init_l_Lean_Meta_synthInstance_x3f___lambda__6___closed__9() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_synthInstance_x3f___lambda__6___closed__3; -x_2 = l_Lean_Meta_synthInstance_x3f___lambda__6___closed__8; -x_3 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; -} +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_153 = lean_ctor_get(x_147, 1); +lean_inc(x_153); +lean_dec(x_147); +lean_inc(x_146); +x_154 = l_Lean_MessageData_ofExpr(x_146); +x_155 = l_Lean_Meta_synthInstance_x3f___lambda__6___closed__2; +x_156 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_156, 0, x_155); +lean_ctor_set(x_156, 1, x_154); +x_157 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; +x_158 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_158, 0, x_156); +lean_ctor_set(x_158, 1, x_157); +x_159 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_3, x_158, x_4, x_5, x_6, x_7, x_153); +x_160 = lean_ctor_get(x_159, 0); +lean_inc(x_160); +x_161 = lean_ctor_get(x_159, 1); +lean_inc(x_161); +lean_dec(x_159); +x_162 = l_Lean_Meta_synthInstance_x3f___lambda__4(x_124, x_146, x_141, x_160, x_4, x_5, x_6, x_7, x_161); +lean_dec(x_160); +x_9 = x_162; +goto block_18; } -static lean_object* _init_l_Lean_Meta_synthInstance_x3f___lambda__6___closed__10() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_synthInstance_x3f___lambda__6___closed__9; -x_2 = l_Lean_Meta_synthInstance_x3f___lambda__6___closed__5; -x_3 = lean_alloc_ctor(7, 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_synthInstance_x3f___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) { -_start: +else { -lean_object* x_9; lean_object* x_10; uint8_t x_11; lean_object* x_12; 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; uint8_t x_30; uint8_t x_31; uint8_t x_32; uint8_t x_33; uint8_t x_34; uint8_t x_35; lean_object* x_36; uint8_t x_37; -x_9 = lean_ctor_get(x_6, 2); -lean_inc(x_9); -x_10 = l_Lean_Meta_synthInstance_x3f___lambda__6___closed__1; -x_11 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_9, x_10); -lean_dec(x_9); -x_22 = lean_ctor_get(x_4, 0); -lean_inc(x_22); -x_23 = lean_ctor_get(x_4, 1); -lean_inc(x_23); -x_24 = lean_ctor_get(x_4, 2); -lean_inc(x_24); -x_25 = lean_ctor_get(x_4, 3); -lean_inc(x_25); -x_26 = lean_ctor_get(x_4, 4); -lean_inc(x_26); -x_27 = lean_ctor_get(x_4, 5); -lean_inc(x_27); -if (lean_is_exclusive(x_4)) { - lean_ctor_release(x_4, 0); - lean_ctor_release(x_4, 1); - lean_ctor_release(x_4, 2); - lean_ctor_release(x_4, 3); - lean_ctor_release(x_4, 4); - lean_ctor_release(x_4, 5); - x_28 = x_4; +lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; +lean_dec(x_131); +lean_dec(x_124); +lean_dec(x_4); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +x_163 = lean_ctor_get(x_134, 0); +lean_inc(x_163); +x_164 = lean_ctor_get(x_134, 1); +lean_inc(x_164); +if (lean_is_exclusive(x_134)) { + lean_ctor_release(x_134, 0); + lean_ctor_release(x_134, 1); + x_165 = x_134; } else { - lean_dec_ref(x_4); - x_28 = lean_box(0); + lean_dec_ref(x_134); + x_165 = lean_box(0); } -x_29 = lean_ctor_get_uint8(x_22, 2); -x_30 = lean_ctor_get_uint8(x_22, 6); -x_31 = lean_ctor_get_uint8(x_22, 7); -x_32 = lean_ctor_get_uint8(x_22, 8); -x_33 = lean_ctor_get_uint8(x_22, 9); -x_34 = lean_ctor_get_uint8(x_22, 10); -x_35 = lean_ctor_get_uint8(x_22, 11); -if (lean_is_exclusive(x_22)) { - x_36 = x_22; +if (lean_is_scalar(x_165)) { + x_166 = lean_alloc_ctor(1, 2, 0); } else { - lean_dec_ref(x_22); - x_36 = lean_box(0); + x_166 = x_165; +} +lean_ctor_set(x_166, 0, x_163); +lean_ctor_set(x_166, 1, x_164); +x_9 = x_166; +goto block_18; } -if (x_11 == 0) -{ -uint8_t x_123; -x_123 = 1; -x_37 = x_123; -goto block_122; } else { -uint8_t x_124; -x_124 = 0; -x_37 = x_124; -goto block_122; +lean_object* x_167; lean_object* x_168; lean_object* x_169; uint8_t x_170; +lean_dec(x_131); +lean_dec(x_2); +x_167 = lean_ctor_get(x_132, 0); +lean_inc(x_167); +lean_dec(x_132); +lean_inc(x_3); +x_168 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_3, x_4, x_5, x_6, x_7, x_128); +x_169 = lean_ctor_get(x_168, 0); +lean_inc(x_169); +x_170 = lean_unbox(x_169); +lean_dec(x_169); +if (x_170 == 0) +{ +lean_object* x_171; lean_object* x_172; lean_object* x_173; +lean_dec(x_3); +x_171 = lean_ctor_get(x_168, 1); +lean_inc(x_171); +lean_dec(x_168); +x_172 = lean_box(0); +x_173 = l_Lean_Meta_synthInstance_x3f___lambda__5(x_167, x_124, x_172, x_4, x_5, x_6, x_7, x_171); +x_9 = x_173; +goto block_18; } -block_21: -{ -if (lean_obj_tag(x_12) == 0) +else { -uint8_t x_13; -x_13 = !lean_is_exclusive(x_12); -if (x_13 == 0) +if (lean_obj_tag(x_167) == 0) { -return x_12; +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_174 = lean_ctor_get(x_168, 1); +lean_inc(x_174); +lean_dec(x_168); +x_175 = l_Lean_Meta_synthInstance_x3f___lambda__6___closed__9; +x_176 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_3, x_175, x_4, x_5, x_6, x_7, 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); +lean_dec(x_176); +x_179 = l_Lean_Meta_synthInstance_x3f___lambda__5(x_167, x_124, x_177, x_4, x_5, x_6, x_7, x_178); +lean_dec(x_177); +x_9 = x_179; +goto block_18; +} +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; +x_180 = lean_ctor_get(x_168, 1); +lean_inc(x_180); +lean_dec(x_168); +x_181 = lean_ctor_get(x_167, 0); +lean_inc(x_181); +x_182 = l_Lean_MessageData_ofExpr(x_181); +x_183 = l_Lean_Meta_synthInstance_x3f___lambda__6___closed__2; +x_184 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_184, 0, x_183); +lean_ctor_set(x_184, 1, x_182); +x_185 = l_Lean_Meta_synthInstance_x3f___lambda__6___closed__4; +x_186 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_186, 0, x_184); +lean_ctor_set(x_186, 1, x_185); +x_187 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_3, x_186, x_4, x_5, x_6, x_7, x_180); +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 = l_Lean_Meta_synthInstance_x3f___lambda__5(x_167, x_124, x_188, x_4, x_5, x_6, x_7, x_189); +lean_dec(x_188); +x_9 = x_190; +goto block_18; +} +} +} } 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_inc(x_14); -lean_dec(x_12); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_14); -lean_ctor_set(x_16, 1, x_15); -return x_16; -} +lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; +lean_dec(x_117); +lean_dec(x_4); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +x_191 = lean_ctor_get(x_123, 0); +lean_inc(x_191); +x_192 = lean_ctor_get(x_123, 1); +lean_inc(x_192); +if (lean_is_exclusive(x_123)) { + lean_ctor_release(x_123, 0); + lean_ctor_release(x_123, 1); + x_193 = x_123; +} else { + lean_dec_ref(x_123); + x_193 = lean_box(0); } -else -{ -uint8_t x_17; -x_17 = !lean_is_exclusive(x_12); -if (x_17 == 0) -{ -return x_12; +if (lean_is_scalar(x_193)) { + x_194 = lean_alloc_ctor(1, 2, 0); +} else { + x_194 = x_193; } -else -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_ctor_get(x_12, 0); -x_19 = lean_ctor_get(x_12, 1); -lean_inc(x_19); -lean_inc(x_18); -lean_dec(x_12); -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_ctor_set(x_194, 0, x_191); +lean_ctor_set(x_194, 1, x_192); +return x_194; } } } -block_122: +else { -uint8_t x_38; uint8_t 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; -x_38 = 1; -x_39 = 0; -x_40 = 3; -if (lean_is_scalar(x_36)) { - x_41 = lean_alloc_ctor(0, 0, 13); +lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; uint8_t x_201; uint8_t x_202; uint8_t x_203; uint8_t x_204; uint8_t x_205; uint8_t x_206; uint8_t x_207; uint8_t x_208; lean_object* x_209; uint8_t x_210; uint8_t x_211; uint8_t 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_195 = lean_ctor_get(x_4, 0); +x_196 = lean_ctor_get(x_4, 1); +x_197 = lean_ctor_get(x_4, 2); +x_198 = lean_ctor_get(x_4, 3); +x_199 = lean_ctor_get(x_4, 4); +x_200 = lean_ctor_get(x_4, 5); +lean_inc(x_200); +lean_inc(x_199); +lean_inc(x_198); +lean_inc(x_197); +lean_inc(x_196); +lean_inc(x_195); +lean_dec(x_4); +x_201 = lean_ctor_get_uint8(x_195, 2); +x_202 = lean_ctor_get_uint8(x_195, 6); +x_203 = lean_ctor_get_uint8(x_195, 7); +x_204 = lean_ctor_get_uint8(x_195, 8); +x_205 = lean_ctor_get_uint8(x_195, 9); +x_206 = lean_ctor_get_uint8(x_195, 10); +x_207 = lean_ctor_get_uint8(x_195, 11); +x_208 = lean_ctor_get_uint8(x_195, 12); +if (lean_is_exclusive(x_195)) { + x_209 = x_195; } else { - x_41 = x_36; -} -lean_ctor_set_uint8(x_41, 0, x_38); -lean_ctor_set_uint8(x_41, 1, x_38); -lean_ctor_set_uint8(x_41, 2, x_29); -lean_ctor_set_uint8(x_41, 3, x_39); -lean_ctor_set_uint8(x_41, 4, x_38); -lean_ctor_set_uint8(x_41, 5, x_40); -lean_ctor_set_uint8(x_41, 6, x_30); -lean_ctor_set_uint8(x_41, 7, x_31); -lean_ctor_set_uint8(x_41, 8, x_32); -lean_ctor_set_uint8(x_41, 9, x_33); -lean_ctor_set_uint8(x_41, 10, x_34); -lean_ctor_set_uint8(x_41, 11, x_35); -lean_ctor_set_uint8(x_41, 12, x_37); -if (lean_is_scalar(x_28)) { - x_42 = lean_alloc_ctor(0, 6, 0); + lean_dec_ref(x_195); + x_209 = lean_box(0); +} +x_210 = 1; +x_211 = 0; +x_212 = 3; +if (lean_is_scalar(x_209)) { + x_213 = lean_alloc_ctor(0, 0, 13); } else { - x_42 = x_28; -} -lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_23); -lean_ctor_set(x_42, 2, x_24); -lean_ctor_set(x_42, 3, x_25); -lean_ctor_set(x_42, 4, x_26); -lean_ctor_set(x_42, 5, x_27); -x_43 = l_Lean_Meta_getLocalInstances(x_42, x_5, x_6, x_7, x_8); -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 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_1, x_42, x_5, x_6, x_7, 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); -lean_dec(x_46); -x_49 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocess___closed__1; + x_213 = x_209; +} +lean_ctor_set_uint8(x_213, 0, x_210); +lean_ctor_set_uint8(x_213, 1, x_210); +lean_ctor_set_uint8(x_213, 2, x_201); +lean_ctor_set_uint8(x_213, 3, x_211); +lean_ctor_set_uint8(x_213, 4, x_210); +lean_ctor_set_uint8(x_213, 5, x_212); +lean_ctor_set_uint8(x_213, 6, x_202); +lean_ctor_set_uint8(x_213, 7, x_203); +lean_ctor_set_uint8(x_213, 8, x_204); +lean_ctor_set_uint8(x_213, 9, x_205); +lean_ctor_set_uint8(x_213, 10, x_206); +lean_ctor_set_uint8(x_213, 11, x_207); +lean_ctor_set_uint8(x_213, 12, x_208); +x_214 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_214, 0, x_213); +lean_ctor_set(x_214, 1, x_196); +lean_ctor_set(x_214, 2, x_197); +lean_ctor_set(x_214, 3, x_198); +lean_ctor_set(x_214, 4, x_199); +lean_ctor_set(x_214, 5, x_200); +x_215 = l_Lean_Meta_getLocalInstances(x_214, x_5, x_6, x_7, x_8); +x_216 = lean_ctor_get(x_215, 0); +lean_inc(x_216); +x_217 = lean_ctor_get(x_215, 1); +lean_inc(x_217); +lean_dec(x_215); +x_218 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_1, x_214, x_5, x_6, x_7, x_217); +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 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocess___closed__1; lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_42); -x_50 = l_Lean_Meta_forallTelescopeReducing___at_Lean_Meta_getParamNames___spec__2___rarg(x_47, x_49, x_42, x_5, x_6, x_7, x_48); -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; -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_st_ref_get(x_5, x_52); -x_54 = lean_ctor_get(x_53, 0); -lean_inc(x_54); -x_55 = lean_ctor_get(x_53, 1); -lean_inc(x_55); -lean_dec(x_53); -x_56 = lean_ctor_get(x_54, 1); -lean_inc(x_56); -lean_dec(x_54); -x_57 = lean_ctor_get(x_56, 2); -lean_inc(x_57); -lean_dec(x_56); -lean_inc(x_51); -x_58 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_58, 0, x_44); -lean_ctor_set(x_58, 1, x_51); -lean_inc(x_58); -x_59 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_synthInstance_x3f___spec__1(x_57, x_58); -if (lean_obj_tag(x_59) == 0) +lean_inc(x_214); +x_222 = l_Lean_Meta_forallTelescopeReducing___at_Lean_Meta_getParamNames___spec__2___rarg(x_219, x_221, x_214, x_5, x_6, x_7, x_220); +if (lean_obj_tag(x_222) == 0) { -lean_object* x_60; lean_object* x_61; -lean_inc(x_51); -x_60 = lean_alloc_closure((void*)(l_Lean_Meta_synthInstance_x3f___lambda__2), 7, 2); -lean_closure_set(x_60, 0, x_51); -lean_closure_set(x_60, 1, x_2); +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_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_5, x_224); +x_226 = lean_ctor_get(x_225, 0); +lean_inc(x_226); +x_227 = lean_ctor_get(x_225, 1); +lean_inc(x_227); +lean_dec(x_225); +x_228 = lean_ctor_get(x_226, 1); +lean_inc(x_228); +lean_dec(x_226); +x_229 = lean_ctor_get(x_228, 2); +lean_inc(x_229); +lean_dec(x_228); +lean_inc(x_223); +x_230 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_230, 0, x_216); +lean_ctor_set(x_230, 1, x_223); +lean_inc(x_230); +x_231 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_synthInstance_x3f___spec__1(x_229, x_230); +if (lean_obj_tag(x_231) == 0) +{ +lean_object* x_232; lean_object* x_233; +lean_inc(x_223); +x_232 = lean_alloc_closure((void*)(l_Lean_Meta_synthInstance_x3f___lambda__2), 7, 2); +lean_closure_set(x_232, 0, x_223); +lean_closure_set(x_232, 1, x_2); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_42); -x_61 = l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Meta_WHNF_0__Lean_Meta_toCtorWhenK___spec__1___rarg(x_60, x_38, x_42, x_5, x_6, x_7, x_55); -if (lean_obj_tag(x_61) == 0) +lean_inc(x_214); +x_233 = l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Meta_WHNF_0__Lean_Meta_toCtorWhenK___spec__1___rarg(x_232, x_210, x_214, x_5, x_6, x_7, x_227); +if (lean_obj_tag(x_233) == 0) { -lean_object* x_62; -x_62 = lean_ctor_get(x_61, 0); -lean_inc(x_62); -if (lean_obj_tag(x_62) == 0) +lean_object* x_234; +x_234 = lean_ctor_get(x_233, 0); +lean_inc(x_234); +if (lean_obj_tag(x_234) == 0) { -lean_object* x_63; lean_object* x_64; lean_object* x_65; -lean_dec(x_51); +lean_object* x_235; lean_object* x_236; lean_object* x_237; +lean_dec(x_223); lean_dec(x_3); -x_63 = lean_ctor_get(x_61, 1); -lean_inc(x_63); -lean_dec(x_61); -x_64 = lean_box(0); -x_65 = l_Lean_Meta_synthInstance_x3f___lambda__3(x_58, x_64, x_42, x_5, x_6, x_7, x_63); +x_235 = lean_ctor_get(x_233, 1); +lean_inc(x_235); +lean_dec(x_233); +x_236 = lean_box(0); +x_237 = l_Lean_Meta_synthInstance_x3f___lambda__3(x_230, x_236, x_214, x_5, x_6, x_7, x_235); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_42); -x_12 = x_65; -goto block_21; +lean_dec(x_214); +x_9 = x_237; +goto block_18; } 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; uint8_t x_76; -x_66 = lean_ctor_get(x_61, 1); -lean_inc(x_66); -lean_dec(x_61); -x_67 = lean_ctor_get(x_62, 0); -lean_inc(x_67); -lean_dec(x_62); -x_68 = lean_alloc_closure((void*)(l_Lean_Meta_synthInstance_x3f___lambda__3___boxed), 7, 1); -lean_closure_set(x_68, 0, x_58); -lean_inc(x_42); -x_69 = l_Lean_Meta_openAbstractMVarsResult(x_67, x_42, x_5, x_6, x_7, x_66); -x_70 = lean_ctor_get(x_69, 0); -lean_inc(x_70); -x_71 = lean_ctor_get(x_70, 1); -lean_inc(x_71); -lean_dec(x_70); -x_72 = lean_ctor_get(x_69, 1); -lean_inc(x_72); -lean_dec(x_69); -x_73 = lean_ctor_get(x_71, 1); -lean_inc(x_73); -lean_dec(x_71); +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; uint8_t x_248; +x_238 = lean_ctor_get(x_233, 1); +lean_inc(x_238); +lean_dec(x_233); +x_239 = lean_ctor_get(x_234, 0); +lean_inc(x_239); +lean_dec(x_234); +x_240 = lean_alloc_closure((void*)(l_Lean_Meta_synthInstance_x3f___lambda__3___boxed), 7, 1); +lean_closure_set(x_240, 0, x_230); +lean_inc(x_214); +x_241 = l_Lean_Meta_openAbstractMVarsResult(x_239, x_214, x_5, x_6, x_7, x_238); +x_242 = lean_ctor_get(x_241, 0); +lean_inc(x_242); +x_243 = lean_ctor_get(x_242, 1); +lean_inc(x_243); +lean_dec(x_242); +x_244 = lean_ctor_get(x_241, 1); +lean_inc(x_244); +lean_dec(x_241); +x_245 = lean_ctor_get(x_243, 1); +lean_inc(x_245); +lean_dec(x_243); lean_inc(x_3); -x_74 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_3, x_42, x_5, x_6, x_7, x_72); -x_75 = lean_ctor_get(x_74, 0); -lean_inc(x_75); -x_76 = lean_unbox(x_75); -lean_dec(x_75); -if (x_76 == 0) -{ -lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_246 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_3, x_214, x_5, x_6, x_7, x_244); +x_247 = lean_ctor_get(x_246, 0); +lean_inc(x_247); +x_248 = lean_unbox(x_247); +lean_dec(x_247); +if (x_248 == 0) +{ +lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_dec(x_3); -x_77 = lean_ctor_get(x_74, 1); -lean_inc(x_77); -lean_dec(x_74); -x_78 = lean_box(0); -x_79 = l_Lean_Meta_synthInstance_x3f___lambda__4(x_51, x_73, x_68, x_78, x_42, x_5, x_6, x_7, x_77); -x_12 = x_79; -goto block_21; -} -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; -x_80 = lean_ctor_get(x_74, 1); -lean_inc(x_80); -lean_dec(x_74); -lean_inc(x_73); -x_81 = l_Lean_MessageData_ofExpr(x_73); -x_82 = l_Lean_Meta_synthInstance_x3f___lambda__6___closed__3; -x_83 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_83, 0, x_82); -lean_ctor_set(x_83, 1, x_81); -x_84 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; -x_85 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_85, 0, x_83); -lean_ctor_set(x_85, 1, x_84); -x_86 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_3, x_85, x_42, x_5, x_6, x_7, x_80); -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 = l_Lean_Meta_synthInstance_x3f___lambda__4(x_51, x_73, x_68, x_87, x_42, x_5, x_6, x_7, x_88); -lean_dec(x_87); -x_12 = x_89; -goto block_21; -} -} +x_249 = lean_ctor_get(x_246, 1); +lean_inc(x_249); +lean_dec(x_246); +x_250 = lean_box(0); +x_251 = l_Lean_Meta_synthInstance_x3f___lambda__4(x_223, x_245, x_240, x_250, x_214, x_5, x_6, x_7, x_249); +x_9 = x_251; +goto block_18; } else { -uint8_t x_90; -lean_dec(x_58); -lean_dec(x_51); -lean_dec(x_42); +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 = lean_ctor_get(x_246, 1); +lean_inc(x_252); +lean_dec(x_246); +lean_inc(x_245); +x_253 = l_Lean_MessageData_ofExpr(x_245); +x_254 = l_Lean_Meta_synthInstance_x3f___lambda__6___closed__2; +x_255 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_255, 0, x_254); +lean_ctor_set(x_255, 1, x_253); +x_256 = l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__3; +x_257 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_257, 0, x_255); +lean_ctor_set(x_257, 1, x_256); +x_258 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_3, x_257, x_214, x_5, x_6, x_7, x_252); +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 = l_Lean_Meta_synthInstance_x3f___lambda__4(x_223, x_245, x_240, x_259, x_214, x_5, x_6, x_7, x_260); +lean_dec(x_259); +x_9 = x_261; +goto block_18; +} +} +} +else +{ +lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; +lean_dec(x_230); +lean_dec(x_223); +lean_dec(x_214); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_90 = !lean_is_exclusive(x_61); -if (x_90 == 0) -{ -x_12 = x_61; -goto block_21; +x_262 = lean_ctor_get(x_233, 0); +lean_inc(x_262); +x_263 = lean_ctor_get(x_233, 1); +lean_inc(x_263); +if (lean_is_exclusive(x_233)) { + lean_ctor_release(x_233, 0); + lean_ctor_release(x_233, 1); + x_264 = x_233; +} else { + lean_dec_ref(x_233); + x_264 = lean_box(0); } -else -{ -lean_object* x_91; lean_object* x_92; lean_object* x_93; -x_91 = lean_ctor_get(x_61, 0); -x_92 = lean_ctor_get(x_61, 1); -lean_inc(x_92); -lean_inc(x_91); -lean_dec(x_61); -x_93 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_93, 0, x_91); -lean_ctor_set(x_93, 1, x_92); -x_12 = x_93; -goto block_21; +if (lean_is_scalar(x_264)) { + x_265 = lean_alloc_ctor(1, 2, 0); +} else { + x_265 = x_264; } +lean_ctor_set(x_265, 0, x_262); +lean_ctor_set(x_265, 1, x_263); +x_9 = x_265; +goto block_18; } } else { -lean_object* x_94; lean_object* x_95; lean_object* x_96; uint8_t x_97; -lean_dec(x_58); +lean_object* x_266; lean_object* x_267; lean_object* x_268; uint8_t x_269; +lean_dec(x_230); lean_dec(x_2); -x_94 = lean_ctor_get(x_59, 0); -lean_inc(x_94); -lean_dec(x_59); +x_266 = lean_ctor_get(x_231, 0); +lean_inc(x_266); +lean_dec(x_231); lean_inc(x_3); -x_95 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_3, x_42, x_5, x_6, x_7, x_55); -x_96 = lean_ctor_get(x_95, 0); -lean_inc(x_96); -x_97 = lean_unbox(x_96); -lean_dec(x_96); -if (x_97 == 0) +x_267 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_3, x_214, x_5, x_6, x_7, x_227); +x_268 = lean_ctor_get(x_267, 0); +lean_inc(x_268); +x_269 = lean_unbox(x_268); +lean_dec(x_268); +if (x_269 == 0) { -lean_object* x_98; lean_object* x_99; lean_object* x_100; +lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_dec(x_3); -x_98 = lean_ctor_get(x_95, 1); -lean_inc(x_98); -lean_dec(x_95); -x_99 = lean_box(0); -x_100 = l_Lean_Meta_synthInstance_x3f___lambda__5(x_94, x_51, x_99, x_42, x_5, x_6, x_7, x_98); -x_12 = x_100; -goto block_21; +x_270 = lean_ctor_get(x_267, 1); +lean_inc(x_270); +lean_dec(x_267); +x_271 = lean_box(0); +x_272 = l_Lean_Meta_synthInstance_x3f___lambda__5(x_266, x_223, x_271, x_214, x_5, x_6, x_7, x_270); +x_9 = x_272; +goto block_18; } else { -if (lean_obj_tag(x_94) == 0) +if (lean_obj_tag(x_266) == 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; -x_101 = lean_ctor_get(x_95, 1); -lean_inc(x_101); -lean_dec(x_95); -x_102 = l_Lean_Meta_synthInstance_x3f___lambda__6___closed__10; -x_103 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_3, x_102, x_42, x_5, x_6, x_7, x_101); -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 = l_Lean_Meta_synthInstance_x3f___lambda__5(x_94, x_51, x_104, x_42, x_5, x_6, x_7, x_105); -lean_dec(x_104); -x_12 = x_106; -goto block_21; +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_273 = lean_ctor_get(x_267, 1); +lean_inc(x_273); +lean_dec(x_267); +x_274 = l_Lean_Meta_synthInstance_x3f___lambda__6___closed__9; +x_275 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_3, x_274, x_214, x_5, x_6, x_7, x_273); +x_276 = lean_ctor_get(x_275, 0); +lean_inc(x_276); +x_277 = lean_ctor_get(x_275, 1); +lean_inc(x_277); +lean_dec(x_275); +x_278 = l_Lean_Meta_synthInstance_x3f___lambda__5(x_266, x_223, x_276, x_214, x_5, x_6, x_7, x_277); +lean_dec(x_276); +x_9 = x_278; +goto block_18; } 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; -x_107 = lean_ctor_get(x_95, 1); -lean_inc(x_107); -lean_dec(x_95); -x_108 = lean_ctor_get(x_94, 0); -lean_inc(x_108); -x_109 = l_Lean_MessageData_ofExpr(x_108); -x_110 = l_Lean_Meta_synthInstance_x3f___lambda__6___closed__3; -x_111 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_111, 0, x_110); -lean_ctor_set(x_111, 1, x_109); -x_112 = l_Lean_Meta_synthInstance_x3f___lambda__6___closed__5; -x_113 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_113, 0, x_111); -lean_ctor_set(x_113, 1, x_112); -x_114 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_3, x_113, x_42, x_5, x_6, x_7, x_107); -x_115 = lean_ctor_get(x_114, 0); -lean_inc(x_115); -x_116 = lean_ctor_get(x_114, 1); -lean_inc(x_116); -lean_dec(x_114); -x_117 = l_Lean_Meta_synthInstance_x3f___lambda__5(x_94, x_51, x_115, x_42, x_5, x_6, x_7, x_116); -lean_dec(x_115); -x_12 = x_117; -goto block_21; +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_279 = lean_ctor_get(x_267, 1); +lean_inc(x_279); +lean_dec(x_267); +x_280 = lean_ctor_get(x_266, 0); +lean_inc(x_280); +x_281 = l_Lean_MessageData_ofExpr(x_280); +x_282 = l_Lean_Meta_synthInstance_x3f___lambda__6___closed__2; +x_283 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_283, 0, x_282); +lean_ctor_set(x_283, 1, x_281); +x_284 = l_Lean_Meta_synthInstance_x3f___lambda__6___closed__4; +x_285 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_285, 0, x_283); +lean_ctor_set(x_285, 1, x_284); +x_286 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_3, x_285, x_214, x_5, x_6, x_7, x_279); +x_287 = lean_ctor_get(x_286, 0); +lean_inc(x_287); +x_288 = lean_ctor_get(x_286, 1); +lean_inc(x_288); +lean_dec(x_286); +x_289 = l_Lean_Meta_synthInstance_x3f___lambda__5(x_266, x_223, x_287, x_214, x_5, x_6, x_7, x_288); +lean_dec(x_287); +x_9 = x_289; +goto block_18; +} +} +} +} +else +{ +lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; +lean_dec(x_216); +lean_dec(x_214); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +x_290 = lean_ctor_get(x_222, 0); +lean_inc(x_290); +x_291 = lean_ctor_get(x_222, 1); +lean_inc(x_291); +if (lean_is_exclusive(x_222)) { + lean_ctor_release(x_222, 0); + lean_ctor_release(x_222, 1); + x_292 = x_222; +} else { + lean_dec_ref(x_222); + x_292 = lean_box(0); +} +if (lean_is_scalar(x_292)) { + x_293 = lean_alloc_ctor(1, 2, 0); +} else { + x_293 = x_292; +} +lean_ctor_set(x_293, 0, x_290); +lean_ctor_set(x_293, 1, x_291); +return x_293; } } +block_18: +{ +if (lean_obj_tag(x_9) == 0) +{ +uint8_t x_10; +x_10 = !lean_is_exclusive(x_9); +if (x_10 == 0) +{ +return x_9; +} +else +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = lean_ctor_get(x_9, 0); +x_12 = lean_ctor_get(x_9, 1); +lean_inc(x_12); +lean_inc(x_11); +lean_dec(x_9); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_11); +lean_ctor_set(x_13, 1, x_12); +return x_13; } } else { -uint8_t x_118; -lean_dec(x_44); -lean_dec(x_42); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -x_118 = !lean_is_exclusive(x_50); -if (x_118 == 0) +uint8_t x_14; +x_14 = !lean_is_exclusive(x_9); +if (x_14 == 0) { -return x_50; +return x_9; } else { -lean_object* x_119; lean_object* x_120; lean_object* x_121; -x_119 = lean_ctor_get(x_50, 0); -x_120 = lean_ctor_get(x_50, 1); -lean_inc(x_120); -lean_inc(x_119); -lean_dec(x_50); -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_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_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; } } } @@ -23771,7 +25070,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_5, 2); lean_inc(x_8); x_9 = l_Lean_Meta_synthInstance_x3f___lambda__7___closed__1; -x_10 = l_Lean_Option_get___at_Std_Format_pretty_x27___spec__1(x_8, x_9); +x_10 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_8, x_9); lean_dec(x_8); lean_inc(x_1); x_11 = lean_alloc_closure((void*)(l_Lean_Meta_synthInstance_x3f___lambda__1___boxed), 7, 1); @@ -23833,7 +25132,7 @@ if (lean_obj_tag(x_11) == 0) lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = l_Lean_Meta_synthInstance_x3f___closed__1; x_13 = lean_box(0); -x_14 = l_Lean_profileitM___at_Lean_Meta_synthInstance_x3f___spec__17___rarg(x_12, x_8, x_9, x_13, x_3, x_4, x_5, x_6, x_7); +x_14 = l_Lean_profileitM___at_Lean_Meta_synthInstance_x3f___spec__16___rarg(x_12, x_8, x_9, x_13, x_3, x_4, x_5, x_6, x_7); lean_dec(x_8); return x_14; } @@ -23844,7 +25143,7 @@ x_15 = lean_ctor_get(x_11, 0); lean_inc(x_15); lean_dec(x_11); x_16 = l_Lean_Meta_synthInstance_x3f___closed__1; -x_17 = l_Lean_profileitM___at_Lean_Meta_synthInstance_x3f___spec__17___rarg(x_16, x_8, x_9, x_15, x_3, x_4, x_5, x_6, x_7); +x_17 = l_Lean_profileitM___at_Lean_Meta_synthInstance_x3f___spec__16___rarg(x_16, x_8, x_9, x_15, x_3, x_4, x_5, x_6, x_7); lean_dec(x_8); return x_17; } @@ -23964,11 +25263,11 @@ x_8 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_synthInstance_x3f___spec return x_8; } } -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_synthInstance_x3f___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_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_synthInstance_x3f___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) { _start: { lean_object* x_7; -x_7 = l_MonadExcept_ofExcept___at_Lean_Meta_synthInstance_x3f___spec__16(x_1, x_2, x_3, x_4, x_5, x_6); +x_7 = l_MonadExcept_ofExcept___at_Lean_Meta_synthInstance_x3f___spec__15(x_1, x_2, x_3, x_4, x_5, x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); @@ -23992,6 +25291,32 @@ lean_dec(x_5); return x_14; } } +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_synthInstance_x3f___spec__13___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: +{ +uint8_t x_15; uint8_t x_16; double x_17; lean_object* x_18; +x_15 = lean_unbox(x_5); +lean_dec(x_5); +x_16 = lean_unbox(x_7); +lean_dec(x_7); +x_17 = lean_unbox_float(x_8); +lean_dec(x_8); +x_18 = l_Lean_withTraceNode___at_Lean_Meta_synthInstance_x3f___spec__13___lambda__2(x_1, x_2, x_3, x_4, x_15, x_6, x_16, x_17, x_9, x_10, x_11, x_12, x_13, x_14); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_synthInstance_x3f___spec__13___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) { +_start: +{ +uint8_t x_13; uint8_t x_14; lean_object* x_15; +x_13 = lean_unbox(x_4); +lean_dec(x_4); +x_14 = lean_unbox(x_6); +lean_dec(x_6); +x_15 = l_Lean_withTraceNode___at_Lean_Meta_synthInstance_x3f___spec__13___lambda__3(x_1, x_2, x_3, x_13, x_5, x_14, x_7, x_8, x_9, x_10, x_11, x_12); +return x_15; +} +} LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_synthInstance_x3f___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) { _start: { @@ -24002,11 +25327,11 @@ x_11 = l_Lean_withTraceNode___at_Lean_Meta_synthInstance_x3f___spec__13(x_1, x_2 return x_11; } } -LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Meta_synthInstance_x3f___spec__17___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Meta_synthInstance_x3f___spec__16___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_10; -x_10 = l_Lean_profileitM___at_Lean_Meta_synthInstance_x3f___spec__17___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_10 = l_Lean_profileitM___at_Lean_Meta_synthInstance_x3f___spec__16___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_2); lean_dec(x_1); return x_10; @@ -25259,7 +26584,7 @@ lean_dec(x_3); return x_9; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__1() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -25269,17 +26594,17 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__2() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__1; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__1; x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_6____closed__8; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__3() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__3() { _start: { lean_object* x_1; @@ -25287,17 +26612,17 @@ x_1 = lean_mk_string_from_bytes("initFn", 6); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__4() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__2; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__3; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__2; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__3; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__5() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__5() { _start: { lean_object* x_1; @@ -25305,37 +26630,37 @@ x_1 = lean_mk_string_from_bytes("_@", 2); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__6() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__4; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__5; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__4; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__7() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__6; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__6; x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_6____closed__7; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__8() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__7; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__7; x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_6____closed__8; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__9() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__9() { _start: { lean_object* x_1; @@ -25343,17 +26668,17 @@ x_1 = lean_mk_string_from_bytes("SynthInstance", 13); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__10() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__8; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__9; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__8; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__9; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__11() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__11() { _start: { lean_object* x_1; @@ -25361,33 +26686,33 @@ x_1 = lean_mk_string_from_bytes("_hyg", 4); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__12() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__10; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__11; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__10; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__11; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__13() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__12; -x_2 = lean_unsigned_to_nat(11355u); +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__12; +x_2 = lean_unsigned_to_nat(11245u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; x_2 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___lambda__2___closed__2; x_3 = 0; -x_4 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__13; +x_4 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__13; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); if (lean_obj_tag(x_5) == 0) { @@ -25650,21 +26975,6 @@ if (lean_io_result_is_error(res)) return res; l_Lean_Meta_synthInstance_maxSize = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Meta_synthInstance_maxSize); lean_dec_ref(res); -}l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_88____closed__1 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_88____closed__1(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_88____closed__1); -l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_88____closed__2 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_88____closed__2(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_88____closed__2); -l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_88____closed__3 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_88____closed__3(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_88____closed__3); -l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_88____closed__4 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_88____closed__4(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_88____closed__4); -l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_88____closed__5 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_88____closed__5(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_88____closed__5); -if (builtin) {res = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_88_(lean_io_mk_world()); -if (lean_io_result_is_error(res)) return res; -l_Lean_Meta_synthInstance_etaExperiment = lean_io_result_get_value(res); -lean_mark_persistent(l_Lean_Meta_synthInstance_etaExperiment); -lean_dec_ref(res); }l_Lean_Meta_SynthInstance_getMaxHeartbeats___closed__1 = _init_l_Lean_Meta_SynthInstance_getMaxHeartbeats___closed__1(); lean_mark_persistent(l_Lean_Meta_SynthInstance_getMaxHeartbeats___closed__1); l_Lean_Meta_SynthInstance_instInhabitedInstance___closed__1 = _init_l_Lean_Meta_SynthInstance_instInhabitedInstance___closed__1(); @@ -25773,19 +27083,19 @@ l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_SynthInstance_ lean_mark_persistent(l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_SynthInstance_newSubgoal___spec__11___rarg___closed__2); l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_SynthInstance_newSubgoal___spec__11___rarg___closed__3 = _init_l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_SynthInstance_newSubgoal___spec__11___rarg___closed__3(); lean_mark_persistent(l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_SynthInstance_newSubgoal___spec__11___rarg___closed__3); -l_Lean_withSeconds___at_Lean_Meta_SynthInstance_newSubgoal___spec__14___closed__1 = _init_l_Lean_withSeconds___at_Lean_Meta_SynthInstance_newSubgoal___spec__14___closed__1(); -l_Lean_withOptProfile___at_Lean_Meta_SynthInstance_newSubgoal___spec__12___closed__1 = _init_l_Lean_withOptProfile___at_Lean_Meta_SynthInstance_newSubgoal___spec__12___closed__1(); -lean_mark_persistent(l_Lean_withOptProfile___at_Lean_Meta_SynthInstance_newSubgoal___spec__12___closed__1); -l_Lean_withOptProfile___at_Lean_Meta_SynthInstance_newSubgoal___spec__12___closed__2 = _init_l_Lean_withOptProfile___at_Lean_Meta_SynthInstance_newSubgoal___spec__12___closed__2(); -lean_mark_persistent(l_Lean_withOptProfile___at_Lean_Meta_SynthInstance_newSubgoal___spec__12___closed__2); -l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___closed__1 = _init_l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___closed__1(); -lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___closed__1); -l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___closed__2 = _init_l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___closed__2(); -lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___closed__2); -l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___closed__3 = _init_l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___closed__3(); -lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___closed__3); -l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___closed__4 = _init_l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___closed__4(); -lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___closed__4); +l_Lean_withSeconds___at_Lean_Meta_SynthInstance_newSubgoal___spec__12___closed__1 = _init_l_Lean_withSeconds___at_Lean_Meta_SynthInstance_newSubgoal___spec__12___closed__1(); +l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__1 = _init_l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__1); +l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__2 = _init_l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__2); +l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__3 = _init_l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__3(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__3); +l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__4 = _init_l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__4(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__4); +l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__5 = _init_l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__5(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3___closed__5); +l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__4___closed__1 = _init_l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__4___closed__1); l_Lean_withTraceNode_x27___at_Lean_Meta_SynthInstance_newSubgoal___spec__8___closed__1 = _init_l_Lean_withTraceNode_x27___at_Lean_Meta_SynthInstance_newSubgoal___spec__8___closed__1(); lean_mark_persistent(l_Lean_withTraceNode_x27___at_Lean_Meta_SynthInstance_newSubgoal___spec__8___closed__1); l_Lean_Meta_SynthInstance_newSubgoal___lambda__2___closed__1 = _init_l_Lean_Meta_SynthInstance_newSubgoal___lambda__2___closed__1(); @@ -25996,8 +27306,6 @@ l_Lean_Meta_synthInstance_x3f___lambda__6___closed__8 = _init_l_Lean_Meta_synthI lean_mark_persistent(l_Lean_Meta_synthInstance_x3f___lambda__6___closed__8); l_Lean_Meta_synthInstance_x3f___lambda__6___closed__9 = _init_l_Lean_Meta_synthInstance_x3f___lambda__6___closed__9(); lean_mark_persistent(l_Lean_Meta_synthInstance_x3f___lambda__6___closed__9); -l_Lean_Meta_synthInstance_x3f___lambda__6___closed__10 = _init_l_Lean_Meta_synthInstance_x3f___lambda__6___closed__10(); -lean_mark_persistent(l_Lean_Meta_synthInstance_x3f___lambda__6___closed__10); l_Lean_Meta_synthInstance_x3f___lambda__7___closed__1 = _init_l_Lean_Meta_synthInstance_x3f___lambda__7___closed__1(); lean_mark_persistent(l_Lean_Meta_synthInstance_x3f___lambda__7___closed__1); l_Lean_Meta_synthInstance_x3f___closed__1 = _init_l_Lean_Meta_synthInstance_x3f___closed__1(); @@ -26020,33 +27328,33 @@ l___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___lambda__2___c lean_mark_persistent(l___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___lambda__2___closed__5); l___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___lambda__2___closed__6 = _init_l___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___lambda__2___closed__6(); lean_mark_persistent(l___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___lambda__2___closed__6); -l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__1 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__1(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__1); -l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__2 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__2(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__2); -l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__3 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__3(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__3); -l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__4 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__4(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__4); -l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__5 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__5(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__5); -l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__6 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__6(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__6); -l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__7 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__7(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__7); -l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__8 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__8(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__8); -l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__9 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__9(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__9); -l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__10 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__10(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__10); -l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__11 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__11(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__11); -l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__12 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__12(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__12); -l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__13 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__13(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355____closed__13); -res = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11355_(lean_io_mk_world()); +l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__1 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__1(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__1); +l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__2 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__2(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__2); +l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__3 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__3(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__3); +l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__4 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__4(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__4); +l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__5 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__5(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__5); +l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__6 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__6(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__6); +l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__7 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__7(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__7); +l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__8 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__8(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__8); +l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__9 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__9(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__9); +l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__10 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__10(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__10); +l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__11 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__11(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__11); +l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__12 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__12(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__12); +l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__13 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__13(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245____closed__13); +res = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_11245_(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/Meta/Tactic/AC/Main.c b/stage0/stdlib/Lean/Meta/Tactic/AC/Main.c index 11fde82d6b08..29b95be05ca8 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/AC/Main.c +++ b/stage0/stdlib/Lean/Meta/Tactic/AC/Main.c @@ -137,7 +137,6 @@ static lean_object* l_Lean_Meta_AC_buildNormProof_mkContext___closed__1; static lean_object* l___regBuiltin_Lean_Meta_AC_acRflTactic_declRange___closed__4; lean_object* lean_array_to_list(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_AC_buildNormProof_convertTarget(lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Data_AC_removeNeutrals_loop___at_Lean_Meta_AC_buildNormProof___spec__3(lean_object*, lean_object*); lean_object* l_Lean_MVarId_withContext___at_Lean_Elab_Tactic_withMainContext___spec__1___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_num___override(lean_object*, lean_object*); @@ -188,6 +187,7 @@ uint8_t lean_nat_dec_eq(lean_object*, lean_object*); lean_object* l_Lean_mkApp3(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_AC_preContext___lambda__1(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_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkListLit(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_AssocList_contains___at_Lean_Meta_AC_toACExpr___spec__7(lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); @@ -5545,7 +5545,7 @@ lean_inc(x_16); x_17 = lean_ctor_get(x_15, 1); lean_inc(x_17); lean_dec(x_15); -x_18 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_16, x_2, x_3, x_4, x_5, x_17); +x_18 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_16, x_2, x_3, x_4, x_5, x_17); x_19 = lean_ctor_get(x_18, 0); lean_inc(x_19); x_20 = lean_ctor_get(x_18, 1); diff --git a/stage0/stdlib/Lean/Meta/Tactic/Apply.c b/stage0/stdlib/Lean/Meta/Tactic/Apply.c index 98d9ddf273a9..9e6e0e915d1b 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Apply.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Apply.c @@ -108,7 +108,6 @@ lean_object* lean_st_mk_ref(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Apply_0__Lean_Meta_dependsOnOthers___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Apply_0__Lean_Meta_partitionDependentMVars___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_to_list(lean_object*, lean_object*); -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Apply_0__Lean_Meta_partitionDependentMVars___closed__1; lean_object* l_Lean_Meta_synthInstance(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_append(lean_object*, lean_object*); @@ -133,6 +132,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_apply(lean_object*, lean_object*, lean_obje LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Apply_0__Lean_Meta_partitionDependentMVars___spec__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_object* l_Lean_MVarId_isAssigned___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_forM_loop___at_Lean_Meta_synthAppInstances___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*); LEAN_EXPORT lean_object* l_Lean_Meta_synthAppInstances___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_Name_mkStr2(lean_object*, lean_object*); @@ -3101,7 +3101,7 @@ lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean x_21 = lean_ctor_get(x_20, 1); lean_inc(x_21); lean_dec(x_20); -x_22 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_6, x_8, x_9, x_10, x_11, x_21); +x_22 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_6, x_8, x_9, x_10, x_11, x_21); x_23 = lean_ctor_get(x_22, 0); lean_inc(x_23); x_24 = lean_ctor_get(x_22, 1); @@ -4626,7 +4626,7 @@ lean_inc(x_11); x_12 = lean_ctor_get(x_10, 1); lean_inc(x_12); lean_dec(x_10); -x_13 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_11, x_3, x_4, x_5, x_6, x_12); +x_13 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_11, x_3, x_4, x_5, x_6, x_12); x_14 = lean_ctor_get(x_13, 0); lean_inc(x_14); x_15 = lean_ctor_get(x_13, 1); diff --git a/stage0/stdlib/Lean/Meta/Tactic/Cases.c b/stage0/stdlib/Lean/Meta/Tactic/Cases.c index a0189f125f23..81fea3d88deb 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Cases.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Cases.c @@ -204,7 +204,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_toByCas LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__15___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_array_to_list(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Cases_cases___lambda__2___closed__8; -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_generalizeIndices___spec__1___closed__1; lean_object* l_Lean_MVarId_induction(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); @@ -301,6 +300,7 @@ LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_v static lean_object* l___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_mkEqAndProof___closed__4; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__42___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_getLocalInstances(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__23___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at___private_Lean_Meta_Tactic_Cases_0__Lean_Meta_Cases_hasIndepIndices___spec__20(lean_object*, lean_object*, lean_object*, lean_object*); @@ -15281,7 +15281,7 @@ LEAN_EXPORT lean_object* l_Lean_MVarId_casesAnd___lambda__1(lean_object* x_1, le { lean_object* x_7; lean_object* x_8; uint8_t x_9; x_7 = l_Lean_LocalDecl_type(x_1); -x_8 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_7, x_2, x_3, x_4, x_5, x_6); +x_8 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_7, x_2, x_3, x_4, x_5, x_6); x_9 = !lean_is_exclusive(x_8); if (x_9 == 0) { @@ -15419,7 +15419,7 @@ LEAN_EXPORT lean_object* l_Lean_MVarId_substEqs___lambda__1(lean_object* x_1, le { lean_object* x_7; lean_object* x_8; uint8_t x_9; x_7 = l_Lean_LocalDecl_type(x_1); -x_8 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_7, x_2, x_3, x_4, x_5, x_6); +x_8 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_7, x_2, x_3, x_4, x_5, x_6); x_9 = !lean_is_exclusive(x_8); if (x_9 == 0) { diff --git a/stage0/stdlib/Lean/Meta/Tactic/Cleanup.c b/stage0/stdlib/Lean/Meta/Tactic/Cleanup.c index 0b77f624c7b9..1095cf703e88 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Cleanup.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Cleanup.c @@ -60,7 +60,6 @@ lean_object* lean_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cleanup_0__Lean_Meta_cleanupCore_addUsedFVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at___private_Lean_Meta_Tactic_Cleanup_0__Lean_Meta_cleanupCore___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_Tactic_Cleanup_0__Lean_Meta_cleanupCore___spec__6___at___private_Lean_Meta_Tactic_Cleanup_0__Lean_Meta_cleanupCore___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Meta_Tactic_Cleanup_0__Lean_Meta_cleanupCore_collectPropsStep___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*); static lean_object* l___private_Lean_Meta_Tactic_Cleanup_0__Lean_Meta_cleanupCore___closed__1; @@ -85,6 +84,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Cleanup_0__Lean_Meta_clean LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at___private_Lean_Meta_Tactic_Cleanup_0__Lean_Meta_cleanupCore___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_RBTree_contains___rarg___boxed(lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_getLocalInstances(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_local_ctx_erase(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at___private_Lean_Meta_Tactic_Cleanup_0__Lean_Meta_cleanupCore___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -3814,7 +3814,7 @@ return x_82; block_50: { lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_28 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_26, x_3, x_4, x_5, x_6, x_27); +x_28 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_26, x_3, x_4, x_5, x_6, x_27); x_29 = lean_ctor_get(x_28, 0); lean_inc(x_29); x_30 = lean_ctor_get(x_28, 1); diff --git a/stage0/stdlib/Lean/Meta/Tactic/Congr.c b/stage0/stdlib/Lean/Meta/Tactic/Congr.c index 7d9bb0f06135..75476b36a1b7 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Congr.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Congr.c @@ -47,6 +47,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Congr_0__Lean_applyCongrTh lean_object* l_Lean_Meta_mkHCongr(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_MVarId_hcongr_x3f___lambda__1___boxed(lean_object*, lean_object*, 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_12510_(uint8_t, uint8_t); static lean_object* l_Lean_MVarId_congr_x3f___closed__2; static lean_object* l_Lean_MVarId_congrImplies_x3f___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_MVarId_congr_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -78,7 +79,6 @@ LEAN_EXPORT lean_object* l_Lean_MVarId_congrN_post(uint8_t, lean_object*, lean_o uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_congrPre___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Congr_0__Lean_applyCongrThm_x3f___closed__2; -uint8_t l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_12336_(uint8_t, uint8_t); 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_MVarId_congr_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); @@ -2374,7 +2374,7 @@ lean_dec(x_9); x_24 = 2; x_25 = lean_unbox(x_22); lean_dec(x_22); -x_26 = l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_12336_(x_25, x_24); +x_26 = l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_12510_(x_25, x_24); if (x_26 == 0) { lean_object* x_27; diff --git a/stage0/stdlib/Lean/Meta/Tactic/Contradiction.c b/stage0/stdlib/Lean/Meta/Tactic/Contradiction.c index 4e1a1773000b..396f49bff43a 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Contradiction.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Contradiction.c @@ -139,7 +139,6 @@ static lean_object* l_Lean_Meta_ElimEmptyInductive_instMonadBacktrackSavedStateM LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Contradiction_0__Lean_Meta_nestedFalseElim(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Contradiction_0__Lean_Meta_nestedFalseElim___lambda__1(lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withMVarContextImp___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_hasLooseBVars(lean_object*); lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); @@ -203,6 +202,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Contradiction_0__Lean_Meta LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Contradiction_0__Lean_Meta_isElimEmptyInductiveCandidate(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Contradiction___hyg_5021____closed__3; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_ElimEmptyInductive_elim___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_ElimEmptyInductive_elim___spec__1___closed__1; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); @@ -4758,7 +4758,7 @@ lean_dec(x_5); x_11 = l_Lean_LocalDecl_toExpr(x_1); lean_dec(x_1); x_12 = l_Lean_mkAppN(x_11, x_2); -x_13 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_12, x_6, x_7, x_8, x_9, x_10); +x_13 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_12, 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); @@ -5859,7 +5859,7 @@ if (x_17 == 0) { lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; lean_inc(x_6); -x_18 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_6, x_9, x_10, x_11, x_12, x_13); +x_18 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_6, x_9, x_10, 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); diff --git a/stage0/stdlib/Lean/Meta/Tactic/ElimInfo.c b/stage0/stdlib/Lean/Meta/Tactic/ElimInfo.c index 97bb858f5017..459c4c71317d 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/ElimInfo.c +++ b/stage0/stdlib/Lean/Meta/Tactic/ElimInfo.c @@ -235,7 +235,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_getCustomEliminators___boxed(lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); static lean_object* l_Lean_Meta_addCustomEliminator___closed__1; lean_object* lean_array_to_list(lean_object*, lean_object*); -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_SMap_insert___at_Lean_Meta_addCustomEliminatorEntry___spec__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_addCustomEliminator___spec__1___closed__13; static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimAltInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_31____closed__6; @@ -373,6 +372,7 @@ uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l_Prod_repr___at___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprCustomEliminators____x40_Lean_Meta_Tactic_ElimInfo___hyg_2218____spec__10___closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_getElimInfo___lambda__1(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_Meta_getElimInfo___spec__6___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Meta_mkCustomEliminator___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addCustomEliminatorEntry___spec__4(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_ElimInfo_0__Lean_Meta_reprElimInfo____x40_Lean_Meta_Tactic_ElimInfo___hyg_192____closed__12; @@ -3542,7 +3542,7 @@ lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean x_11 = lean_array_uget(x_3, x_2); x_12 = lean_unsigned_to_nat(0u); x_13 = lean_array_uset(x_3, x_2, x_12); -x_14 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_11, x_4, x_5, x_6, x_7, x_8); +x_14 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_11, x_4, x_5, 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); @@ -13307,7 +13307,7 @@ lean_inc(x_25); x_26 = lean_ctor_get(x_24, 1); lean_inc(x_26); lean_dec(x_24); -x_27 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_25, x_6, x_7, x_8, x_9, x_26); +x_27 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_25, x_6, x_7, 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); diff --git a/stage0/stdlib/Lean/Meta/Tactic/Generalize.c b/stage0/stdlib/Lean/Meta/Tactic/Generalize.c index ef06d2ea2445..a0094620c587 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Generalize.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Generalize.c @@ -54,7 +54,6 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_MVarId_generalizeHyp LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_MVarId_generalizeHyp___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_mkHEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_hasLooseBVars(lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_generalizeHyp(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_MVarId_generalizeHyp___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -73,6 +72,7 @@ static lean_object* l___private_Lean_Meta_Tactic_Generalize_0__Lean_Meta_general 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_Array_foldlMUnsafe_fold___at_Lean_MVarId_generalizeHyp___spec__4(size_t, lean_object*, lean_object*, size_t, size_t, 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_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Generalize_0__Lean_Meta_generalizeCore___spec__1(size_t, size_t, lean_object*); lean_object* l_Lean_Meta_mkEqRefl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_generalizeHyp___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -257,7 +257,7 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean x_12 = lean_array_fget(x_1, x_3); x_13 = lean_ctor_get(x_12, 0); lean_inc(x_13); -x_14 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_13, x_4, x_5, x_6, x_7, x_8); +x_14 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_13, x_4, x_5, 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); @@ -277,7 +277,7 @@ lean_inc(x_18); x_19 = lean_ctor_get(x_17, 1); lean_inc(x_19); lean_dec(x_17); -x_20 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_18, x_4, x_5, x_6, x_7, x_19); +x_20 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_18, x_4, x_5, x_6, x_7, x_19); x_21 = lean_ctor_get(x_20, 0); lean_inc(x_21); x_22 = lean_ctor_get(x_20, 1); @@ -600,7 +600,7 @@ lean_dec(x_22); x_25 = lean_ctor_get(x_15, 0); lean_inc(x_25); lean_dec(x_15); -x_26 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_25, x_5, x_6, x_7, x_8, x_24); +x_26 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_25, x_5, x_6, x_7, x_8, x_24); x_27 = lean_ctor_get(x_26, 0); lean_inc(x_27); x_28 = lean_ctor_get(x_26, 1); @@ -620,7 +620,7 @@ lean_inc(x_30); x_31 = lean_ctor_get(x_29, 1); lean_inc(x_31); lean_dec(x_29); -x_32 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_30, x_5, x_6, x_7, x_8, x_31); +x_32 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(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); @@ -1466,7 +1466,7 @@ lean_inc(x_15); x_16 = lean_ctor_get(x_14, 1); lean_inc(x_16); lean_dec(x_14); -x_17 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_15, x_4, x_5, x_6, x_7, x_16); +x_17 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_15, x_4, x_5, x_6, x_7, x_16); x_18 = lean_ctor_get(x_17, 0); lean_inc(x_18); x_19 = lean_ctor_get(x_17, 1); @@ -1778,7 +1778,7 @@ if (x_14 == 0) { lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; size_t x_19; size_t x_20; lean_object* x_21; x_15 = lean_ctor_get(x_11, 0); -x_16 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_15, x_4, x_5, x_6, x_7, x_8); +x_16 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_15, x_4, x_5, x_6, x_7, x_8); x_17 = lean_ctor_get(x_16, 0); lean_inc(x_17); x_18 = lean_ctor_get(x_16, 1); @@ -1803,7 +1803,7 @@ lean_inc(x_25); lean_inc(x_24); lean_inc(x_23); lean_dec(x_11); -x_26 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_23, x_4, x_5, x_6, x_7, x_8); +x_26 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_23, x_4, x_5, x_6, x_7, x_8); x_27 = lean_ctor_get(x_26, 0); lean_inc(x_27); x_28 = lean_ctor_get(x_26, 1); @@ -2128,7 +2128,7 @@ lean_inc(x_15); x_16 = lean_ctor_get(x_14, 1); lean_inc(x_16); lean_dec(x_14); -x_17 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_15, x_7, x_8, x_9, x_10, x_16); +x_17 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_15, x_7, x_8, x_9, x_10, x_16); x_18 = lean_ctor_get(x_17, 0); lean_inc(x_18); x_19 = lean_ctor_get(x_17, 1); diff --git a/stage0/stdlib/Lean/Meta/Tactic/Intro.c b/stage0/stdlib/Lean/Meta/Tactic/Intro.c index d567504a3522..fad1119a4bc2 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Intro.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Intro.c @@ -13,7 +13,6 @@ #ifdef __cplusplus extern "C" { #endif -uint8_t l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(lean_object*, lean_object*); lean_object* l___private_Init_Util_0__outOfBounds___rarg(lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Intro_0__Lean_Meta_introNImp_loop___rarg___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Intro_0__Lean_Meta_mkFreshBinderNameForTacticCore___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -62,7 +61,6 @@ lean_object* l_ReaderT_instMonadLiftReaderT(lean_object*, lean_object*, lean_obj lean_object* l_Lean_MVarId_assign___rarg(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_LocalContext_mkLetDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Intro_0__Lean_Meta_introNImp___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Intro_0__Lean_Meta_getIntrosSize(lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Intro_0__Lean_Meta_mkAuxNameImp___closed__1; static lean_object* l___private_Lean_Meta_Tactic_Intro_0__Lean_Meta_introNImp_loop___rarg___closed__7; @@ -73,6 +71,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_mkFreshBinderNameForTactic(lean_object*, le uint8_t lean_name_eq(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Intro_0__Lean_Meta_introNImp___rarg___lambda__1___closed__1; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Intro_0__Lean_Meta_introNImp_loop___rarg___closed__5; static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Intro___hyg_1052____closed__4; static lean_object* l___private_Lean_Meta_Tactic_Intro_0__Lean_Meta_introNImp_loop___rarg___closed__6; @@ -80,7 +79,6 @@ extern lean_object* l_Lean_Meta_instMonadMetaM; LEAN_EXPORT lean_object* l_Lean_MVarId_introN(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_LocalContext_mkLocalDecl(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t); extern lean_object* l_Lean_instInhabitedFVarId; -lean_object* l_Lean_Option_register___at_Std_Format_initFn____x40_Lean_Data_Format___hyg_95____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Intro_0__Lean_Meta_getIntrosSize___boxed(lean_object*); LEAN_EXPORT 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_object* l_instMonadControlT__1___rarg(lean_object*); @@ -92,12 +90,14 @@ uint8_t lean_nat_dec_eq(lean_object*, lean_object*); extern lean_object* l_Lean_Core_instMonadNameGeneratorCoreM; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Intro_0__Lean_Meta_introNImp_loop___rarg___lambda__1___closed__2; +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_introN(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); uint8_t l_Lean_BinderInfo_isExplicit(uint8_t); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Intro_0__Lean_Meta_introNImp(lean_object*); lean_object* l_StateRefT_x27_lift(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); +lean_object* l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Intro___hyg_1052____closed__8; static lean_object* l___private_Lean_Meta_Tactic_Intro_0__Lean_Meta_introNImp_loop___rarg___closed__3; LEAN_EXPORT lean_object* l_Lean_MVarId_intro(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1323,7 +1323,7 @@ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Intro___hyg_1052____closed__3; x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Intro___hyg_1052____closed__5; x_4 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Intro___hyg_1052____closed__8; -x_5 = l_Lean_Option_register___at_Std_Format_initFn____x40_Lean_Data_Format___hyg_95____spec__1(x_2, x_3, x_4, x_1); +x_5 = l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__1(x_2, x_3, x_4, x_1); return x_5; } } @@ -1378,7 +1378,7 @@ x_7 = lean_ctor_get(x_2, 1); lean_inc(x_7); x_8 = lean_ctor_get(x_4, 2); x_9 = l_Lean_Meta_mkFreshBinderNameForTactic___closed__1; -x_10 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_8, x_9); +x_10 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_8, x_9); x_11 = l___private_Lean_Meta_Tactic_Intro_0__Lean_Meta_mkFreshBinderNameForTacticCore(x_7, x_1, x_10, x_2, x_3, x_4, x_5, x_6); lean_dec(x_2); return x_11; @@ -1610,7 +1610,7 @@ lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; uint8_t x x_11 = lean_ctor_get(x_8, 2); lean_inc(x_11); x_12 = l_Lean_Meta_mkFreshBinderNameForTactic___closed__1; -x_13 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_11, x_12); +x_13 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_11, x_12); lean_dec(x_11); x_14 = lean_unsigned_to_nat(0u); x_15 = lean_nat_dec_eq(x_2, x_14); @@ -2170,7 +2170,7 @@ lean_inc(x_8); x_9 = lean_ctor_get(x_7, 1); lean_inc(x_9); lean_dec(x_7); -x_10 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_8, x_2, x_3, x_4, x_5, x_9); +x_10 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_8, x_2, x_3, x_4, x_5, x_9); x_11 = !lean_is_exclusive(x_10); if (x_11 == 0) { diff --git a/stage0/stdlib/Lean/Meta/Tactic/Refl.c b/stage0/stdlib/Lean/Meta/Tactic/Refl.c index 40fc64ec4e90..ac733e0458fc 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Refl.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Refl.c @@ -59,7 +59,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Refl_0__Lean_Meta_useKerne lean_object* l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_MVarId_refl___lambda__4___closed__3; lean_object* lean_st_ref_get(lean_object*, lean_object*); -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_MVarId_refl___lambda__3___closed__6; LEAN_EXPORT lean_object* l_Lean_ofExceptKernelException___at_Lean_MVarId_refl___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_MVarId_refl___lambda__3___closed__1; @@ -84,6 +83,7 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_MVarId_refl___spec__3___box lean_object* l_Lean_Expr_constLevels_x21(lean_object*); static lean_object* l_Lean_MVarId_eqOfHEq___lambda__1___closed__1; static lean_object* l_Lean_MVarId_refl___lambda__2___closed__8; +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); lean_object* l_Lean_indentExpr(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*); @@ -607,14 +607,14 @@ lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_ lean_dec(x_3); x_9 = l_Lean_Expr_appFn_x21(x_1); x_10 = l_Lean_Expr_appArg_x21(x_9); -x_11 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_10, x_4, x_5, x_6, x_7, x_8); +x_11 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_10, x_4, x_5, x_6, x_7, x_8); 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 = l_Lean_Expr_appArg_x21(x_1); -x_15 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_14, x_4, x_5, x_6, x_7, x_13); +x_15 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_14, x_4, x_5, x_6, x_7, x_13); x_16 = lean_ctor_get(x_15, 0); lean_inc(x_16); x_17 = lean_ctor_get(x_15, 1); diff --git a/stage0/stdlib/Lean/Meta/Tactic/Rewrite.c b/stage0/stdlib/Lean/Meta/Tactic/Rewrite.c index b28e57d9a638..f21a91940a25 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Rewrite.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Rewrite.c @@ -56,7 +56,6 @@ uint8_t l_Lean_Expr_isMVar(lean_object*); lean_object* l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_to_list(lean_object*, lean_object*); lean_object* l_Lean_Meta_mkEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_MVarId_rewrite___lambda__13___closed__8; uint8_t l_Lean_Expr_hasLooseBVars(lean_object*); static lean_object* l_Lean_MVarId_rewrite___lambda__3___closed__3; @@ -87,6 +86,7 @@ LEAN_EXPORT lean_object* l_Lean_MVarId_rewrite___lambda__2(lean_object*, lean_ob lean_object* l_Lean_MVarId_isAssigned___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_rewrite___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_object*, lean_object*, lean_object*); static lean_object* l_Lean_MVarId_rewrite___lambda__13___closed__5; +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_rewrite___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_Meta_mkEqRefl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_indentExpr(lean_object*); @@ -793,7 +793,7 @@ lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean lean_dec(x_10); x_16 = lean_expr_instantiate1(x_1, x_2); lean_dec(x_2); -x_17 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_16, x_11, x_12, x_13, x_14, x_15); +x_17 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_16, x_11, x_12, x_13, x_14, x_15); x_18 = lean_ctor_get(x_17, 0); lean_inc(x_18); x_19 = lean_ctor_get(x_17, 1); @@ -951,7 +951,7 @@ LEAN_EXPORT lean_object* l_Lean_MVarId_rewrite___lambda__3(lean_object* x_1, lea { 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; lean_dec(x_12); -x_18 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_1, x_13, x_14, x_15, x_16, x_17); +x_18 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_1, x_13, x_14, x_15, x_16, x_17); x_19 = lean_ctor_get(x_13, 0); lean_inc(x_19); x_20 = lean_ctor_get(x_18, 0); @@ -1567,7 +1567,7 @@ lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean lean_dec(x_10); x_16 = lean_expr_instantiate1(x_1, x_2); lean_dec(x_2); -x_17 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_16, x_11, x_12, x_13, x_14, x_15); +x_17 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_16, x_11, x_12, x_13, x_14, x_15); x_18 = lean_ctor_get(x_17, 0); lean_inc(x_18); x_19 = lean_ctor_get(x_17, 1); @@ -1691,7 +1691,7 @@ LEAN_EXPORT lean_object* l_Lean_MVarId_rewrite___lambda__6(lean_object* x_1, lea { 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; lean_dec(x_12); -x_18 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_1, x_13, x_14, x_15, x_16, x_17); +x_18 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_1, x_13, x_14, x_15, x_16, x_17); x_19 = lean_ctor_get(x_13, 0); lean_inc(x_19); x_20 = lean_ctor_get(x_18, 0); @@ -2307,7 +2307,7 @@ lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean lean_dec(x_10); x_16 = lean_expr_instantiate1(x_1, x_2); lean_dec(x_2); -x_17 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_16, x_11, x_12, x_13, x_14, x_15); +x_17 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_16, x_11, x_12, x_13, x_14, x_15); x_18 = lean_ctor_get(x_17, 0); lean_inc(x_18); x_19 = lean_ctor_get(x_17, 1); @@ -2431,7 +2431,7 @@ LEAN_EXPORT lean_object* l_Lean_MVarId_rewrite___lambda__9(lean_object* x_1, lea { 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; lean_dec(x_12); -x_18 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_1, x_13, x_14, x_15, x_16, x_17); +x_18 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_1, x_13, x_14, x_15, x_16, x_17); x_19 = lean_ctor_get(x_13, 0); lean_inc(x_19); x_20 = lean_ctor_get(x_18, 0); @@ -3047,7 +3047,7 @@ lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean lean_dec(x_10); x_16 = lean_expr_instantiate1(x_1, x_2); lean_dec(x_2); -x_17 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_16, x_11, x_12, x_13, x_14, x_15); +x_17 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_16, x_11, x_12, x_13, x_14, x_15); x_18 = lean_ctor_get(x_17, 0); lean_inc(x_18); x_19 = lean_ctor_get(x_17, 1); @@ -3171,7 +3171,7 @@ LEAN_EXPORT lean_object* l_Lean_MVarId_rewrite___lambda__12(lean_object* x_1, le { 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; lean_dec(x_12); -x_18 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_1, x_13, x_14, x_15, x_16, x_17); +x_18 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_1, x_13, x_14, x_15, x_16, x_17); x_19 = lean_ctor_get(x_13, 0); lean_inc(x_19); x_20 = lean_ctor_get(x_18, 0); @@ -3580,7 +3580,7 @@ lean_inc(x_16); x_17 = lean_ctor_get(x_15, 1); lean_inc(x_17); lean_dec(x_15); -x_18 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_16, x_8, x_9, x_10, x_11, x_17); +x_18 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_16, x_8, x_9, x_10, x_11, x_17); x_19 = lean_ctor_get(x_18, 0); lean_inc(x_19); x_20 = lean_ctor_get(x_18, 1); diff --git a/stage0/stdlib/Lean/Meta/Tactic/Simp/Main.c b/stage0/stdlib/Lean/Meta/Tactic/Simp/Main.c index 283bc2c16a10..145c674145f7 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Simp/Main.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Simp/Main.c @@ -376,6 +376,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Simp_simp_simpForall___lambda__1(lean_objec LEAN_EXPORT lean_object* l_Lean_Meta_Simp_simp_trySimpCongrTheorem_x3f___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_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_reduceProj(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_simp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_profileitM___at_Lean_Meta_synthInstance_x3f___spec__16___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isMVar(lean_object*); lean_object* l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_simpTarget(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -420,7 +421,6 @@ static lean_object* l_Lean_Meta_Simp_simp_simpForall___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_simp_simpForall___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_Meta_Simp_simp_tryAutoCongrTheorem_x3f___lambda__5___closed__2; lean_object* l_Lean_Meta_mkEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_reduce___lambda__4(lean_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_Simp_simp_congr___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_forInUnsafe_loop___at_Lean_Meta_Simp_simp_trySimpCongrTheorem_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*); @@ -622,10 +622,10 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Sim LEAN_EXPORT lean_object* l_Lean_HashMapImp_moveEntries___at_Lean_Meta_Simp_simp_cacheResult___spec__4(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_applySimpResultToLocalDecl(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -lean_object* l_Lean_profileitM___at_Lean_Meta_synthInstance_x3f___spec__17___rarg(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_simpTargetCore___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Simp_recordSimpTheorem(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_simp_trySimpCongrTheorem_x3f___lambda__6(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_object*, lean_object*, lean_object*); +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_simp_simpArrow___lambda__3___closed__1; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_Meta_Simp_simp_simpProj___spec__4(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_simp_trySimpCongrTheorem_x3f___lambda__2___closed__2; @@ -840,6 +840,7 @@ lean_object* l_Lean_Meta_SimpTheoremsArray_addTheorem(lean_object*, lean_object* LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimp___spec__9___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_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_main___closed__2; lean_object* l_Lean_Meta_mkImpCongrCtx(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Expr_consumeMData(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_simpLocalDecl___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_uset(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_simpGoal___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -2042,7 +2043,7 @@ static lean_object* _init_l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Si 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_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_mkImpCongr___closed__1; x_2 = l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_mkImpCongr___closed__2; -x_3 = lean_unsigned_to_nat(1540u); +x_3 = lean_unsigned_to_nat(1544u); x_4 = lean_unsigned_to_nat(24u); x_5 = l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_mkImpCongr___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -41141,7 +41142,7 @@ static lean_object* _init_l_Lean_Meta_Simp_simp_simpProj___lambda__1___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___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_mkImpCongr___closed__1; x_2 = l_Lean_Meta_Simp_simp_simpProj___lambda__1___closed__1; -x_3 = lean_unsigned_to_nat(1514u); +x_3 = lean_unsigned_to_nat(1518u); x_4 = lean_unsigned_to_nat(18u); x_5 = l_Lean_Meta_Simp_simp_simpProj___lambda__1___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -44923,10 +44924,12 @@ goto _start; } else { -lean_object* x_12; uint8_t x_13; -x_12 = l_Lean_Meta_Simp_isEqnThmHypothesis_go___closed__2; -x_13 = l_Lean_Expr_isConstOf(x_1, x_12); -return x_13; +lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_12 = l_Lean_Expr_consumeMData(x_1); +x_13 = l_Lean_Meta_Simp_isEqnThmHypothesis_go___closed__2; +x_14 = l_Lean_Expr_isConstOf(x_12, x_13); +lean_dec(x_12); +return x_14; } } } @@ -46724,517 +46727,523 @@ uint8_t x_22; 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; uint8_t x_27; +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_ctor_get(x_21, 0); x_24 = lean_ctor_get(x_21, 1); x_25 = lean_ctor_get(x_23, 0); lean_inc(x_25); -x_26 = l_Lean_Meta_Simp_DefaultMethods_discharge_x3f___lambda__1___closed__6; -x_27 = l_Lean_Expr_isConstOf(x_25, x_26); +x_26 = l_Lean_Expr_consumeMData(x_25); lean_dec(x_25); -if (x_27 == 0) +x_27 = l_Lean_Meta_Simp_DefaultMethods_discharge_x3f___lambda__1___closed__6; +x_28 = l_Lean_Expr_isConstOf(x_26, x_27); +lean_dec(x_26); +if (x_28 == 0) { -lean_object* x_28; +lean_object* x_29; lean_dec(x_23); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -x_28 = lean_box(0); -lean_ctor_set(x_21, 0, x_28); +x_29 = lean_box(0); +lean_ctor_set(x_21, 0, x_29); return x_21; } else { -lean_object* x_29; +lean_object* x_30; lean_free_object(x_21); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); -x_29 = l_Lean_Meta_Simp_Result_getProof(x_23, x_7, x_8, x_9, x_10, x_24); -if (lean_obj_tag(x_29) == 0) +x_30 = l_Lean_Meta_Simp_Result_getProof(x_23, x_7, x_8, x_9, x_10, x_24); +if (lean_obj_tag(x_30) == 0) { -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_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_30, 0); lean_inc(x_31); -lean_dec(x_29); -x_32 = l_Lean_Meta_mkOfEqTrue(x_30, x_7, x_8, x_9, x_10, x_31); -if (lean_obj_tag(x_32) == 0) +x_32 = lean_ctor_get(x_30, 1); +lean_inc(x_32); +lean_dec(x_30); +x_33 = l_Lean_Meta_mkOfEqTrue(x_31, x_7, x_8, x_9, x_10, x_32); +if (lean_obj_tag(x_33) == 0) { -uint8_t x_33; -x_33 = !lean_is_exclusive(x_32); -if (x_33 == 0) +uint8_t x_34; +x_34 = !lean_is_exclusive(x_33); +if (x_34 == 0) { -lean_object* x_34; lean_object* x_35; -x_34 = lean_ctor_get(x_32, 0); -x_35 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_35, 0, x_34); -lean_ctor_set(x_32, 0, x_35); -return x_32; +lean_object* x_35; lean_object* x_36; +x_35 = lean_ctor_get(x_33, 0); +x_36 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_33, 0, x_36); +return x_33; } 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_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_37 = lean_ctor_get(x_33, 0); +x_38 = lean_ctor_get(x_33, 1); +lean_inc(x_38); lean_inc(x_37); -lean_inc(x_36); -lean_dec(x_32); -x_38 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_38, 0, 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_dec(x_33); +x_39 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_39, 0, 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_38); +return x_40; } } else { -uint8_t x_40; -x_40 = !lean_is_exclusive(x_32); -if (x_40 == 0) +uint8_t x_41; +x_41 = !lean_is_exclusive(x_33); +if (x_41 == 0) { -lean_object* x_41; lean_object* x_42; -x_41 = lean_ctor_get(x_32, 0); -lean_dec(x_41); -x_42 = lean_box(0); -lean_ctor_set_tag(x_32, 0); -lean_ctor_set(x_32, 0, x_42); -return x_32; +lean_object* x_42; lean_object* x_43; +x_42 = lean_ctor_get(x_33, 0); +lean_dec(x_42); +x_43 = lean_box(0); +lean_ctor_set_tag(x_33, 0); +lean_ctor_set(x_33, 0, x_43); +return x_33; } else { -lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_43 = lean_ctor_get(x_32, 1); -lean_inc(x_43); -lean_dec(x_32); -x_44 = lean_box(0); -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_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_33, 1); +lean_inc(x_44); +lean_dec(x_33); +x_45 = lean_box(0); +x_46 = lean_alloc_ctor(0, 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_46; +uint8_t x_47; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -x_46 = !lean_is_exclusive(x_29); -if (x_46 == 0) +x_47 = !lean_is_exclusive(x_30); +if (x_47 == 0) { -lean_object* x_47; lean_object* x_48; -x_47 = lean_ctor_get(x_29, 0); -lean_dec(x_47); -x_48 = lean_box(0); -lean_ctor_set_tag(x_29, 0); -lean_ctor_set(x_29, 0, x_48); -return x_29; +lean_object* x_48; lean_object* x_49; +x_48 = lean_ctor_get(x_30, 0); +lean_dec(x_48); +x_49 = lean_box(0); +lean_ctor_set_tag(x_30, 0); +lean_ctor_set(x_30, 0, x_49); +return x_30; } else { -lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_49 = lean_ctor_get(x_29, 1); -lean_inc(x_49); -lean_dec(x_29); -x_50 = lean_box(0); -x_51 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_51, 0, x_50); -lean_ctor_set(x_51, 1, x_49); -return x_51; +lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_50 = lean_ctor_get(x_30, 1); +lean_inc(x_50); +lean_dec(x_30); +x_51 = lean_box(0); +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 { -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_21, 0); -x_53 = lean_ctor_get(x_21, 1); +lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; uint8_t x_58; +x_53 = lean_ctor_get(x_21, 0); +x_54 = lean_ctor_get(x_21, 1); +lean_inc(x_54); lean_inc(x_53); -lean_inc(x_52); lean_dec(x_21); -x_54 = lean_ctor_get(x_52, 0); -lean_inc(x_54); -x_55 = l_Lean_Meta_Simp_DefaultMethods_discharge_x3f___lambda__1___closed__6; -x_56 = l_Lean_Expr_isConstOf(x_54, x_55); -lean_dec(x_54); -if (x_56 == 0) +x_55 = lean_ctor_get(x_53, 0); +lean_inc(x_55); +x_56 = l_Lean_Expr_consumeMData(x_55); +lean_dec(x_55); +x_57 = l_Lean_Meta_Simp_DefaultMethods_discharge_x3f___lambda__1___closed__6; +x_58 = l_Lean_Expr_isConstOf(x_56, x_57); +lean_dec(x_56); +if (x_58 == 0) { -lean_object* x_57; lean_object* x_58; -lean_dec(x_52); +lean_object* x_59; lean_object* x_60; +lean_dec(x_53); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -x_57 = lean_box(0); -x_58 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_58, 0, x_57); -lean_ctor_set(x_58, 1, x_53); -return x_58; +x_59 = lean_box(0); +x_60 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_60, 0, x_59); +lean_ctor_set(x_60, 1, x_54); +return x_60; } else { -lean_object* x_59; +lean_object* x_61; lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); -x_59 = l_Lean_Meta_Simp_Result_getProof(x_52, x_7, x_8, x_9, x_10, x_53); -if (lean_obj_tag(x_59) == 0) -{ -lean_object* x_60; lean_object* x_61; lean_object* x_62; -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 = l_Lean_Meta_mkOfEqTrue(x_60, x_7, x_8, x_9, x_10, x_61); -if (lean_obj_tag(x_62) == 0) +x_61 = l_Lean_Meta_Simp_Result_getProof(x_53, x_7, x_8, x_9, x_10, x_54); +if (lean_obj_tag(x_61) == 0) { -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, 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); -if (lean_is_exclusive(x_62)) { - lean_ctor_release(x_62, 0); - lean_ctor_release(x_62, 1); - x_65 = x_62; +lean_dec(x_61); +x_64 = l_Lean_Meta_mkOfEqTrue(x_62, x_7, x_8, x_9, x_10, x_63); +if (lean_obj_tag(x_64) == 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, 0); +lean_inc(x_65); +x_66 = lean_ctor_get(x_64, 1); +lean_inc(x_66); +if (lean_is_exclusive(x_64)) { + lean_ctor_release(x_64, 0); + lean_ctor_release(x_64, 1); + x_67 = x_64; } else { - lean_dec_ref(x_62); - x_65 = lean_box(0); + lean_dec_ref(x_64); + x_67 = lean_box(0); } -x_66 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_66, 0, x_63); -if (lean_is_scalar(x_65)) { - x_67 = lean_alloc_ctor(0, 2, 0); +x_68 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_68, 0, x_65); +if (lean_is_scalar(x_67)) { + x_69 = lean_alloc_ctor(0, 2, 0); } else { - x_67 = x_65; + x_69 = x_67; } -lean_ctor_set(x_67, 0, x_66); -lean_ctor_set(x_67, 1, x_64); -return x_67; +lean_ctor_set(x_69, 0, x_68); +lean_ctor_set(x_69, 1, x_66); +return x_69; } else { -lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; -x_68 = lean_ctor_get(x_62, 1); -lean_inc(x_68); -if (lean_is_exclusive(x_62)) { - lean_ctor_release(x_62, 0); - lean_ctor_release(x_62, 1); - x_69 = x_62; +lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; +x_70 = lean_ctor_get(x_64, 1); +lean_inc(x_70); +if (lean_is_exclusive(x_64)) { + lean_ctor_release(x_64, 0); + lean_ctor_release(x_64, 1); + x_71 = x_64; } else { - lean_dec_ref(x_62); - x_69 = lean_box(0); + lean_dec_ref(x_64); + x_71 = lean_box(0); } -x_70 = lean_box(0); -if (lean_is_scalar(x_69)) { - x_71 = lean_alloc_ctor(0, 2, 0); +x_72 = lean_box(0); +if (lean_is_scalar(x_71)) { + x_73 = lean_alloc_ctor(0, 2, 0); } else { - x_71 = x_69; - lean_ctor_set_tag(x_71, 0); + x_73 = x_71; + lean_ctor_set_tag(x_73, 0); } -lean_ctor_set(x_71, 0, x_70); -lean_ctor_set(x_71, 1, x_68); -return x_71; +lean_ctor_set(x_73, 0, x_72); +lean_ctor_set(x_73, 1, x_70); +return x_73; } } else { -lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; +lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -x_72 = lean_ctor_get(x_59, 1); -lean_inc(x_72); -if (lean_is_exclusive(x_59)) { - lean_ctor_release(x_59, 0); - lean_ctor_release(x_59, 1); - x_73 = x_59; +x_74 = lean_ctor_get(x_61, 1); +lean_inc(x_74); +if (lean_is_exclusive(x_61)) { + lean_ctor_release(x_61, 0); + lean_ctor_release(x_61, 1); + x_75 = x_61; } else { - lean_dec_ref(x_59); - x_73 = lean_box(0); + lean_dec_ref(x_61); + x_75 = lean_box(0); } -x_74 = lean_box(0); -if (lean_is_scalar(x_73)) { - x_75 = lean_alloc_ctor(0, 2, 0); +x_76 = lean_box(0); +if (lean_is_scalar(x_75)) { + x_77 = lean_alloc_ctor(0, 2, 0); } else { - x_75 = x_73; - lean_ctor_set_tag(x_75, 0); + x_77 = x_75; + lean_ctor_set_tag(x_77, 0); } -lean_ctor_set(x_75, 0, x_74); -lean_ctor_set(x_75, 1, x_72); -return x_75; +lean_ctor_set(x_77, 0, x_76); +lean_ctor_set(x_77, 1, x_74); +return x_77; } } } } else { -uint8_t x_76; +uint8_t x_78; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -x_76 = !lean_is_exclusive(x_21); -if (x_76 == 0) +x_78 = !lean_is_exclusive(x_21); +if (x_78 == 0) { return x_21; } else { -lean_object* x_77; lean_object* x_78; lean_object* x_79; -x_77 = lean_ctor_get(x_21, 0); -x_78 = lean_ctor_get(x_21, 1); -lean_inc(x_78); -lean_inc(x_77); +lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_79 = lean_ctor_get(x_21, 0); +x_80 = lean_ctor_get(x_21, 1); +lean_inc(x_80); +lean_inc(x_79); lean_dec(x_21); -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; +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_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 = lean_ctor_get(x_5, 0); -x_81 = lean_ctor_get(x_5, 1); -x_82 = lean_ctor_get(x_5, 2); -x_83 = lean_ctor_get(x_5, 3); -x_84 = lean_ctor_get(x_5, 4); +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_82 = lean_ctor_get(x_5, 0); +x_83 = lean_ctor_get(x_5, 1); +x_84 = lean_ctor_get(x_5, 2); +x_85 = lean_ctor_get(x_5, 3); +x_86 = lean_ctor_get(x_5, 4); +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_inc(x_80); lean_dec(x_5); -x_85 = lean_unsigned_to_nat(1u); -x_86 = lean_nat_add(x_84, x_85); -lean_dec(x_84); -x_87 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_87, 0, x_80); -lean_ctor_set(x_87, 1, x_81); -lean_ctor_set(x_87, 2, x_82); -lean_ctor_set(x_87, 3, x_83); -lean_ctor_set(x_87, 4, x_86); -x_88 = l_Lean_Meta_Simp_DefaultMethods_discharge_x3f___lambda__1___closed__4; +x_87 = lean_unsigned_to_nat(1u); +x_88 = lean_nat_add(x_86, x_87); +lean_dec(x_86); +x_89 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_89, 0, x_82); +lean_ctor_set(x_89, 1, x_83); +lean_ctor_set(x_89, 2, x_84); +lean_ctor_set(x_89, 3, x_85); +lean_ctor_set(x_89, 4, x_88); +x_90 = l_Lean_Meta_Simp_DefaultMethods_discharge_x3f___lambda__1___closed__4; lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); -x_89 = l_Lean_Meta_Simp_simp(x_2, x_88, x_87, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_89) == 0) +x_91 = l_Lean_Meta_Simp_simp(x_2, x_90, x_89, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_91) == 0) { -lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; uint8_t x_95; -x_90 = lean_ctor_get(x_89, 0); -lean_inc(x_90); -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; +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_92 = lean_ctor_get(x_91, 0); +lean_inc(x_92); +x_93 = lean_ctor_get(x_91, 1); +lean_inc(x_93); +if (lean_is_exclusive(x_91)) { + lean_ctor_release(x_91, 0); + lean_ctor_release(x_91, 1); + x_94 = x_91; } else { - lean_dec_ref(x_89); - x_92 = lean_box(0); + lean_dec_ref(x_91); + x_94 = lean_box(0); } -x_93 = lean_ctor_get(x_90, 0); -lean_inc(x_93); -x_94 = l_Lean_Meta_Simp_DefaultMethods_discharge_x3f___lambda__1___closed__6; -x_95 = l_Lean_Expr_isConstOf(x_93, x_94); -lean_dec(x_93); -if (x_95 == 0) +x_95 = lean_ctor_get(x_92, 0); +lean_inc(x_95); +x_96 = l_Lean_Expr_consumeMData(x_95); +lean_dec(x_95); +x_97 = l_Lean_Meta_Simp_DefaultMethods_discharge_x3f___lambda__1___closed__6; +x_98 = l_Lean_Expr_isConstOf(x_96, x_97); +lean_dec(x_96); +if (x_98 == 0) { -lean_object* x_96; lean_object* x_97; -lean_dec(x_90); +lean_object* x_99; lean_object* x_100; +lean_dec(x_92); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -x_96 = lean_box(0); -if (lean_is_scalar(x_92)) { - x_97 = lean_alloc_ctor(0, 2, 0); +x_99 = lean_box(0); +if (lean_is_scalar(x_94)) { + x_100 = lean_alloc_ctor(0, 2, 0); } else { - x_97 = x_92; + x_100 = x_94; } -lean_ctor_set(x_97, 0, x_96); -lean_ctor_set(x_97, 1, x_91); -return x_97; +lean_ctor_set(x_100, 0, x_99); +lean_ctor_set(x_100, 1, x_93); +return x_100; } else { -lean_object* x_98; -lean_dec(x_92); +lean_object* x_101; +lean_dec(x_94); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); -x_98 = l_Lean_Meta_Simp_Result_getProof(x_90, x_7, x_8, x_9, x_10, x_91); -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, 0); -lean_inc(x_99); -x_100 = lean_ctor_get(x_98, 1); -lean_inc(x_100); -lean_dec(x_98); -x_101 = l_Lean_Meta_mkOfEqTrue(x_99, x_7, x_8, x_9, x_10, x_100); +x_101 = l_Lean_Meta_Simp_Result_getProof(x_92, x_7, x_8, x_9, x_10, x_93); 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, 0); lean_inc(x_102); x_103 = lean_ctor_get(x_101, 1); lean_inc(x_103); -if (lean_is_exclusive(x_101)) { - lean_ctor_release(x_101, 0); - lean_ctor_release(x_101, 1); - x_104 = x_101; +lean_dec(x_101); +x_104 = l_Lean_Meta_mkOfEqTrue(x_102, x_7, x_8, x_9, x_10, x_103); +if (lean_obj_tag(x_104) == 0) +{ +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, 0); +lean_inc(x_105); +x_106 = lean_ctor_get(x_104, 1); +lean_inc(x_106); +if (lean_is_exclusive(x_104)) { + lean_ctor_release(x_104, 0); + lean_ctor_release(x_104, 1); + x_107 = x_104; } else { - lean_dec_ref(x_101); - x_104 = lean_box(0); + lean_dec_ref(x_104); + x_107 = lean_box(0); } -x_105 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_105, 0, x_102); -if (lean_is_scalar(x_104)) { - x_106 = lean_alloc_ctor(0, 2, 0); +x_108 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_108, 0, x_105); +if (lean_is_scalar(x_107)) { + x_109 = lean_alloc_ctor(0, 2, 0); } else { - x_106 = x_104; + x_109 = x_107; } -lean_ctor_set(x_106, 0, x_105); -lean_ctor_set(x_106, 1, x_103); -return x_106; +lean_ctor_set(x_109, 0, x_108); +lean_ctor_set(x_109, 1, x_106); +return x_109; } else { -lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; -x_107 = lean_ctor_get(x_101, 1); -lean_inc(x_107); -if (lean_is_exclusive(x_101)) { - lean_ctor_release(x_101, 0); - lean_ctor_release(x_101, 1); - x_108 = x_101; +lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; +x_110 = lean_ctor_get(x_104, 1); +lean_inc(x_110); +if (lean_is_exclusive(x_104)) { + lean_ctor_release(x_104, 0); + lean_ctor_release(x_104, 1); + x_111 = x_104; } else { - lean_dec_ref(x_101); - x_108 = lean_box(0); + lean_dec_ref(x_104); + x_111 = lean_box(0); } -x_109 = lean_box(0); -if (lean_is_scalar(x_108)) { - x_110 = lean_alloc_ctor(0, 2, 0); +x_112 = lean_box(0); +if (lean_is_scalar(x_111)) { + x_113 = lean_alloc_ctor(0, 2, 0); } else { - x_110 = x_108; - lean_ctor_set_tag(x_110, 0); + x_113 = x_111; + lean_ctor_set_tag(x_113, 0); } -lean_ctor_set(x_110, 0, x_109); -lean_ctor_set(x_110, 1, x_107); -return x_110; +lean_ctor_set(x_113, 0, x_112); +lean_ctor_set(x_113, 1, x_110); +return x_113; } } else { -lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; +lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -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; +x_114 = lean_ctor_get(x_101, 1); +lean_inc(x_114); +if (lean_is_exclusive(x_101)) { + lean_ctor_release(x_101, 0); + lean_ctor_release(x_101, 1); + x_115 = x_101; } else { - lean_dec_ref(x_98); - x_112 = lean_box(0); + lean_dec_ref(x_101); + x_115 = lean_box(0); } -x_113 = lean_box(0); -if (lean_is_scalar(x_112)) { - x_114 = lean_alloc_ctor(0, 2, 0); +x_116 = lean_box(0); +if (lean_is_scalar(x_115)) { + x_117 = lean_alloc_ctor(0, 2, 0); } else { - x_114 = x_112; - lean_ctor_set_tag(x_114, 0); + x_117 = x_115; + lean_ctor_set_tag(x_117, 0); } -lean_ctor_set(x_114, 0, x_113); -lean_ctor_set(x_114, 1, x_111); -return x_114; +lean_ctor_set(x_117, 0, x_116); +lean_ctor_set(x_117, 1, x_114); +return x_117; } } } else { -lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; +lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -x_115 = lean_ctor_get(x_89, 0); -lean_inc(x_115); -x_116 = lean_ctor_get(x_89, 1); -lean_inc(x_116); -if (lean_is_exclusive(x_89)) { - lean_ctor_release(x_89, 0); - lean_ctor_release(x_89, 1); - x_117 = x_89; +x_118 = lean_ctor_get(x_91, 0); +lean_inc(x_118); +x_119 = lean_ctor_get(x_91, 1); +lean_inc(x_119); +if (lean_is_exclusive(x_91)) { + lean_ctor_release(x_91, 0); + lean_ctor_release(x_91, 1); + x_120 = x_91; } else { - lean_dec_ref(x_89); - x_117 = lean_box(0); + lean_dec_ref(x_91); + x_120 = lean_box(0); } -if (lean_is_scalar(x_117)) { - x_118 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_120)) { + x_121 = lean_alloc_ctor(1, 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); -return x_118; +lean_ctor_set(x_121, 0, x_118); +lean_ctor_set(x_121, 1, x_119); +return x_121; } } } else { -lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; uint8_t x_123; +lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; uint8_t x_126; lean_dec(x_2); lean_inc(x_3); -x_119 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Simp_synthesizeArgs_synthesizeInstance___spec__1(x_3, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -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 = l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimp___lambda__3___closed__1; -x_123 = lean_unbox(x_120); -lean_dec(x_120); -if (x_123 == 0) +x_122 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Simp_synthesizeArgs_synthesizeInstance___spec__1(x_3, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +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 = l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimp___lambda__3___closed__1; +x_126 = lean_unbox(x_123); +lean_dec(x_123); +if (x_126 == 0) { -lean_object* x_124; lean_object* x_125; +lean_object* x_127; lean_object* x_128; lean_dec(x_3); -x_124 = lean_box(0); -x_125 = lean_apply_8(x_122, x_124, x_5, x_6, x_7, x_8, x_9, x_10, x_121); -return x_125; +x_127 = lean_box(0); +x_128 = lean_apply_8(x_125, x_127, x_5, x_6, x_7, x_8, x_9, x_10, x_124); +return x_128; } else { -lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; -x_126 = l_Lean_Meta_Simp_DefaultMethods_discharge_x3f___lambda__1___closed__8; -x_127 = l_Lean_addTrace___at_Lean_Meta_Simp_synthesizeArgs_synthesizeInstance___spec__6(x_3, x_126, x_5, x_6, x_7, x_8, x_9, x_10, x_121); -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_apply_8(x_122, x_128, x_5, x_6, x_7, x_8, x_9, x_10, x_129); -return x_130; +lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; +x_129 = l_Lean_Meta_Simp_DefaultMethods_discharge_x3f___lambda__1___closed__8; +x_130 = l_Lean_addTrace___at_Lean_Meta_Simp_synthesizeArgs_synthesizeInstance___spec__6(x_3, x_129, x_5, x_6, x_7, x_8, x_9, x_10, x_124); +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_130); +x_133 = lean_apply_8(x_125, x_131, x_5, x_6, x_7, x_8, x_9, x_10, x_132); +return x_133; } } } @@ -47659,7 +47668,7 @@ lean_closure_set(x_12, 2, x_4); lean_closure_set(x_12, 3, x_11); x_13 = l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_reduce___lambda__2___closed__3; x_14 = lean_box(0); -x_15 = l_Lean_profileitM___at_Lean_Meta_synthInstance_x3f___spec__17___rarg(x_13, x_10, x_12, x_14, x_5, x_6, x_7, x_8, x_9); +x_15 = l_Lean_profileitM___at_Lean_Meta_synthInstance_x3f___spec__16___rarg(x_13, x_10, x_12, x_14, x_5, x_6, x_7, x_8, x_9); lean_dec(x_10); return x_15; } @@ -47688,7 +47697,7 @@ lean_closure_set(x_21, 2, x_4); lean_closure_set(x_21, 3, x_20); x_22 = l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_reduce___lambda__2___closed__3; x_23 = lean_box(0); -x_24 = l_Lean_profileitM___at_Lean_Meta_synthInstance_x3f___spec__17___rarg(x_22, x_16, x_21, x_23, x_5, x_6, x_7, x_8, x_9); +x_24 = l_Lean_profileitM___at_Lean_Meta_synthInstance_x3f___spec__16___rarg(x_22, x_16, x_21, x_23, x_5, x_6, x_7, x_8, x_9); lean_dec(x_16); return x_24; } @@ -47708,7 +47717,7 @@ lean_closure_set(x_11, 2, x_3); lean_closure_set(x_11, 3, x_10); x_12 = l_Lean_Meta_Simp_dsimpMain___closed__1; x_13 = lean_box(0); -x_14 = l_Lean_profileitM___at_Lean_Meta_synthInstance_x3f___spec__17___rarg(x_12, x_9, x_11, x_13, x_4, x_5, x_6, x_7, x_8); +x_14 = l_Lean_profileitM___at_Lean_Meta_synthInstance_x3f___spec__16___rarg(x_12, x_9, x_11, x_13, x_4, x_5, x_6, x_7, x_8); lean_dec(x_9); return x_14; } @@ -47825,7 +47834,7 @@ lean_inc(x_12); x_13 = lean_ctor_get(x_11, 1); lean_inc(x_13); lean_dec(x_11); -x_14 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_12, x_6, x_7, x_8, x_9, x_13); +x_14 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_12, x_6, x_7, x_8, x_9, x_13); x_15 = lean_ctor_get(x_14, 0); lean_inc(x_15); x_16 = lean_ctor_get(x_14, 1); @@ -47987,160 +47996,162 @@ lean_dec(x_17); x_49 = !lean_is_exclusive(x_18); if (x_49 == 0) { -lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; uint8_t x_54; +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_50 = lean_ctor_get(x_18, 0); x_51 = lean_ctor_get(x_18, 1); x_52 = lean_ctor_get(x_50, 0); lean_inc(x_52); -x_53 = l_Lean_Meta_Simp_DefaultMethods_discharge_x3f___lambda__1___closed__6; -x_54 = l_Lean_Expr_isConstOf(x_52, x_53); +x_53 = l_Lean_Expr_consumeMData(x_52); lean_dec(x_52); -if (x_54 == 0) +x_54 = l_Lean_Meta_Simp_DefaultMethods_discharge_x3f___lambda__1___closed__6; +x_55 = l_Lean_Expr_isConstOf(x_53, x_54); +lean_dec(x_53); +if (x_55 == 0) { -lean_object* x_55; -x_55 = l_Lean_Meta_applySimpResultToTarget(x_1, x_15, x_50, x_6, x_7, x_8, x_9, x_48); +lean_object* x_56; +x_56 = l_Lean_Meta_applySimpResultToTarget(x_1, x_15, x_50, x_6, x_7, x_8, x_9, x_48); lean_dec(x_15); -if (lean_obj_tag(x_55) == 0) +if (lean_obj_tag(x_56) == 0) { -uint8_t x_56; -x_56 = !lean_is_exclusive(x_55); -if (x_56 == 0) +uint8_t x_57; +x_57 = !lean_is_exclusive(x_56); +if (x_57 == 0) { -lean_object* x_57; lean_object* x_58; -x_57 = lean_ctor_get(x_55, 0); -x_58 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_58, 0, x_57); -lean_ctor_set(x_18, 0, x_58); -lean_ctor_set(x_55, 0, x_18); -return x_55; +lean_object* x_58; lean_object* x_59; +x_58 = lean_ctor_get(x_56, 0); +x_59 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_59, 0, x_58); +lean_ctor_set(x_18, 0, x_59); +lean_ctor_set(x_56, 0, x_18); +return x_56; } 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_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; +x_60 = lean_ctor_get(x_56, 0); +x_61 = lean_ctor_get(x_56, 1); +lean_inc(x_61); lean_inc(x_60); -lean_inc(x_59); -lean_dec(x_55); -x_61 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_61, 0, x_59); -lean_ctor_set(x_18, 0, x_61); -x_62 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_62, 0, x_18); -lean_ctor_set(x_62, 1, x_60); -return x_62; +lean_dec(x_56); +x_62 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_62, 0, x_60); +lean_ctor_set(x_18, 0, x_62); +x_63 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_63, 0, x_18); +lean_ctor_set(x_63, 1, x_61); +return x_63; } } else { -uint8_t x_63; +uint8_t x_64; lean_free_object(x_18); lean_dec(x_51); -x_63 = !lean_is_exclusive(x_55); -if (x_63 == 0) +x_64 = !lean_is_exclusive(x_56); +if (x_64 == 0) { -return x_55; +return x_56; } 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_object* x_65; lean_object* x_66; lean_object* x_67; +x_65 = lean_ctor_get(x_56, 0); +x_66 = lean_ctor_get(x_56, 1); +lean_inc(x_66); lean_inc(x_65); -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_dec(x_56); +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_67; +lean_object* x_68; lean_free_object(x_18); lean_dec(x_15); -x_67 = lean_ctor_get(x_50, 1); -lean_inc(x_67); +x_68 = lean_ctor_get(x_50, 1); +lean_inc(x_68); lean_dec(x_50); -if (lean_obj_tag(x_67) == 0) +if (lean_obj_tag(x_68) == 0) { -lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; -x_68 = l_Lean_Meta_simpTargetCore___closed__3; -x_69 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_1, x_68, x_6, x_7, x_8, x_9, x_48); -x_70 = lean_ctor_get(x_69, 0); -lean_inc(x_70); -x_71 = lean_ctor_get(x_69, 1); +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_simpTargetCore___closed__3; +x_70 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_1, x_69, x_6, x_7, x_8, x_9, x_48); +x_71 = lean_ctor_get(x_70, 0); lean_inc(x_71); -lean_dec(x_69); -x_72 = l_Lean_Meta_simpTargetCore___lambda__1(x_51, x_70, x_6, x_7, x_8, x_9, x_71); +x_72 = lean_ctor_get(x_70, 1); +lean_inc(x_72); +lean_dec(x_70); +x_73 = l_Lean_Meta_simpTargetCore___lambda__1(x_51, x_71, x_6, x_7, x_8, x_9, x_72); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_70); -return x_72; +lean_dec(x_71); +return x_73; } else { -lean_object* x_73; lean_object* x_74; -x_73 = lean_ctor_get(x_67, 0); -lean_inc(x_73); -lean_dec(x_67); +lean_object* x_74; lean_object* x_75; +x_74 = lean_ctor_get(x_68, 0); +lean_inc(x_74); +lean_dec(x_68); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_74 = l_Lean_Meta_mkOfEqTrue(x_73, x_6, x_7, x_8, x_9, x_48); -if (lean_obj_tag(x_74) == 0) +x_75 = l_Lean_Meta_mkOfEqTrue(x_74, x_6, x_7, x_8, x_9, x_48); +if (lean_obj_tag(x_75) == 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; -x_75 = lean_ctor_get(x_74, 0); -lean_inc(x_75); -x_76 = lean_ctor_get(x_74, 1); +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_76 = lean_ctor_get(x_75, 0); lean_inc(x_76); -lean_dec(x_74); -x_77 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_1, x_75, x_6, x_7, x_8, x_9, x_76); -x_78 = lean_ctor_get(x_77, 0); -lean_inc(x_78); -x_79 = lean_ctor_get(x_77, 1); +x_77 = lean_ctor_get(x_75, 1); +lean_inc(x_77); +lean_dec(x_75); +x_78 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_1, x_76, x_6, x_7, x_8, x_9, x_77); +x_79 = lean_ctor_get(x_78, 0); lean_inc(x_79); -lean_dec(x_77); -x_80 = l_Lean_Meta_simpTargetCore___lambda__1(x_51, x_78, x_6, x_7, x_8, x_9, x_79); +x_80 = lean_ctor_get(x_78, 1); +lean_inc(x_80); +lean_dec(x_78); +x_81 = l_Lean_Meta_simpTargetCore___lambda__1(x_51, x_79, x_6, x_7, x_8, x_9, x_80); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_78); -return x_80; +lean_dec(x_79); +return x_81; } else { -uint8_t x_81; +uint8_t x_82; lean_dec(x_51); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_1); -x_81 = !lean_is_exclusive(x_74); -if (x_81 == 0) +x_82 = !lean_is_exclusive(x_75); +if (x_82 == 0) { -return x_74; +return x_75; } else { -lean_object* x_82; lean_object* x_83; lean_object* x_84; -x_82 = lean_ctor_get(x_74, 0); -x_83 = lean_ctor_get(x_74, 1); +lean_object* x_83; lean_object* x_84; lean_object* x_85; +x_83 = lean_ctor_get(x_75, 0); +x_84 = lean_ctor_get(x_75, 1); +lean_inc(x_84); lean_inc(x_83); -lean_inc(x_82); -lean_dec(x_74); -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_dec(x_75); +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; } } } @@ -48148,164 +48159,166 @@ return x_84; } else { -lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; uint8_t x_89; -x_85 = lean_ctor_get(x_18, 0); -x_86 = lean_ctor_get(x_18, 1); +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_86 = lean_ctor_get(x_18, 0); +x_87 = lean_ctor_get(x_18, 1); +lean_inc(x_87); lean_inc(x_86); -lean_inc(x_85); lean_dec(x_18); -x_87 = lean_ctor_get(x_85, 0); -lean_inc(x_87); -x_88 = l_Lean_Meta_Simp_DefaultMethods_discharge_x3f___lambda__1___closed__6; -x_89 = l_Lean_Expr_isConstOf(x_87, x_88); -lean_dec(x_87); -if (x_89 == 0) +x_88 = lean_ctor_get(x_86, 0); +lean_inc(x_88); +x_89 = l_Lean_Expr_consumeMData(x_88); +lean_dec(x_88); +x_90 = l_Lean_Meta_Simp_DefaultMethods_discharge_x3f___lambda__1___closed__6; +x_91 = l_Lean_Expr_isConstOf(x_89, x_90); +lean_dec(x_89); +if (x_91 == 0) { -lean_object* x_90; -x_90 = l_Lean_Meta_applySimpResultToTarget(x_1, x_15, x_85, x_6, x_7, x_8, x_9, x_48); +lean_object* x_92; +x_92 = l_Lean_Meta_applySimpResultToTarget(x_1, x_15, x_86, x_6, x_7, x_8, x_9, x_48); lean_dec(x_15); -if (lean_obj_tag(x_90) == 0) +if (lean_obj_tag(x_92) == 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; -x_91 = lean_ctor_get(x_90, 0); -lean_inc(x_91); -x_92 = lean_ctor_get(x_90, 1); -lean_inc(x_92); -if (lean_is_exclusive(x_90)) { - lean_ctor_release(x_90, 0); - lean_ctor_release(x_90, 1); - x_93 = x_90; +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_93 = lean_ctor_get(x_92, 0); +lean_inc(x_93); +x_94 = lean_ctor_get(x_92, 1); +lean_inc(x_94); +if (lean_is_exclusive(x_92)) { + lean_ctor_release(x_92, 0); + lean_ctor_release(x_92, 1); + x_95 = x_92; } else { - lean_dec_ref(x_90); - x_93 = lean_box(0); + lean_dec_ref(x_92); + x_95 = lean_box(0); } -x_94 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_94, 0, x_91); -x_95 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_95, 0, x_94); -lean_ctor_set(x_95, 1, x_86); -if (lean_is_scalar(x_93)) { - x_96 = lean_alloc_ctor(0, 2, 0); +x_96 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_96, 0, x_93); +x_97 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_97, 0, x_96); +lean_ctor_set(x_97, 1, x_87); +if (lean_is_scalar(x_95)) { + x_98 = lean_alloc_ctor(0, 2, 0); } else { - x_96 = x_93; + x_98 = x_95; } -lean_ctor_set(x_96, 0, x_95); -lean_ctor_set(x_96, 1, x_92); -return x_96; +lean_ctor_set(x_98, 0, x_97); +lean_ctor_set(x_98, 1, x_94); +return x_98; } else { -lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; -lean_dec(x_86); -x_97 = lean_ctor_get(x_90, 0); -lean_inc(x_97); -x_98 = lean_ctor_get(x_90, 1); -lean_inc(x_98); -if (lean_is_exclusive(x_90)) { - lean_ctor_release(x_90, 0); - lean_ctor_release(x_90, 1); - x_99 = x_90; +lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; +lean_dec(x_87); +x_99 = lean_ctor_get(x_92, 0); +lean_inc(x_99); +x_100 = lean_ctor_get(x_92, 1); +lean_inc(x_100); +if (lean_is_exclusive(x_92)) { + lean_ctor_release(x_92, 0); + lean_ctor_release(x_92, 1); + x_101 = x_92; } else { - lean_dec_ref(x_90); - x_99 = lean_box(0); + lean_dec_ref(x_92); + x_101 = lean_box(0); } -if (lean_is_scalar(x_99)) { - x_100 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_101)) { + x_102 = lean_alloc_ctor(1, 2, 0); } else { - x_100 = x_99; + x_102 = x_101; } -lean_ctor_set(x_100, 0, x_97); -lean_ctor_set(x_100, 1, x_98); -return x_100; +lean_ctor_set(x_102, 0, x_99); +lean_ctor_set(x_102, 1, x_100); +return x_102; } } else { -lean_object* x_101; +lean_object* x_103; lean_dec(x_15); -x_101 = lean_ctor_get(x_85, 1); -lean_inc(x_101); -lean_dec(x_85); -if (lean_obj_tag(x_101) == 0) +x_103 = lean_ctor_get(x_86, 1); +lean_inc(x_103); +lean_dec(x_86); +if (lean_obj_tag(x_103) == 0) { -lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; -x_102 = l_Lean_Meta_simpTargetCore___closed__3; -x_103 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_1, x_102, x_6, x_7, x_8, x_9, x_48); -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 = l_Lean_Meta_simpTargetCore___lambda__1(x_86, x_104, x_6, x_7, x_8, x_9, x_105); +lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; +x_104 = l_Lean_Meta_simpTargetCore___closed__3; +x_105 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_1, x_104, x_6, x_7, x_8, x_9, x_48); +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 = l_Lean_Meta_simpTargetCore___lambda__1(x_87, x_106, x_6, x_7, x_8, x_9, x_107); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_104); -return x_106; +lean_dec(x_106); +return x_108; } else { -lean_object* x_107; lean_object* x_108; -x_107 = lean_ctor_get(x_101, 0); -lean_inc(x_107); -lean_dec(x_101); +lean_object* x_109; lean_object* x_110; +x_109 = lean_ctor_get(x_103, 0); +lean_inc(x_109); +lean_dec(x_103); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_108 = l_Lean_Meta_mkOfEqTrue(x_107, x_6, x_7, x_8, x_9, x_48); -if (lean_obj_tag(x_108) == 0) +x_110 = l_Lean_Meta_mkOfEqTrue(x_109, x_6, x_7, x_8, x_9, x_48); +if (lean_obj_tag(x_110) == 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; -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 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_1, x_109, x_6, x_7, x_8, x_9, x_110); -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; lean_object* 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_Meta_simpTargetCore___lambda__1(x_86, x_112, x_6, x_7, x_8, x_9, x_113); +lean_dec(x_110); +x_113 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_1, x_111, x_6, x_7, x_8, x_9, x_112); +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 = l_Lean_Meta_simpTargetCore___lambda__1(x_87, x_114, x_6, x_7, x_8, x_9, x_115); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_112); -return x_114; +lean_dec(x_114); +return x_116; } else { -lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; -lean_dec(x_86); +lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; +lean_dec(x_87); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_1); -x_115 = lean_ctor_get(x_108, 0); -lean_inc(x_115); -x_116 = lean_ctor_get(x_108, 1); -lean_inc(x_116); -if (lean_is_exclusive(x_108)) { - lean_ctor_release(x_108, 0); - lean_ctor_release(x_108, 1); - x_117 = x_108; +x_117 = lean_ctor_get(x_110, 0); +lean_inc(x_117); +x_118 = lean_ctor_get(x_110, 1); +lean_inc(x_118); +if (lean_is_exclusive(x_110)) { + lean_ctor_release(x_110, 0); + lean_ctor_release(x_110, 1); + x_119 = x_110; } else { - lean_dec_ref(x_108); - x_117 = lean_box(0); + lean_dec_ref(x_110); + x_119 = lean_box(0); } -if (lean_is_scalar(x_117)) { - x_118 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_119)) { + x_120 = lean_alloc_ctor(1, 2, 0); } else { - x_118 = x_117; + x_120 = x_119; } -lean_ctor_set(x_118, 0, x_115); -lean_ctor_set(x_118, 1, x_116); -return x_118; +lean_ctor_set(x_120, 0, x_117); +lean_ctor_set(x_120, 1, x_118); +return x_120; } } } @@ -48314,36 +48327,36 @@ return x_118; } else { -uint8_t x_119; +uint8_t x_121; lean_dec(x_15); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_1); -x_119 = !lean_is_exclusive(x_17); -if (x_119 == 0) +x_121 = !lean_is_exclusive(x_17); +if (x_121 == 0) { return x_17; } else { -lean_object* x_120; lean_object* x_121; lean_object* x_122; -x_120 = lean_ctor_get(x_17, 0); -x_121 = lean_ctor_get(x_17, 1); -lean_inc(x_121); -lean_inc(x_120); +lean_object* x_122; lean_object* x_123; lean_object* x_124; +x_122 = lean_ctor_get(x_17, 0); +x_123 = lean_ctor_get(x_17, 1); +lean_inc(x_123); +lean_inc(x_122); lean_dec(x_17); -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; +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_123; +uint8_t x_125; lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -48352,23 +48365,23 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_123 = !lean_is_exclusive(x_11); -if (x_123 == 0) +x_125 = !lean_is_exclusive(x_11); +if (x_125 == 0) { return x_11; } else { -lean_object* x_124; lean_object* x_125; lean_object* x_126; -x_124 = lean_ctor_get(x_11, 0); -x_125 = lean_ctor_get(x_11, 1); -lean_inc(x_125); -lean_inc(x_124); +lean_object* x_126; lean_object* x_127; lean_object* x_128; +x_126 = lean_ctor_get(x_11, 0); +x_127 = lean_ctor_get(x_11, 1); +lean_inc(x_127); +lean_inc(x_126); lean_dec(x_11); -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; +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; } } } @@ -48516,254 +48529,256 @@ goto block_61; } else { -lean_object* x_63; lean_object* x_64; uint8_t x_65; +lean_object* x_63; lean_object* x_64; lean_object* x_65; uint8_t x_66; x_63 = lean_ctor_get(x_4, 0); lean_inc(x_63); -x_64 = l_Lean_Meta_Simp_isEqnThmHypothesis_go___closed__2; -x_65 = l_Lean_Expr_isConstOf(x_63, x_64); +x_64 = l_Lean_Expr_consumeMData(x_63); lean_dec(x_63); -if (x_65 == 0) +x_65 = l_Lean_Meta_Simp_isEqnThmHypothesis_go___closed__2; +x_66 = l_Lean_Expr_isConstOf(x_64, x_65); +lean_dec(x_64); +if (x_66 == 0) { -lean_object* x_66; +lean_object* x_67; lean_dec(x_1); -x_66 = lean_box(0); -x_11 = x_66; +x_67 = lean_box(0); +x_11 = x_67; goto block_61; } else { -lean_object* x_67; lean_object* x_68; -x_67 = l_Lean_Meta_applySimpResultToProp___closed__1; -x_68 = lean_ctor_get(x_4, 1); -lean_inc(x_68); +lean_object* x_68; lean_object* x_69; +x_68 = l_Lean_Meta_applySimpResultToProp___closed__1; +x_69 = lean_ctor_get(x_4, 1); +lean_inc(x_69); lean_dec(x_4); -if (lean_obj_tag(x_68) == 0) +if (lean_obj_tag(x_69) == 0) { -lean_object* x_69; +lean_object* x_70; lean_inc(x_1); -x_69 = l_Lean_MVarId_getType(x_1, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_69) == 0) +x_70 = l_Lean_MVarId_getType(x_1, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(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_object* x_71; lean_object* x_72; lean_object* x_73; +x_71 = lean_ctor_get(x_70, 0); lean_inc(x_71); -lean_dec(x_69); +x_72 = lean_ctor_get(x_70, 1); +lean_inc(x_72); +lean_dec(x_70); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_72 = l_Lean_Meta_mkFalseElim(x_70, x_2, x_6, x_7, x_8, x_9, x_71); -if (lean_obj_tag(x_72) == 0) +x_73 = l_Lean_Meta_mkFalseElim(x_71, x_2, x_6, x_7, x_8, x_9, x_72); +if (lean_obj_tag(x_73) == 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; -x_73 = lean_ctor_get(x_72, 0); -lean_inc(x_73); -x_74 = lean_ctor_get(x_72, 1); +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_74 = lean_ctor_get(x_73, 0); lean_inc(x_74); -lean_dec(x_72); -x_75 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_1, x_73, x_6, x_7, x_8, x_9, x_74); -x_76 = lean_ctor_get(x_75, 0); -lean_inc(x_76); -x_77 = lean_ctor_get(x_75, 1); +x_75 = lean_ctor_get(x_73, 1); +lean_inc(x_75); +lean_dec(x_73); +x_76 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_1, x_74, x_6, x_7, x_8, x_9, x_75); +x_77 = lean_ctor_get(x_76, 0); lean_inc(x_77); -lean_dec(x_75); -x_78 = lean_apply_6(x_67, x_76, x_6, x_7, x_8, x_9, x_77); -return x_78; +x_78 = lean_ctor_get(x_76, 1); +lean_inc(x_78); +lean_dec(x_76); +x_79 = lean_apply_6(x_68, x_77, x_6, x_7, x_8, x_9, x_78); +return x_79; } else { -uint8_t x_79; +uint8_t x_80; lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_1); -x_79 = !lean_is_exclusive(x_72); -if (x_79 == 0) +x_80 = !lean_is_exclusive(x_73); +if (x_80 == 0) { -return x_72; +return x_73; } else { -lean_object* x_80; lean_object* x_81; lean_object* x_82; -x_80 = lean_ctor_get(x_72, 0); -x_81 = lean_ctor_get(x_72, 1); +lean_object* x_81; lean_object* x_82; lean_object* x_83; +x_81 = lean_ctor_get(x_73, 0); +x_82 = lean_ctor_get(x_73, 1); +lean_inc(x_82); lean_inc(x_81); -lean_inc(x_80); -lean_dec(x_72); -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_dec(x_73); +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 { -uint8_t x_83; +uint8_t x_84; 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_83 = !lean_is_exclusive(x_69); -if (x_83 == 0) +x_84 = !lean_is_exclusive(x_70); +if (x_84 == 0) { -return x_69; +return x_70; } else { -lean_object* x_84; lean_object* x_85; lean_object* x_86; -x_84 = lean_ctor_get(x_69, 0); -x_85 = lean_ctor_get(x_69, 1); +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_inc(x_84); -lean_dec(x_69); -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_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; } } } else { -lean_object* x_87; lean_object* x_88; -x_87 = lean_ctor_get(x_68, 0); -lean_inc(x_87); -lean_dec(x_68); +lean_object* x_88; lean_object* x_89; +x_88 = lean_ctor_get(x_69, 0); +lean_inc(x_88); +lean_dec(x_69); lean_inc(x_1); -x_88 = l_Lean_MVarId_getType(x_1, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_88) == 0) +x_89 = l_Lean_MVarId_getType(x_1, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_89) == 0) { -lean_object* x_89; lean_object* x_90; lean_object* x_91; -x_89 = lean_ctor_get(x_88, 0); -lean_inc(x_89); -x_90 = lean_ctor_get(x_88, 1); +lean_object* x_90; lean_object* x_91; lean_object* x_92; +x_90 = lean_ctor_get(x_89, 0); lean_inc(x_90); -lean_dec(x_88); +x_91 = lean_ctor_get(x_89, 1); +lean_inc(x_91); +lean_dec(x_89); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_91 = l_Lean_Meta_mkEqMP(x_87, x_2, x_6, x_7, x_8, x_9, x_90); -if (lean_obj_tag(x_91) == 0) +x_92 = l_Lean_Meta_mkEqMP(x_88, x_2, x_6, x_7, x_8, x_9, x_91); +if (lean_obj_tag(x_92) == 0) { -lean_object* x_92; lean_object* x_93; lean_object* x_94; -x_92 = lean_ctor_get(x_91, 0); -lean_inc(x_92); -x_93 = lean_ctor_get(x_91, 1); +lean_object* x_93; lean_object* x_94; lean_object* x_95; +x_93 = lean_ctor_get(x_92, 0); lean_inc(x_93); -lean_dec(x_91); +x_94 = lean_ctor_get(x_92, 1); +lean_inc(x_94); +lean_dec(x_92); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_94 = l_Lean_Meta_mkFalseElim(x_89, x_92, x_6, x_7, x_8, x_9, x_93); -if (lean_obj_tag(x_94) == 0) +x_95 = l_Lean_Meta_mkFalseElim(x_90, x_93, x_6, x_7, x_8, x_9, x_94); +if (lean_obj_tag(x_95) == 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; -x_95 = lean_ctor_get(x_94, 0); -lean_inc(x_95); -x_96 = lean_ctor_get(x_94, 1); +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_96 = lean_ctor_get(x_95, 0); lean_inc(x_96); -lean_dec(x_94); -x_97 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_1, x_95, x_6, x_7, x_8, x_9, x_96); -x_98 = lean_ctor_get(x_97, 0); -lean_inc(x_98); -x_99 = lean_ctor_get(x_97, 1); +x_97 = lean_ctor_get(x_95, 1); +lean_inc(x_97); +lean_dec(x_95); +x_98 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_1, x_96, x_6, x_7, x_8, x_9, x_97); +x_99 = lean_ctor_get(x_98, 0); lean_inc(x_99); -lean_dec(x_97); -x_100 = lean_apply_6(x_67, x_98, x_6, x_7, x_8, x_9, x_99); -return x_100; +x_100 = lean_ctor_get(x_98, 1); +lean_inc(x_100); +lean_dec(x_98); +x_101 = lean_apply_6(x_68, x_99, x_6, x_7, x_8, x_9, x_100); +return x_101; } else { -uint8_t x_101; +uint8_t x_102; lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_1); -x_101 = !lean_is_exclusive(x_94); -if (x_101 == 0) +x_102 = !lean_is_exclusive(x_95); +if (x_102 == 0) { -return x_94; +return x_95; } else { -lean_object* x_102; lean_object* x_103; lean_object* x_104; -x_102 = lean_ctor_get(x_94, 0); -x_103 = lean_ctor_get(x_94, 1); +lean_object* x_103; lean_object* x_104; lean_object* x_105; +x_103 = lean_ctor_get(x_95, 0); +x_104 = lean_ctor_get(x_95, 1); +lean_inc(x_104); lean_inc(x_103); -lean_inc(x_102); -lean_dec(x_94); -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_95); +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; } } } else { -uint8_t x_105; -lean_dec(x_89); +uint8_t x_106; +lean_dec(x_90); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_1); -x_105 = !lean_is_exclusive(x_91); -if (x_105 == 0) +x_106 = !lean_is_exclusive(x_92); +if (x_106 == 0) { -return x_91; +return x_92; } else { -lean_object* x_106; lean_object* x_107; lean_object* x_108; -x_106 = lean_ctor_get(x_91, 0); -x_107 = lean_ctor_get(x_91, 1); +lean_object* x_107; lean_object* x_108; lean_object* x_109; +x_107 = lean_ctor_get(x_92, 0); +x_108 = lean_ctor_get(x_92, 1); +lean_inc(x_108); lean_inc(x_107); -lean_inc(x_106); -lean_dec(x_91); -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; +lean_dec(x_92); +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; } } } else { -uint8_t x_109; -lean_dec(x_87); +uint8_t x_110; +lean_dec(x_88); 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_109 = !lean_is_exclusive(x_88); -if (x_109 == 0) +x_110 = !lean_is_exclusive(x_89); +if (x_110 == 0) { -return x_88; +return x_89; } else { -lean_object* x_110; lean_object* x_111; lean_object* x_112; -x_110 = lean_ctor_get(x_88, 0); -x_111 = lean_ctor_get(x_88, 1); +lean_object* x_111; lean_object* x_112; lean_object* x_113; +x_111 = lean_ctor_get(x_89, 0); +x_112 = lean_ctor_get(x_89, 1); +lean_inc(x_112); lean_inc(x_111); -lean_inc(x_110); -lean_dec(x_88); -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_dec(x_89); +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; } } } @@ -50344,280 +50359,284 @@ uint8_t x_22; x_22 = !lean_is_exclusive(x_12); if (x_22 == 0) { -lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; x_23 = lean_ctor_get(x_12, 0); x_24 = lean_ctor_get(x_12, 1); -x_25 = l_Lean_Meta_Simp_isEqnThmHypothesis_go___closed__2; -x_26 = l_Lean_Expr_isConstOf(x_11, x_25); +x_25 = l_Lean_Expr_consumeMData(x_11); lean_dec(x_11); -if (x_26 == 0) +x_26 = l_Lean_Meta_Simp_isEqnThmHypothesis_go___closed__2; +x_27 = l_Lean_Expr_isConstOf(x_25, x_26); +lean_dec(x_25); +if (x_27 == 0) { -lean_object* x_27; lean_object* x_28; +lean_object* x_28; lean_object* x_29; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_27 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_27, 0, x_2); -lean_ctor_set(x_27, 1, x_23); -x_28 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_28, 0, x_27); -lean_ctor_set(x_12, 0, x_28); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_2); +lean_ctor_set(x_28, 1, x_23); +x_29 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_12, 0, x_29); return x_12; } else { -lean_object* x_29; +lean_object* x_30; lean_free_object(x_12); lean_inc(x_23); -x_29 = l_Lean_MVarId_getType(x_23, x_5, x_6, x_7, x_8, x_24); -if (lean_obj_tag(x_29) == 0) +x_30 = l_Lean_MVarId_getType(x_23, x_5, x_6, x_7, x_8, x_24); +if (lean_obj_tag(x_30) == 0) { -lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; -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; +x_31 = lean_ctor_get(x_30, 0); lean_inc(x_31); -lean_dec(x_29); -x_32 = l_Lean_Expr_fvar___override(x_2); +x_32 = lean_ctor_get(x_30, 1); +lean_inc(x_32); +lean_dec(x_30); +x_33 = l_Lean_Expr_fvar___override(x_2); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_33 = l_Lean_Meta_mkFalseElim(x_30, x_32, x_5, x_6, x_7, x_8, x_31); -if (lean_obj_tag(x_33) == 0) +x_34 = l_Lean_Meta_mkFalseElim(x_31, x_33, x_5, x_6, x_7, x_8, x_32); +if (lean_obj_tag(x_34) == 0) { -lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; -x_34 = lean_ctor_get(x_33, 0); -lean_inc(x_34); -x_35 = lean_ctor_get(x_33, 1); +lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; +x_35 = lean_ctor_get(x_34, 0); lean_inc(x_35); -lean_dec(x_33); -x_36 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_23, x_34, x_5, x_6, x_7, x_8, x_35); +x_36 = lean_ctor_get(x_34, 1); +lean_inc(x_36); +lean_dec(x_34); +x_37 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_23, x_35, x_5, x_6, x_7, x_8, x_36); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_37 = !lean_is_exclusive(x_36); -if (x_37 == 0) +x_38 = !lean_is_exclusive(x_37); +if (x_38 == 0) { -lean_object* x_38; lean_object* x_39; -x_38 = lean_ctor_get(x_36, 0); -lean_dec(x_38); -x_39 = lean_box(0); -lean_ctor_set(x_36, 0, x_39); -return x_36; +lean_object* x_39; lean_object* x_40; +x_39 = lean_ctor_get(x_37, 0); +lean_dec(x_39); +x_40 = lean_box(0); +lean_ctor_set(x_37, 0, x_40); +return x_37; } else { -lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_40 = lean_ctor_get(x_36, 1); -lean_inc(x_40); -lean_dec(x_36); -x_41 = lean_box(0); -x_42 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_40); -return x_42; +lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = lean_ctor_get(x_37, 1); +lean_inc(x_41); +lean_dec(x_37); +x_42 = lean_box(0); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_41); +return x_43; } } else { -uint8_t x_43; +uint8_t x_44; lean_dec(x_23); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_43 = !lean_is_exclusive(x_33); -if (x_43 == 0) +x_44 = !lean_is_exclusive(x_34); +if (x_44 == 0) { -return x_33; +return x_34; } else { -lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_44 = lean_ctor_get(x_33, 0); -x_45 = lean_ctor_get(x_33, 1); +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_34, 0); +x_46 = lean_ctor_get(x_34, 1); +lean_inc(x_46); lean_inc(x_45); -lean_inc(x_44); -lean_dec(x_33); -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_dec(x_34); +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 { -uint8_t x_47; +uint8_t x_48; lean_dec(x_23); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_2); -x_47 = !lean_is_exclusive(x_29); -if (x_47 == 0) +x_48 = !lean_is_exclusive(x_30); +if (x_48 == 0) { -return x_29; +return x_30; } else { -lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_48 = lean_ctor_get(x_29, 0); -x_49 = lean_ctor_get(x_29, 1); +lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_49 = lean_ctor_get(x_30, 0); +x_50 = lean_ctor_get(x_30, 1); +lean_inc(x_50); lean_inc(x_49); -lean_inc(x_48); -lean_dec(x_29); -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; +lean_dec(x_30); +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_51; lean_object* x_52; lean_object* x_53; uint8_t x_54; -x_51 = lean_ctor_get(x_12, 0); -x_52 = lean_ctor_get(x_12, 1); +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_12, 0); +x_53 = lean_ctor_get(x_12, 1); +lean_inc(x_53); lean_inc(x_52); -lean_inc(x_51); lean_dec(x_12); -x_53 = l_Lean_Meta_Simp_isEqnThmHypothesis_go___closed__2; -x_54 = l_Lean_Expr_isConstOf(x_11, x_53); +x_54 = l_Lean_Expr_consumeMData(x_11); lean_dec(x_11); -if (x_54 == 0) +x_55 = l_Lean_Meta_Simp_isEqnThmHypothesis_go___closed__2; +x_56 = l_Lean_Expr_isConstOf(x_54, x_55); +lean_dec(x_54); +if (x_56 == 0) { -lean_object* x_55; lean_object* x_56; lean_object* x_57; +lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_55 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_55, 0, x_2); -lean_ctor_set(x_55, 1, x_51); -x_56 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_56, 0, x_55); x_57 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_57, 0, x_56); +lean_ctor_set(x_57, 0, x_2); lean_ctor_set(x_57, 1, x_52); -return x_57; +x_58 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_58, 0, 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_53); +return x_59; } else { -lean_object* x_58; -lean_inc(x_51); -x_58 = l_Lean_MVarId_getType(x_51, x_5, x_6, x_7, x_8, x_52); -if (lean_obj_tag(x_58) == 0) +lean_object* x_60; +lean_inc(x_52); +x_60 = l_Lean_MVarId_getType(x_52, x_5, x_6, x_7, x_8, x_53); +if (lean_obj_tag(x_60) == 0) { -lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_59 = lean_ctor_get(x_58, 0); -lean_inc(x_59); -x_60 = lean_ctor_get(x_58, 1); -lean_inc(x_60); -lean_dec(x_58); -x_61 = l_Lean_Expr_fvar___override(x_2); +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +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_63 = l_Lean_Expr_fvar___override(x_2); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_62 = l_Lean_Meta_mkFalseElim(x_59, x_61, x_5, x_6, x_7, x_8, x_60); -if (lean_obj_tag(x_62) == 0) +x_64 = l_Lean_Meta_mkFalseElim(x_61, x_63, x_5, x_6, x_7, x_8, x_62); +if (lean_obj_tag(x_64) == 0) { -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_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_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_51, x_63, x_5, x_6, x_7, x_8, 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_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_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_52, x_65, x_5, x_6, x_7, x_8, x_66); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_66 = lean_ctor_get(x_65, 1); -lean_inc(x_66); -if (lean_is_exclusive(x_65)) { - lean_ctor_release(x_65, 0); - lean_ctor_release(x_65, 1); - x_67 = x_65; +x_68 = lean_ctor_get(x_67, 1); +lean_inc(x_68); +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_65); - x_67 = lean_box(0); + lean_dec_ref(x_67); + x_69 = lean_box(0); } -x_68 = lean_box(0); -if (lean_is_scalar(x_67)) { - x_69 = lean_alloc_ctor(0, 2, 0); +x_70 = lean_box(0); +if (lean_is_scalar(x_69)) { + x_71 = lean_alloc_ctor(0, 2, 0); } else { - x_69 = x_67; + x_71 = x_69; } -lean_ctor_set(x_69, 0, x_68); -lean_ctor_set(x_69, 1, x_66); -return x_69; +lean_ctor_set(x_71, 0, x_70); +lean_ctor_set(x_71, 1, x_68); +return x_71; } else { -lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; -lean_dec(x_51); +lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; +lean_dec(x_52); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_70 = lean_ctor_get(x_62, 0); -lean_inc(x_70); -x_71 = lean_ctor_get(x_62, 1); -lean_inc(x_71); -if (lean_is_exclusive(x_62)) { - lean_ctor_release(x_62, 0); - lean_ctor_release(x_62, 1); - x_72 = x_62; +x_72 = lean_ctor_get(x_64, 0); +lean_inc(x_72); +x_73 = lean_ctor_get(x_64, 1); +lean_inc(x_73); +if (lean_is_exclusive(x_64)) { + lean_ctor_release(x_64, 0); + lean_ctor_release(x_64, 1); + x_74 = x_64; } else { - lean_dec_ref(x_62); - x_72 = lean_box(0); + lean_dec_ref(x_64); + x_74 = lean_box(0); } -if (lean_is_scalar(x_72)) { - x_73 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_74)) { + x_75 = lean_alloc_ctor(1, 2, 0); } else { - x_73 = x_72; + x_75 = x_74; } -lean_ctor_set(x_73, 0, x_70); -lean_ctor_set(x_73, 1, x_71); -return x_73; +lean_ctor_set(x_75, 0, x_72); +lean_ctor_set(x_75, 1, x_73); +return x_75; } } else { -lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; -lean_dec(x_51); +lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; +lean_dec(x_52); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_2); -x_74 = lean_ctor_get(x_58, 0); -lean_inc(x_74); -x_75 = lean_ctor_get(x_58, 1); -lean_inc(x_75); -if (lean_is_exclusive(x_58)) { - lean_ctor_release(x_58, 0); - lean_ctor_release(x_58, 1); - x_76 = x_58; +x_76 = lean_ctor_get(x_60, 0); +lean_inc(x_76); +x_77 = lean_ctor_get(x_60, 1); +lean_inc(x_77); +if (lean_is_exclusive(x_60)) { + lean_ctor_release(x_60, 0); + lean_ctor_release(x_60, 1); + x_78 = x_60; } else { - lean_dec_ref(x_58); - x_76 = lean_box(0); + lean_dec_ref(x_60); + x_78 = lean_box(0); } -if (lean_is_scalar(x_76)) { - x_77 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_78)) { + x_79 = lean_alloc_ctor(1, 2, 0); } else { - x_77 = x_76; + x_79 = x_78; } -lean_ctor_set(x_77, 0, x_74); -lean_ctor_set(x_77, 1, x_75); -return x_77; +lean_ctor_set(x_79, 0, x_76); +lean_ctor_set(x_79, 1, x_77); +return x_79; } } } @@ -50625,36 +50644,36 @@ return x_77; } else { -uint8_t x_78; +uint8_t x_80; lean_dec(x_11); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_2); -x_78 = !lean_is_exclusive(x_12); -if (x_78 == 0) +x_80 = !lean_is_exclusive(x_12); +if (x_80 == 0) { return x_12; } else { -lean_object* x_79; lean_object* x_80; lean_object* x_81; -x_79 = lean_ctor_get(x_12, 0); -x_80 = lean_ctor_get(x_12, 1); -lean_inc(x_80); -lean_inc(x_79); +lean_object* x_81; lean_object* x_82; lean_object* x_83; +x_81 = lean_ctor_get(x_12, 0); +x_82 = lean_ctor_get(x_12, 1); +lean_inc(x_82); +lean_inc(x_81); lean_dec(x_12); -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_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_82; +lean_object* x_84; lean_dec(x_10); lean_inc(x_8); lean_inc(x_7); @@ -50662,44 +50681,44 @@ lean_inc(x_6); lean_inc(x_5); lean_inc(x_2); lean_inc(x_1); -x_82 = l_Lean_Meta_applySimpResultToFVarId(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_82) == 0) +x_84 = l_Lean_Meta_applySimpResultToFVarId(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_84) == 0) { -lean_object* x_83; lean_object* x_84; lean_object* x_85; -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 = l_Lean_Meta_applySimpResultToLocalDeclCore(x_1, x_2, x_83, x_5, x_6, x_7, x_8, x_84); -return x_85; +lean_object* x_85; lean_object* x_86; lean_object* x_87; +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 = l_Lean_Meta_applySimpResultToLocalDeclCore(x_1, x_2, x_85, x_5, x_6, x_7, x_8, x_86); +return x_87; } else { -uint8_t x_86; +uint8_t x_88; 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_86 = !lean_is_exclusive(x_82); -if (x_86 == 0) +x_88 = !lean_is_exclusive(x_84); +if (x_88 == 0) { -return x_82; +return x_84; } else { -lean_object* x_87; lean_object* x_88; lean_object* x_89; -x_87 = lean_ctor_get(x_82, 0); -x_88 = lean_ctor_get(x_82, 1); -lean_inc(x_88); -lean_inc(x_87); -lean_dec(x_82); -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_object* x_89; lean_object* x_90; lean_object* x_91; +x_89 = lean_ctor_get(x_84, 0); +x_90 = lean_ctor_get(x_84, 1); +lean_inc(x_90); +lean_inc(x_89); +lean_dec(x_84); +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; } } } @@ -50740,7 +50759,7 @@ lean_inc(x_16); x_17 = lean_ctor_get(x_15, 1); lean_inc(x_17); lean_dec(x_15); -x_18 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_16, x_8, x_9, x_10, x_11, x_17); +x_18 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_16, x_8, x_9, x_10, x_11, x_17); x_19 = lean_ctor_get(x_18, 0); lean_inc(x_19); x_20 = lean_ctor_get(x_18, 1); @@ -51173,7 +51192,7 @@ x_39 = lean_ctor_get(x_37, 1); lean_inc(x_39); lean_dec(x_37); x_40 = l_Lean_LocalDecl_type(x_38); -x_41 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_40, x_8, x_9, x_10, x_11, x_39); +x_41 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_40, x_8, x_9, x_10, x_11, x_39); x_42 = lean_ctor_get(x_41, 0); lean_inc(x_42); x_43 = lean_ctor_get(x_41, 1); @@ -51226,44 +51245,46 @@ lean_dec(x_53); x_58 = !lean_is_exclusive(x_54); if (x_58 == 0) { -lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; +lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; uint8_t x_64; x_59 = lean_ctor_get(x_54, 1); x_60 = lean_ctor_get(x_54, 0); lean_dec(x_60); x_61 = lean_ctor_get(x_55, 0); lean_inc(x_61); -x_62 = l_Lean_Meta_Simp_isEqnThmHypothesis_go___closed__2; -x_63 = l_Lean_Expr_isConstOf(x_61, x_62); +x_62 = l_Lean_Expr_consumeMData(x_61); lean_dec(x_61); -if (x_63 == 0) +x_63 = l_Lean_Meta_Simp_isEqnThmHypothesis_go___closed__2; +x_64 = l_Lean_Expr_isConstOf(x_62, x_63); +lean_dec(x_62); +if (x_64 == 0) { -lean_object* x_64; lean_object* x_65; +lean_object* x_65; lean_object* x_66; lean_free_object(x_54); lean_free_object(x_27); lean_free_object(x_26); lean_free_object(x_25); -x_64 = lean_box(0); +x_65 = lean_box(0); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_3); -x_65 = l_Array_forInUnsafe_loop___at_Lean_Meta_simpGoal___spec__1___lambda__1(x_55, x_15, x_35, x_59, x_3, x_29, x_32, x_64, x_8, x_9, x_10, x_11, x_57); -if (lean_obj_tag(x_65) == 0) +x_66 = l_Array_forInUnsafe_loop___at_Lean_Meta_simpGoal___spec__1___lambda__1(x_55, x_15, x_35, x_59, x_3, x_29, x_32, x_65, x_8, x_9, x_10, x_11, x_57); +if (lean_obj_tag(x_66) == 0) { -lean_object* x_66; lean_object* x_67; -x_66 = lean_ctor_get(x_65, 0); -lean_inc(x_66); -x_67 = lean_ctor_get(x_65, 1); +lean_object* x_67; lean_object* x_68; +x_67 = lean_ctor_get(x_66, 0); lean_inc(x_67); -lean_dec(x_65); -x_16 = x_66; -x_17 = x_67; +x_68 = lean_ctor_get(x_66, 1); +lean_inc(x_68); +lean_dec(x_66); +x_16 = x_67; +x_17 = x_68; goto block_24; } else { -uint8_t x_68; +uint8_t x_69; lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -51271,77 +51292,77 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_68 = !lean_is_exclusive(x_65); -if (x_68 == 0) +x_69 = !lean_is_exclusive(x_66); +if (x_69 == 0) { -return x_65; +return x_66; } else { -lean_object* x_69; lean_object* x_70; lean_object* x_71; -x_69 = lean_ctor_get(x_65, 0); -x_70 = lean_ctor_get(x_65, 1); +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_inc(x_69); -lean_dec(x_65); -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; +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; } } } else { -lean_object* x_72; +lean_object* x_73; lean_dec(x_55); lean_inc(x_29); -x_72 = l_Lean_MVarId_getType(x_29, x_8, x_9, x_10, x_11, x_57); -if (lean_obj_tag(x_72) == 0) +x_73 = l_Lean_MVarId_getType(x_29, x_8, x_9, x_10, x_11, x_57); +if (lean_obj_tag(x_73) == 0) { -lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; -x_73 = lean_ctor_get(x_72, 0); -lean_inc(x_73); -x_74 = lean_ctor_get(x_72, 1); +lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; +x_74 = lean_ctor_get(x_73, 0); lean_inc(x_74); -lean_dec(x_72); -x_75 = l_Lean_Expr_fvar___override(x_15); +x_75 = lean_ctor_get(x_73, 1); +lean_inc(x_75); +lean_dec(x_73); +x_76 = l_Lean_Expr_fvar___override(x_15); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -x_76 = l_Lean_Meta_mkFalseElim(x_73, x_75, x_8, x_9, x_10, x_11, x_74); -if (lean_obj_tag(x_76) == 0) +x_77 = l_Lean_Meta_mkFalseElim(x_74, x_76, x_8, x_9, x_10, x_11, x_75); +if (lean_obj_tag(x_77) == 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; -x_77 = lean_ctor_get(x_76, 0); -lean_inc(x_77); -x_78 = lean_ctor_get(x_76, 1); +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_78 = lean_ctor_get(x_77, 0); lean_inc(x_78); -lean_dec(x_76); +x_79 = lean_ctor_get(x_77, 1); +lean_inc(x_79); +lean_dec(x_77); lean_inc(x_29); -x_79 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_29, x_77, x_8, x_9, x_10, x_11, x_78); -x_80 = lean_ctor_get(x_79, 1); -lean_inc(x_80); -lean_dec(x_79); -x_81 = lean_box(0); +x_80 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_29, x_78, x_8, x_9, x_10, x_11, x_79); +x_81 = lean_ctor_get(x_80, 1); +lean_inc(x_81); +lean_dec(x_80); +x_82 = lean_box(0); lean_inc(x_59); -lean_ctor_set(x_54, 0, x_81); -x_82 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_82, 0, x_54); +lean_ctor_set(x_54, 0, x_82); +x_83 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_83, 0, x_54); lean_ctor_set(x_27, 1, x_59); -x_83 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_83, 0, x_82); -lean_ctor_set(x_83, 1, x_25); -x_84 = lean_alloc_ctor(0, 1, 0); +x_84 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_84, 0, x_83); -x_16 = x_84; -x_17 = x_80; +lean_ctor_set(x_84, 1, x_25); +x_85 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_85, 0, x_84); +x_16 = x_85; +x_17 = x_81; goto block_24; } else { -uint8_t x_85; +uint8_t x_86; lean_free_object(x_54); lean_dec(x_59); lean_free_object(x_27); @@ -51357,29 +51378,29 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_85 = !lean_is_exclusive(x_76); -if (x_85 == 0) +x_86 = !lean_is_exclusive(x_77); +if (x_86 == 0) { -return x_76; +return x_77; } else { -lean_object* x_86; lean_object* x_87; lean_object* x_88; -x_86 = lean_ctor_get(x_76, 0); -x_87 = lean_ctor_get(x_76, 1); +lean_object* x_87; lean_object* x_88; lean_object* x_89; +x_87 = lean_ctor_get(x_77, 0); +x_88 = lean_ctor_get(x_77, 1); +lean_inc(x_88); lean_inc(x_87); -lean_inc(x_86); -lean_dec(x_76); -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_dec(x_77); +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_89; +uint8_t x_90; lean_free_object(x_54); lean_dec(x_59); lean_free_object(x_27); @@ -51396,66 +51417,68 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_89 = !lean_is_exclusive(x_72); -if (x_89 == 0) +x_90 = !lean_is_exclusive(x_73); +if (x_90 == 0) { -return x_72; +return x_73; } else { -lean_object* x_90; lean_object* x_91; lean_object* x_92; -x_90 = lean_ctor_get(x_72, 0); -x_91 = lean_ctor_get(x_72, 1); +lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_91 = lean_ctor_get(x_73, 0); +x_92 = lean_ctor_get(x_73, 1); +lean_inc(x_92); lean_inc(x_91); -lean_inc(x_90); -lean_dec(x_72); -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_dec(x_73); +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_93; lean_object* x_94; lean_object* x_95; uint8_t x_96; -x_93 = lean_ctor_get(x_54, 1); -lean_inc(x_93); -lean_dec(x_54); -x_94 = lean_ctor_get(x_55, 0); +lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; uint8_t x_98; +x_94 = lean_ctor_get(x_54, 1); lean_inc(x_94); -x_95 = l_Lean_Meta_Simp_isEqnThmHypothesis_go___closed__2; -x_96 = l_Lean_Expr_isConstOf(x_94, x_95); -lean_dec(x_94); -if (x_96 == 0) +lean_dec(x_54); +x_95 = lean_ctor_get(x_55, 0); +lean_inc(x_95); +x_96 = l_Lean_Expr_consumeMData(x_95); +lean_dec(x_95); +x_97 = l_Lean_Meta_Simp_isEqnThmHypothesis_go___closed__2; +x_98 = l_Lean_Expr_isConstOf(x_96, x_97); +lean_dec(x_96); +if (x_98 == 0) { -lean_object* x_97; lean_object* x_98; +lean_object* x_99; lean_object* x_100; lean_free_object(x_27); lean_free_object(x_26); lean_free_object(x_25); -x_97 = lean_box(0); +x_99 = lean_box(0); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_3); -x_98 = l_Array_forInUnsafe_loop___at_Lean_Meta_simpGoal___spec__1___lambda__1(x_55, x_15, x_35, x_93, x_3, x_29, x_32, x_97, x_8, x_9, x_10, x_11, x_57); -if (lean_obj_tag(x_98) == 0) +x_100 = l_Array_forInUnsafe_loop___at_Lean_Meta_simpGoal___spec__1___lambda__1(x_55, x_15, x_35, x_94, x_3, x_29, x_32, x_99, x_8, x_9, x_10, x_11, x_57); +if (lean_obj_tag(x_100) == 0) { -lean_object* x_99; lean_object* x_100; -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_16 = x_99; -x_17 = x_100; +lean_object* x_101; lean_object* x_102; +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_16 = x_101; +x_17 = x_102; goto block_24; } else { -lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; +lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -51463,82 +51486,82 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_101 = lean_ctor_get(x_98, 0); -lean_inc(x_101); -x_102 = lean_ctor_get(x_98, 1); -lean_inc(x_102); -if (lean_is_exclusive(x_98)) { - lean_ctor_release(x_98, 0); - lean_ctor_release(x_98, 1); - x_103 = x_98; +x_103 = lean_ctor_get(x_100, 0); +lean_inc(x_103); +x_104 = lean_ctor_get(x_100, 1); +lean_inc(x_104); +if (lean_is_exclusive(x_100)) { + lean_ctor_release(x_100, 0); + lean_ctor_release(x_100, 1); + x_105 = x_100; } else { - lean_dec_ref(x_98); - x_103 = lean_box(0); + lean_dec_ref(x_100); + x_105 = lean_box(0); } -if (lean_is_scalar(x_103)) { - x_104 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_105)) { + x_106 = lean_alloc_ctor(1, 2, 0); } else { - x_104 = x_103; + x_106 = x_105; } -lean_ctor_set(x_104, 0, x_101); -lean_ctor_set(x_104, 1, x_102); -return x_104; +lean_ctor_set(x_106, 0, x_103); +lean_ctor_set(x_106, 1, x_104); +return x_106; } } else { -lean_object* x_105; +lean_object* x_107; lean_dec(x_55); lean_inc(x_29); -x_105 = l_Lean_MVarId_getType(x_29, x_8, x_9, x_10, x_11, x_57); -if (lean_obj_tag(x_105) == 0) +x_107 = l_Lean_MVarId_getType(x_29, x_8, x_9, x_10, x_11, x_57); +if (lean_obj_tag(x_107) == 0) { -lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; -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 = l_Lean_Expr_fvar___override(x_15); +lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; +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 = l_Lean_Expr_fvar___override(x_15); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -x_109 = l_Lean_Meta_mkFalseElim(x_106, x_108, x_8, x_9, x_10, x_11, x_107); -if (lean_obj_tag(x_109) == 0) +x_111 = l_Lean_Meta_mkFalseElim(x_108, x_110, x_8, x_9, x_10, x_11, x_109); +if (lean_obj_tag(x_111) == 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; -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); -lean_inc(x_29); -x_112 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_29, x_110, x_8, x_9, x_10, x_11, x_111); -x_113 = lean_ctor_get(x_112, 1); +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_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_112); -x_114 = lean_box(0); -lean_inc(x_93); -x_115 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_115, 0, x_114); -lean_ctor_set(x_115, 1, x_93); -x_116 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_116, 0, x_115); -lean_ctor_set(x_27, 1, x_93); +lean_dec(x_111); +lean_inc(x_29); +x_114 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_29, x_112, x_8, x_9, x_10, x_11, x_113); +x_115 = lean_ctor_get(x_114, 1); +lean_inc(x_115); +lean_dec(x_114); +x_116 = lean_box(0); +lean_inc(x_94); x_117 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_117, 0, x_116); -lean_ctor_set(x_117, 1, x_25); -x_118 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_117, 1, x_94); +x_118 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_118, 0, x_117); -x_16 = x_118; -x_17 = x_113; +lean_ctor_set(x_27, 1, x_94); +x_119 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_119, 0, x_118); +lean_ctor_set(x_119, 1, x_25); +x_120 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_120, 0, x_119); +x_16 = x_120; +x_17 = x_115; goto block_24; } else { -lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; -lean_dec(x_93); +lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; +lean_dec(x_94); lean_free_object(x_27); lean_dec(x_35); lean_free_object(x_26); @@ -51552,32 +51575,32 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_119 = lean_ctor_get(x_109, 0); -lean_inc(x_119); -x_120 = lean_ctor_get(x_109, 1); -lean_inc(x_120); -if (lean_is_exclusive(x_109)) { - lean_ctor_release(x_109, 0); - lean_ctor_release(x_109, 1); - x_121 = x_109; +x_121 = lean_ctor_get(x_111, 0); +lean_inc(x_121); +x_122 = lean_ctor_get(x_111, 1); +lean_inc(x_122); +if (lean_is_exclusive(x_111)) { + lean_ctor_release(x_111, 0); + lean_ctor_release(x_111, 1); + x_123 = x_111; } else { - lean_dec_ref(x_109); - x_121 = lean_box(0); + lean_dec_ref(x_111); + x_123 = lean_box(0); } -if (lean_is_scalar(x_121)) { - x_122 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_123)) { + x_124 = lean_alloc_ctor(1, 2, 0); } else { - x_122 = x_121; + x_124 = x_123; } -lean_ctor_set(x_122, 0, x_119); -lean_ctor_set(x_122, 1, x_120); -return x_122; +lean_ctor_set(x_124, 0, x_121); +lean_ctor_set(x_124, 1, x_122); +return x_124; } } else { -lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; -lean_dec(x_93); +lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; +lean_dec(x_94); lean_free_object(x_27); lean_dec(x_35); lean_free_object(x_26); @@ -51592,125 +51615,125 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_123 = lean_ctor_get(x_105, 0); -lean_inc(x_123); -x_124 = lean_ctor_get(x_105, 1); -lean_inc(x_124); -if (lean_is_exclusive(x_105)) { - lean_ctor_release(x_105, 0); - lean_ctor_release(x_105, 1); - x_125 = x_105; +x_125 = lean_ctor_get(x_107, 0); +lean_inc(x_125); +x_126 = lean_ctor_get(x_107, 1); +lean_inc(x_126); +if (lean_is_exclusive(x_107)) { + lean_ctor_release(x_107, 0); + lean_ctor_release(x_107, 1); + x_127 = x_107; } else { - lean_dec_ref(x_105); - x_125 = lean_box(0); + lean_dec_ref(x_107); + x_127 = lean_box(0); } -if (lean_is_scalar(x_125)) { - x_126 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_127)) { + x_128 = lean_alloc_ctor(1, 2, 0); } else { - x_126 = x_125; + x_128 = x_127; } -lean_ctor_set(x_126, 0, x_123); -lean_ctor_set(x_126, 1, x_124); -return x_126; +lean_ctor_set(x_128, 0, x_125); +lean_ctor_set(x_128, 1, x_126); +return x_128; } } } } else { -uint8_t x_127; -x_127 = !lean_is_exclusive(x_56); -if (x_127 == 0) +uint8_t x_129; +x_129 = !lean_is_exclusive(x_56); +if (x_129 == 0) { -lean_object* x_128; lean_object* x_129; uint8_t x_130; -x_128 = lean_ctor_get(x_56, 0); -lean_dec(x_128); -x_129 = lean_ctor_get(x_53, 1); -lean_inc(x_129); +lean_object* x_130; lean_object* x_131; uint8_t x_132; +x_130 = lean_ctor_get(x_56, 0); +lean_dec(x_130); +x_131 = lean_ctor_get(x_53, 1); +lean_inc(x_131); lean_dec(x_53); -x_130 = !lean_is_exclusive(x_54); -if (x_130 == 0) +x_132 = !lean_is_exclusive(x_54); +if (x_132 == 0) { -lean_object* x_131; lean_object* x_132; lean_object* x_133; uint8_t x_134; lean_object* x_135; -x_131 = lean_ctor_get(x_54, 1); -x_132 = lean_ctor_get(x_54, 0); -lean_dec(x_132); -x_133 = l_Lean_Expr_fvar___override(x_15); -x_134 = 1; +lean_object* x_133; lean_object* x_134; lean_object* x_135; uint8_t x_136; lean_object* x_137; +x_133 = lean_ctor_get(x_54, 1); +x_134 = lean_ctor_get(x_54, 0); +lean_dec(x_134); +x_135 = l_Lean_Expr_fvar___override(x_15); +x_136 = 1; lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_29); -x_135 = l_Lean_Meta_applySimpResultToProp(x_29, x_133, x_42, x_55, x_134, x_8, x_9, x_10, x_11, x_129); +x_137 = l_Lean_Meta_applySimpResultToProp(x_29, x_135, x_42, x_55, x_136, x_8, x_9, x_10, x_11, x_131); lean_dec(x_42); -if (lean_obj_tag(x_135) == 0) +if (lean_obj_tag(x_137) == 0) { -lean_object* x_136; -x_136 = lean_ctor_get(x_135, 0); -lean_inc(x_136); -if (lean_obj_tag(x_136) == 0) +lean_object* x_138; +x_138 = lean_ctor_get(x_137, 0); +lean_inc(x_138); +if (lean_obj_tag(x_138) == 0) { -lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; +lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_dec(x_38); -x_137 = lean_ctor_get(x_135, 1); -lean_inc(x_137); -lean_dec(x_135); -x_138 = lean_box(0); -lean_inc(x_131); -lean_ctor_set(x_54, 0, x_138); +x_139 = lean_ctor_get(x_137, 1); +lean_inc(x_139); +lean_dec(x_137); +x_140 = lean_box(0); +lean_inc(x_133); +lean_ctor_set(x_54, 0, x_140); lean_ctor_set(x_56, 0, x_54); -lean_ctor_set(x_27, 1, x_131); -x_139 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_139, 0, x_56); -lean_ctor_set(x_139, 1, x_25); -x_140 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_140, 0, x_139); -x_16 = x_140; -x_17 = x_137; +lean_ctor_set(x_27, 1, x_133); +x_141 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_141, 0, x_56); +lean_ctor_set(x_141, 1, x_25); +x_142 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_142, 0, x_141); +x_16 = x_142; +x_17 = x_139; goto block_24; } 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_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_free_object(x_54); lean_free_object(x_56); -x_141 = lean_ctor_get(x_136, 0); -lean_inc(x_141); -lean_dec(x_136); -x_142 = lean_ctor_get(x_135, 1); -lean_inc(x_142); -lean_dec(x_135); -x_143 = lean_ctor_get(x_141, 0); +x_143 = lean_ctor_get(x_138, 0); lean_inc(x_143); -x_144 = lean_ctor_get(x_141, 1); +lean_dec(x_138); +x_144 = lean_ctor_get(x_137, 1); lean_inc(x_144); -lean_dec(x_141); -x_145 = l_Lean_LocalDecl_userName(x_38); +lean_dec(x_137); +x_145 = lean_ctor_get(x_143, 0); +lean_inc(x_145); +x_146 = lean_ctor_get(x_143, 1); +lean_inc(x_146); +lean_dec(x_143); +x_147 = l_Lean_LocalDecl_userName(x_38); lean_dec(x_38); -x_146 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_146, 0, x_145); -lean_ctor_set(x_146, 1, x_144); -lean_ctor_set(x_146, 2, x_143); -x_147 = lean_array_push(x_35, x_146); -lean_ctor_set(x_27, 1, x_131); -lean_ctor_set(x_27, 0, x_147); +x_148 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_148, 0, x_147); +lean_ctor_set(x_148, 1, x_146); +lean_ctor_set(x_148, 2, x_145); +x_149 = lean_array_push(x_35, x_148); +lean_ctor_set(x_27, 1, x_133); +lean_ctor_set(x_27, 0, x_149); lean_inc(x_3); -x_148 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_148, 0, x_3); -lean_ctor_set(x_148, 1, x_25); -x_149 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_149, 0, x_148); -x_16 = x_149; -x_17 = x_142; +x_150 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_150, 0, x_3); +lean_ctor_set(x_150, 1, x_25); +x_151 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_151, 0, x_150); +x_16 = x_151; +x_17 = x_144; goto block_24; } } else { -uint8_t x_150; +uint8_t x_152; lean_free_object(x_54); -lean_dec(x_131); +lean_dec(x_133); lean_free_object(x_56); lean_dec(x_38); lean_free_object(x_27); @@ -51726,108 +51749,108 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_150 = !lean_is_exclusive(x_135); -if (x_150 == 0) +x_152 = !lean_is_exclusive(x_137); +if (x_152 == 0) { -return x_135; +return x_137; } else { -lean_object* x_151; lean_object* x_152; lean_object* x_153; -x_151 = lean_ctor_get(x_135, 0); -x_152 = lean_ctor_get(x_135, 1); -lean_inc(x_152); -lean_inc(x_151); -lean_dec(x_135); -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_object* x_153; lean_object* x_154; lean_object* x_155; +x_153 = lean_ctor_get(x_137, 0); +x_154 = lean_ctor_get(x_137, 1); +lean_inc(x_154); +lean_inc(x_153); +lean_dec(x_137); +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; } } } else { -lean_object* x_154; lean_object* x_155; uint8_t x_156; lean_object* x_157; -x_154 = lean_ctor_get(x_54, 1); -lean_inc(x_154); +lean_object* x_156; lean_object* x_157; uint8_t x_158; lean_object* x_159; +x_156 = lean_ctor_get(x_54, 1); +lean_inc(x_156); lean_dec(x_54); -x_155 = l_Lean_Expr_fvar___override(x_15); -x_156 = 1; +x_157 = l_Lean_Expr_fvar___override(x_15); +x_158 = 1; lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_29); -x_157 = l_Lean_Meta_applySimpResultToProp(x_29, x_155, x_42, x_55, x_156, x_8, x_9, x_10, x_11, x_129); +x_159 = l_Lean_Meta_applySimpResultToProp(x_29, x_157, x_42, x_55, x_158, x_8, x_9, x_10, x_11, x_131); lean_dec(x_42); -if (lean_obj_tag(x_157) == 0) +if (lean_obj_tag(x_159) == 0) { -lean_object* x_158; -x_158 = lean_ctor_get(x_157, 0); -lean_inc(x_158); -if (lean_obj_tag(x_158) == 0) +lean_object* x_160; +x_160 = lean_ctor_get(x_159, 0); +lean_inc(x_160); +if (lean_obj_tag(x_160) == 0) { -lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; +lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_dec(x_38); -x_159 = lean_ctor_get(x_157, 1); -lean_inc(x_159); -lean_dec(x_157); -x_160 = lean_box(0); -lean_inc(x_154); -x_161 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_161, 0, x_160); -lean_ctor_set(x_161, 1, x_154); -lean_ctor_set(x_56, 0, x_161); -lean_ctor_set(x_27, 1, x_154); -x_162 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_162, 0, x_56); -lean_ctor_set(x_162, 1, x_25); -x_163 = lean_alloc_ctor(0, 1, 0); +x_161 = lean_ctor_get(x_159, 1); +lean_inc(x_161); +lean_dec(x_159); +x_162 = lean_box(0); +lean_inc(x_156); +x_163 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_163, 0, x_162); -x_16 = x_163; -x_17 = x_159; +lean_ctor_set(x_163, 1, x_156); +lean_ctor_set(x_56, 0, x_163); +lean_ctor_set(x_27, 1, x_156); +x_164 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_164, 0, x_56); +lean_ctor_set(x_164, 1, x_25); +x_165 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_165, 0, x_164); +x_16 = x_165; +x_17 = x_161; goto block_24; } 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_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_free_object(x_56); -x_164 = lean_ctor_get(x_158, 0); -lean_inc(x_164); -lean_dec(x_158); -x_165 = lean_ctor_get(x_157, 1); -lean_inc(x_165); -lean_dec(x_157); -x_166 = lean_ctor_get(x_164, 0); +x_166 = lean_ctor_get(x_160, 0); lean_inc(x_166); -x_167 = lean_ctor_get(x_164, 1); +lean_dec(x_160); +x_167 = lean_ctor_get(x_159, 1); lean_inc(x_167); -lean_dec(x_164); -x_168 = l_Lean_LocalDecl_userName(x_38); +lean_dec(x_159); +x_168 = lean_ctor_get(x_166, 0); +lean_inc(x_168); +x_169 = lean_ctor_get(x_166, 1); +lean_inc(x_169); +lean_dec(x_166); +x_170 = l_Lean_LocalDecl_userName(x_38); lean_dec(x_38); -x_169 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_169, 0, x_168); -lean_ctor_set(x_169, 1, x_167); -lean_ctor_set(x_169, 2, x_166); -x_170 = lean_array_push(x_35, x_169); -lean_ctor_set(x_27, 1, x_154); -lean_ctor_set(x_27, 0, x_170); +x_171 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_171, 0, x_170); +lean_ctor_set(x_171, 1, x_169); +lean_ctor_set(x_171, 2, x_168); +x_172 = lean_array_push(x_35, x_171); +lean_ctor_set(x_27, 1, x_156); +lean_ctor_set(x_27, 0, x_172); lean_inc(x_3); -x_171 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_171, 0, x_3); -lean_ctor_set(x_171, 1, x_25); -x_172 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_172, 0, x_171); -x_16 = x_172; -x_17 = x_165; +x_173 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_173, 0, x_3); +lean_ctor_set(x_173, 1, x_25); +x_174 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_174, 0, x_173); +x_16 = x_174; +x_17 = x_167; goto block_24; } } else { -lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; -lean_dec(x_154); +lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; +lean_dec(x_156); lean_free_object(x_56); lean_dec(x_38); lean_free_object(x_27); @@ -51843,128 +51866,128 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_173 = lean_ctor_get(x_157, 0); -lean_inc(x_173); -x_174 = lean_ctor_get(x_157, 1); -lean_inc(x_174); -if (lean_is_exclusive(x_157)) { - lean_ctor_release(x_157, 0); - lean_ctor_release(x_157, 1); - x_175 = x_157; +x_175 = lean_ctor_get(x_159, 0); +lean_inc(x_175); +x_176 = lean_ctor_get(x_159, 1); +lean_inc(x_176); +if (lean_is_exclusive(x_159)) { + lean_ctor_release(x_159, 0); + lean_ctor_release(x_159, 1); + x_177 = x_159; } else { - lean_dec_ref(x_157); - x_175 = lean_box(0); + lean_dec_ref(x_159); + x_177 = lean_box(0); } -if (lean_is_scalar(x_175)) { - x_176 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_177)) { + x_178 = lean_alloc_ctor(1, 2, 0); } else { - x_176 = x_175; + x_178 = x_177; } -lean_ctor_set(x_176, 0, x_173); -lean_ctor_set(x_176, 1, x_174); -return x_176; +lean_ctor_set(x_178, 0, x_175); +lean_ctor_set(x_178, 1, x_176); +return x_178; } } } else { -lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; uint8_t x_181; lean_object* x_182; +lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; uint8_t x_183; lean_object* x_184; lean_dec(x_56); -x_177 = lean_ctor_get(x_53, 1); -lean_inc(x_177); +x_179 = lean_ctor_get(x_53, 1); +lean_inc(x_179); lean_dec(x_53); -x_178 = lean_ctor_get(x_54, 1); -lean_inc(x_178); +x_180 = lean_ctor_get(x_54, 1); +lean_inc(x_180); if (lean_is_exclusive(x_54)) { lean_ctor_release(x_54, 0); lean_ctor_release(x_54, 1); - x_179 = x_54; + x_181 = x_54; } else { lean_dec_ref(x_54); - x_179 = lean_box(0); + x_181 = lean_box(0); } -x_180 = l_Lean_Expr_fvar___override(x_15); -x_181 = 1; +x_182 = l_Lean_Expr_fvar___override(x_15); +x_183 = 1; lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_29); -x_182 = l_Lean_Meta_applySimpResultToProp(x_29, x_180, x_42, x_55, x_181, x_8, x_9, x_10, x_11, x_177); +x_184 = l_Lean_Meta_applySimpResultToProp(x_29, x_182, x_42, x_55, x_183, x_8, x_9, x_10, x_11, x_179); lean_dec(x_42); -if (lean_obj_tag(x_182) == 0) +if (lean_obj_tag(x_184) == 0) { -lean_object* x_183; -x_183 = lean_ctor_get(x_182, 0); -lean_inc(x_183); -if (lean_obj_tag(x_183) == 0) +lean_object* x_185; +x_185 = lean_ctor_get(x_184, 0); +lean_inc(x_185); +if (lean_obj_tag(x_185) == 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_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_dec(x_38); -x_184 = lean_ctor_get(x_182, 1); -lean_inc(x_184); -lean_dec(x_182); -x_185 = lean_box(0); -lean_inc(x_178); -if (lean_is_scalar(x_179)) { - x_186 = lean_alloc_ctor(0, 2, 0); +x_186 = lean_ctor_get(x_184, 1); +lean_inc(x_186); +lean_dec(x_184); +x_187 = lean_box(0); +lean_inc(x_180); +if (lean_is_scalar(x_181)) { + x_188 = lean_alloc_ctor(0, 2, 0); } else { - x_186 = x_179; + x_188 = x_181; } -lean_ctor_set(x_186, 0, x_185); -lean_ctor_set(x_186, 1, x_178); -x_187 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_187, 0, x_186); -lean_ctor_set(x_27, 1, x_178); -x_188 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_188, 0, x_187); -lean_ctor_set(x_188, 1, x_25); -x_189 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_188, 1, x_180); +x_189 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_189, 0, x_188); -x_16 = x_189; -x_17 = x_184; +lean_ctor_set(x_27, 1, x_180); +x_190 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_190, 0, x_189); +lean_ctor_set(x_190, 1, x_25); +x_191 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_191, 0, x_190); +x_16 = x_191; +x_17 = x_186; goto block_24; } 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_dec(x_179); -x_190 = lean_ctor_get(x_183, 0); -lean_inc(x_190); -lean_dec(x_183); -x_191 = lean_ctor_get(x_182, 1); -lean_inc(x_191); -lean_dec(x_182); -x_192 = lean_ctor_get(x_190, 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_dec(x_181); +x_192 = lean_ctor_get(x_185, 0); lean_inc(x_192); -x_193 = lean_ctor_get(x_190, 1); +lean_dec(x_185); +x_193 = lean_ctor_get(x_184, 1); lean_inc(x_193); -lean_dec(x_190); -x_194 = l_Lean_LocalDecl_userName(x_38); +lean_dec(x_184); +x_194 = lean_ctor_get(x_192, 0); +lean_inc(x_194); +x_195 = lean_ctor_get(x_192, 1); +lean_inc(x_195); +lean_dec(x_192); +x_196 = l_Lean_LocalDecl_userName(x_38); lean_dec(x_38); -x_195 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_195, 0, x_194); -lean_ctor_set(x_195, 1, x_193); -lean_ctor_set(x_195, 2, x_192); -x_196 = lean_array_push(x_35, x_195); -lean_ctor_set(x_27, 1, x_178); -lean_ctor_set(x_27, 0, x_196); +x_197 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_197, 0, x_196); +lean_ctor_set(x_197, 1, x_195); +lean_ctor_set(x_197, 2, x_194); +x_198 = lean_array_push(x_35, x_197); +lean_ctor_set(x_27, 1, x_180); +lean_ctor_set(x_27, 0, x_198); lean_inc(x_3); -x_197 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_197, 0, x_3); -lean_ctor_set(x_197, 1, x_25); -x_198 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_198, 0, x_197); -x_16 = x_198; -x_17 = x_191; +x_199 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_199, 0, x_3); +lean_ctor_set(x_199, 1, x_25); +x_200 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_200, 0, x_199); +x_16 = x_200; +x_17 = x_193; goto block_24; } } else { -lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; -lean_dec(x_179); -lean_dec(x_178); +lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; +lean_dec(x_181); +lean_dec(x_180); lean_dec(x_38); lean_free_object(x_27); lean_dec(x_35); @@ -51979,33 +52002,33 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_199 = lean_ctor_get(x_182, 0); -lean_inc(x_199); -x_200 = lean_ctor_get(x_182, 1); -lean_inc(x_200); -if (lean_is_exclusive(x_182)) { - lean_ctor_release(x_182, 0); - lean_ctor_release(x_182, 1); - x_201 = x_182; +x_201 = lean_ctor_get(x_184, 0); +lean_inc(x_201); +x_202 = lean_ctor_get(x_184, 1); +lean_inc(x_202); +if (lean_is_exclusive(x_184)) { + lean_ctor_release(x_184, 0); + lean_ctor_release(x_184, 1); + x_203 = x_184; } else { - lean_dec_ref(x_182); - x_201 = lean_box(0); + lean_dec_ref(x_184); + x_203 = lean_box(0); } -if (lean_is_scalar(x_201)) { - x_202 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_203)) { + x_204 = lean_alloc_ctor(1, 2, 0); } else { - x_202 = x_201; + x_204 = x_203; } -lean_ctor_set(x_202, 0, x_199); -lean_ctor_set(x_202, 1, x_200); -return x_202; +lean_ctor_set(x_204, 0, x_201); +lean_ctor_set(x_204, 1, x_202); +return x_204; } } } } else { -uint8_t x_203; +uint8_t x_205; lean_dec(x_42); lean_dec(x_38); lean_free_object(x_27); @@ -52022,29 +52045,29 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_203 = !lean_is_exclusive(x_53); -if (x_203 == 0) +x_205 = !lean_is_exclusive(x_53); +if (x_205 == 0) { return x_53; } else { -lean_object* x_204; lean_object* x_205; lean_object* x_206; -x_204 = lean_ctor_get(x_53, 0); -x_205 = lean_ctor_get(x_53, 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_53, 0); +x_207 = lean_ctor_get(x_53, 1); +lean_inc(x_207); +lean_inc(x_206); lean_dec(x_53); -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_free_object(x_27); lean_dec(x_36); lean_dec(x_35); @@ -52060,139 +52083,141 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_207 = !lean_is_exclusive(x_37); -if (x_207 == 0) +x_209 = !lean_is_exclusive(x_37); +if (x_209 == 0) { return x_37; } else { -lean_object* x_208; lean_object* x_209; lean_object* x_210; -x_208 = lean_ctor_get(x_37, 0); -x_209 = lean_ctor_get(x_37, 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_37, 0); +x_211 = lean_ctor_get(x_37, 1); +lean_inc(x_211); +lean_inc(x_210); lean_dec(x_37); -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; } } } else { -lean_object* x_211; lean_object* x_212; lean_object* x_213; -x_211 = lean_ctor_get(x_27, 0); -x_212 = lean_ctor_get(x_27, 1); -lean_inc(x_212); -lean_inc(x_211); +lean_object* x_213; lean_object* x_214; lean_object* x_215; +x_213 = lean_ctor_get(x_27, 0); +x_214 = lean_ctor_get(x_27, 1); +lean_inc(x_214); +lean_inc(x_213); lean_dec(x_27); lean_inc(x_8); lean_inc(x_15); -x_213 = l_Lean_FVarId_getDecl(x_15, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_213) == 0) +x_215 = l_Lean_FVarId_getDecl(x_15, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_215) == 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; 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_214 = lean_ctor_get(x_213, 0); -lean_inc(x_214); -x_215 = lean_ctor_get(x_213, 1); -lean_inc(x_215); -lean_dec(x_213); -x_216 = l_Lean_LocalDecl_type(x_214); -x_217 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_216, x_8, x_9, x_10, x_11, x_215); -x_218 = lean_ctor_get(x_217, 0); -lean_inc(x_218); -x_219 = lean_ctor_get(x_217, 1); -lean_inc(x_219); -lean_dec(x_217); -x_220 = lean_ctor_get(x_1, 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; 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_216 = lean_ctor_get(x_215, 0); +lean_inc(x_216); +x_217 = lean_ctor_get(x_215, 1); +lean_inc(x_217); +lean_dec(x_215); +x_218 = l_Lean_LocalDecl_type(x_216); +x_219 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_218, x_8, x_9, x_10, x_11, x_217); +x_220 = lean_ctor_get(x_219, 0); lean_inc(x_220); -x_221 = lean_ctor_get(x_1, 1); +x_221 = lean_ctor_get(x_219, 1); lean_inc(x_221); -x_222 = lean_ctor_get(x_1, 2); +lean_dec(x_219); +x_222 = lean_ctor_get(x_1, 0); lean_inc(x_222); -x_223 = lean_ctor_get(x_1, 3); +x_223 = lean_ctor_get(x_1, 1); lean_inc(x_223); -x_224 = lean_ctor_get(x_1, 4); +x_224 = lean_ctor_get(x_1, 2); lean_inc(x_224); -x_225 = l_Lean_LocalDecl_fvarId(x_214); -x_226 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_226, 0, x_225); -x_227 = l_Lean_Meta_SimpTheoremsArray_eraseTheorem(x_221, x_226); -x_228 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_228, 0, x_220); -lean_ctor_set(x_228, 1, x_227); -lean_ctor_set(x_228, 2, x_222); -lean_ctor_set(x_228, 3, x_223); -lean_ctor_set(x_228, 4, x_224); +x_225 = lean_ctor_get(x_1, 3); +lean_inc(x_225); +x_226 = lean_ctor_get(x_1, 4); +lean_inc(x_226); +x_227 = l_Lean_LocalDecl_fvarId(x_216); +x_228 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_228, 0, x_227); +x_229 = l_Lean_Meta_SimpTheoremsArray_eraseTheorem(x_223, x_228); +x_230 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_230, 0, x_222); +lean_ctor_set(x_230, 1, x_229); +lean_ctor_set(x_230, 2, x_224); +lean_ctor_set(x_230, 3, x_225); +lean_ctor_set(x_230, 4, x_226); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_2); -lean_inc(x_218); -x_229 = l_Lean_Meta_simp(x_218, x_228, x_2, x_212, x_8, x_9, x_10, x_11, x_219); -if (lean_obj_tag(x_229) == 0) +lean_inc(x_220); +x_231 = l_Lean_Meta_simp(x_220, x_230, x_2, x_214, x_8, x_9, x_10, x_11, x_221); +if (lean_obj_tag(x_231) == 0) { -lean_object* x_230; lean_object* x_231; lean_object* x_232; -x_230 = lean_ctor_get(x_229, 0); -lean_inc(x_230); -x_231 = lean_ctor_get(x_230, 0); -lean_inc(x_231); -x_232 = lean_ctor_get(x_231, 1); +lean_object* x_232; lean_object* x_233; lean_object* x_234; +x_232 = lean_ctor_get(x_231, 0); lean_inc(x_232); -if (lean_obj_tag(x_232) == 0) -{ -lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; uint8_t x_238; -lean_dec(x_218); -lean_dec(x_214); -x_233 = lean_ctor_get(x_229, 1); +x_233 = lean_ctor_get(x_232, 0); lean_inc(x_233); -lean_dec(x_229); -x_234 = lean_ctor_get(x_230, 1); +x_234 = lean_ctor_get(x_233, 1); lean_inc(x_234); -if (lean_is_exclusive(x_230)) { - lean_ctor_release(x_230, 0); - lean_ctor_release(x_230, 1); - x_235 = x_230; +if (lean_obj_tag(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; uint8_t x_241; +lean_dec(x_220); +lean_dec(x_216); +x_235 = lean_ctor_get(x_231, 1); +lean_inc(x_235); +lean_dec(x_231); +x_236 = lean_ctor_get(x_232, 1); +lean_inc(x_236); +if (lean_is_exclusive(x_232)) { + lean_ctor_release(x_232, 0); + lean_ctor_release(x_232, 1); + x_237 = x_232; } else { - lean_dec_ref(x_230); - x_235 = lean_box(0); + lean_dec_ref(x_232); + x_237 = lean_box(0); } -x_236 = lean_ctor_get(x_231, 0); -lean_inc(x_236); -x_237 = l_Lean_Meta_Simp_isEqnThmHypothesis_go___closed__2; -x_238 = l_Lean_Expr_isConstOf(x_236, x_237); -lean_dec(x_236); -if (x_238 == 0) +x_238 = lean_ctor_get(x_233, 0); +lean_inc(x_238); +x_239 = l_Lean_Expr_consumeMData(x_238); +lean_dec(x_238); +x_240 = l_Lean_Meta_Simp_isEqnThmHypothesis_go___closed__2; +x_241 = l_Lean_Expr_isConstOf(x_239, x_240); +lean_dec(x_239); +if (x_241 == 0) { -lean_object* x_239; lean_object* x_240; -lean_dec(x_235); +lean_object* x_242; lean_object* x_243; +lean_dec(x_237); lean_free_object(x_26); lean_free_object(x_25); -x_239 = lean_box(0); +x_242 = lean_box(0); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_3); -x_240 = l_Array_forInUnsafe_loop___at_Lean_Meta_simpGoal___spec__1___lambda__1(x_231, x_15, x_211, x_234, x_3, x_29, x_32, x_239, x_8, x_9, x_10, x_11, x_233); -if (lean_obj_tag(x_240) == 0) +x_243 = l_Array_forInUnsafe_loop___at_Lean_Meta_simpGoal___spec__1___lambda__1(x_233, x_15, x_213, x_236, x_3, x_29, x_32, x_242, x_8, x_9, x_10, x_11, x_235); +if (lean_obj_tag(x_243) == 0) { -lean_object* x_241; lean_object* x_242; -x_241 = lean_ctor_get(x_240, 0); -lean_inc(x_241); -x_242 = lean_ctor_get(x_240, 1); -lean_inc(x_242); -lean_dec(x_240); -x_16 = x_241; -x_17 = x_242; +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_16 = x_244; +x_17 = x_245; goto block_24; } else { -lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; +lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -52200,91 +52225,91 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_243 = lean_ctor_get(x_240, 0); -lean_inc(x_243); -x_244 = lean_ctor_get(x_240, 1); -lean_inc(x_244); -if (lean_is_exclusive(x_240)) { - lean_ctor_release(x_240, 0); - lean_ctor_release(x_240, 1); - x_245 = x_240; +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; } else { - lean_dec_ref(x_240); - x_245 = lean_box(0); + lean_dec_ref(x_243); + x_248 = lean_box(0); } -if (lean_is_scalar(x_245)) { - x_246 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_248)) { + x_249 = lean_alloc_ctor(1, 2, 0); } else { - x_246 = x_245; + x_249 = x_248; } -lean_ctor_set(x_246, 0, x_243); -lean_ctor_set(x_246, 1, x_244); -return x_246; +lean_ctor_set(x_249, 0, x_246); +lean_ctor_set(x_249, 1, x_247); +return x_249; } } else { -lean_object* x_247; -lean_dec(x_231); +lean_object* x_250; +lean_dec(x_233); lean_inc(x_29); -x_247 = l_Lean_MVarId_getType(x_29, x_8, x_9, x_10, x_11, x_233); -if (lean_obj_tag(x_247) == 0) +x_250 = l_Lean_MVarId_getType(x_29, x_8, x_9, x_10, x_11, x_235); +if (lean_obj_tag(x_250) == 0) { -lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; -x_248 = lean_ctor_get(x_247, 0); -lean_inc(x_248); -x_249 = lean_ctor_get(x_247, 1); -lean_inc(x_249); -lean_dec(x_247); -x_250 = l_Lean_Expr_fvar___override(x_15); +lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; +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 = l_Lean_Expr_fvar___override(x_15); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -x_251 = l_Lean_Meta_mkFalseElim(x_248, x_250, x_8, x_9, x_10, x_11, x_249); -if (lean_obj_tag(x_251) == 0) +x_254 = l_Lean_Meta_mkFalseElim(x_251, x_253, x_8, x_9, x_10, x_11, x_252); +if (lean_obj_tag(x_254) == 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 = lean_ctor_get(x_251, 0); -lean_inc(x_252); -x_253 = lean_ctor_get(x_251, 1); -lean_inc(x_253); -lean_dec(x_251); -lean_inc(x_29); -x_254 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_29, x_252, x_8, x_9, x_10, x_11, x_253); -x_255 = lean_ctor_get(x_254, 1); +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_255 = lean_ctor_get(x_254, 0); lean_inc(x_255); +x_256 = lean_ctor_get(x_254, 1); +lean_inc(x_256); lean_dec(x_254); -x_256 = lean_box(0); -lean_inc(x_234); -if (lean_is_scalar(x_235)) { - x_257 = lean_alloc_ctor(0, 2, 0); -} else { - x_257 = x_235; -} -lean_ctor_set(x_257, 0, x_256); -lean_ctor_set(x_257, 1, x_234); -x_258 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_258, 0, x_257); -x_259 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_259, 0, x_211); -lean_ctor_set(x_259, 1, x_234); -lean_ctor_set(x_26, 1, 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_25); -x_261 = lean_alloc_ctor(0, 1, 0); +lean_inc(x_29); +x_257 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_29, x_255, x_8, x_9, x_10, x_11, x_256); +x_258 = lean_ctor_get(x_257, 1); +lean_inc(x_258); +lean_dec(x_257); +x_259 = lean_box(0); +lean_inc(x_236); +if (lean_is_scalar(x_237)) { + x_260 = lean_alloc_ctor(0, 2, 0); +} else { + x_260 = x_237; +} +lean_ctor_set(x_260, 0, x_259); +lean_ctor_set(x_260, 1, x_236); +x_261 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_261, 0, x_260); -x_16 = x_261; -x_17 = x_255; +x_262 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_262, 0, x_213); +lean_ctor_set(x_262, 1, x_236); +lean_ctor_set(x_26, 1, x_262); +x_263 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_263, 0, x_261); +lean_ctor_set(x_263, 1, x_25); +x_264 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_264, 0, x_263); +x_16 = x_264; +x_17 = x_258; goto block_24; } else { -lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; -lean_dec(x_235); -lean_dec(x_234); -lean_dec(x_211); +lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; +lean_dec(x_237); +lean_dec(x_236); +lean_dec(x_213); lean_free_object(x_26); lean_dec(x_32); lean_free_object(x_25); @@ -52296,34 +52321,34 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_262 = lean_ctor_get(x_251, 0); -lean_inc(x_262); -x_263 = lean_ctor_get(x_251, 1); -lean_inc(x_263); -if (lean_is_exclusive(x_251)) { - lean_ctor_release(x_251, 0); - lean_ctor_release(x_251, 1); - x_264 = x_251; +x_265 = lean_ctor_get(x_254, 0); +lean_inc(x_265); +x_266 = lean_ctor_get(x_254, 1); +lean_inc(x_266); +if (lean_is_exclusive(x_254)) { + lean_ctor_release(x_254, 0); + lean_ctor_release(x_254, 1); + x_267 = x_254; } else { - lean_dec_ref(x_251); - x_264 = lean_box(0); + lean_dec_ref(x_254); + x_267 = lean_box(0); } -if (lean_is_scalar(x_264)) { - x_265 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_267)) { + x_268 = lean_alloc_ctor(1, 2, 0); } else { - x_265 = x_264; + x_268 = x_267; } -lean_ctor_set(x_265, 0, x_262); -lean_ctor_set(x_265, 1, x_263); -return x_265; +lean_ctor_set(x_268, 0, x_265); +lean_ctor_set(x_268, 1, x_266); +return x_268; } } else { -lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; -lean_dec(x_235); -lean_dec(x_234); -lean_dec(x_211); +lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; +lean_dec(x_237); +lean_dec(x_236); +lean_dec(x_213); lean_free_object(x_26); lean_dec(x_32); lean_free_object(x_25); @@ -52336,147 +52361,147 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_266 = lean_ctor_get(x_247, 0); -lean_inc(x_266); -x_267 = lean_ctor_get(x_247, 1); -lean_inc(x_267); -if (lean_is_exclusive(x_247)) { - lean_ctor_release(x_247, 0); - lean_ctor_release(x_247, 1); - x_268 = x_247; +x_269 = lean_ctor_get(x_250, 0); +lean_inc(x_269); +x_270 = lean_ctor_get(x_250, 1); +lean_inc(x_270); +if (lean_is_exclusive(x_250)) { + lean_ctor_release(x_250, 0); + lean_ctor_release(x_250, 1); + x_271 = x_250; } else { - lean_dec_ref(x_247); - x_268 = lean_box(0); + lean_dec_ref(x_250); + x_271 = lean_box(0); } -if (lean_is_scalar(x_268)) { - x_269 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_271)) { + x_272 = lean_alloc_ctor(1, 2, 0); } else { - x_269 = x_268; + x_272 = x_271; } -lean_ctor_set(x_269, 0, x_266); -lean_ctor_set(x_269, 1, x_267); -return x_269; +lean_ctor_set(x_272, 0, x_269); +lean_ctor_set(x_272, 1, x_270); +return x_272; } } } else { -lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; uint8_t x_275; lean_object* x_276; +lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; uint8_t x_278; lean_object* x_279; +if (lean_is_exclusive(x_234)) { + lean_ctor_release(x_234, 0); + x_273 = x_234; +} else { + lean_dec_ref(x_234); + x_273 = lean_box(0); +} +x_274 = lean_ctor_get(x_231, 1); +lean_inc(x_274); +lean_dec(x_231); +x_275 = lean_ctor_get(x_232, 1); +lean_inc(x_275); if (lean_is_exclusive(x_232)) { lean_ctor_release(x_232, 0); - x_270 = x_232; + lean_ctor_release(x_232, 1); + x_276 = x_232; } else { lean_dec_ref(x_232); - x_270 = lean_box(0); + x_276 = lean_box(0); } -x_271 = lean_ctor_get(x_229, 1); -lean_inc(x_271); -lean_dec(x_229); -x_272 = lean_ctor_get(x_230, 1); -lean_inc(x_272); -if (lean_is_exclusive(x_230)) { - lean_ctor_release(x_230, 0); - lean_ctor_release(x_230, 1); - x_273 = x_230; -} else { - lean_dec_ref(x_230); - x_273 = lean_box(0); -} -x_274 = l_Lean_Expr_fvar___override(x_15); -x_275 = 1; +x_277 = l_Lean_Expr_fvar___override(x_15); +x_278 = 1; lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_29); -x_276 = l_Lean_Meta_applySimpResultToProp(x_29, x_274, x_218, x_231, x_275, x_8, x_9, x_10, x_11, x_271); -lean_dec(x_218); -if (lean_obj_tag(x_276) == 0) +x_279 = l_Lean_Meta_applySimpResultToProp(x_29, x_277, x_220, x_233, x_278, x_8, x_9, x_10, x_11, x_274); +lean_dec(x_220); +if (lean_obj_tag(x_279) == 0) { -lean_object* x_277; -x_277 = lean_ctor_get(x_276, 0); -lean_inc(x_277); -if (lean_obj_tag(x_277) == 0) +lean_object* x_280; +x_280 = lean_ctor_get(x_279, 0); +lean_inc(x_280); +if (lean_obj_tag(x_280) == 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_dec(x_214); -x_278 = lean_ctor_get(x_276, 1); -lean_inc(x_278); -lean_dec(x_276); -x_279 = lean_box(0); -lean_inc(x_272); -if (lean_is_scalar(x_273)) { - x_280 = lean_alloc_ctor(0, 2, 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_dec(x_216); +x_281 = lean_ctor_get(x_279, 1); +lean_inc(x_281); +lean_dec(x_279); +x_282 = lean_box(0); +lean_inc(x_275); +if (lean_is_scalar(x_276)) { + x_283 = lean_alloc_ctor(0, 2, 0); } else { - x_280 = x_273; + x_283 = x_276; } -lean_ctor_set(x_280, 0, x_279); -lean_ctor_set(x_280, 1, x_272); -if (lean_is_scalar(x_270)) { - x_281 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_283, 0, x_282); +lean_ctor_set(x_283, 1, x_275); +if (lean_is_scalar(x_273)) { + x_284 = lean_alloc_ctor(1, 1, 0); } else { - x_281 = x_270; + x_284 = x_273; } -lean_ctor_set(x_281, 0, x_280); -x_282 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_282, 0, x_211); -lean_ctor_set(x_282, 1, x_272); -lean_ctor_set(x_26, 1, x_282); -x_283 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_283, 0, x_281); -lean_ctor_set(x_283, 1, x_25); -x_284 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_284, 0, x_283); -x_16 = x_284; -x_17 = x_278; +x_285 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_285, 0, x_213); +lean_ctor_set(x_285, 1, x_275); +lean_ctor_set(x_26, 1, x_285); +x_286 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_286, 0, x_284); +lean_ctor_set(x_286, 1, x_25); +x_287 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_287, 0, x_286); +x_16 = x_287; +x_17 = x_281; goto block_24; } else { -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_273); -lean_dec(x_270); -x_285 = lean_ctor_get(x_277, 0); -lean_inc(x_285); -lean_dec(x_277); -x_286 = lean_ctor_get(x_276, 1); -lean_inc(x_286); +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_dec(x_276); -x_287 = lean_ctor_get(x_285, 0); -lean_inc(x_287); -x_288 = lean_ctor_get(x_285, 1); +lean_dec(x_273); +x_288 = lean_ctor_get(x_280, 0); lean_inc(x_288); -lean_dec(x_285); -x_289 = l_Lean_LocalDecl_userName(x_214); -lean_dec(x_214); -x_290 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_290, 0, x_289); -lean_ctor_set(x_290, 1, x_288); -lean_ctor_set(x_290, 2, x_287); -x_291 = lean_array_push(x_211, x_290); -x_292 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_292, 0, x_291); -lean_ctor_set(x_292, 1, x_272); -lean_ctor_set(x_26, 1, x_292); +lean_dec(x_280); +x_289 = lean_ctor_get(x_279, 1); +lean_inc(x_289); +lean_dec(x_279); +x_290 = lean_ctor_get(x_288, 0); +lean_inc(x_290); +x_291 = lean_ctor_get(x_288, 1); +lean_inc(x_291); +lean_dec(x_288); +x_292 = l_Lean_LocalDecl_userName(x_216); +lean_dec(x_216); +x_293 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_293, 0, x_292); +lean_ctor_set(x_293, 1, x_291); +lean_ctor_set(x_293, 2, x_290); +x_294 = lean_array_push(x_213, x_293); +x_295 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_295, 0, x_294); +lean_ctor_set(x_295, 1, x_275); +lean_ctor_set(x_26, 1, x_295); lean_inc(x_3); -x_293 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_293, 0, x_3); -lean_ctor_set(x_293, 1, x_25); -x_294 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_294, 0, x_293); -x_16 = x_294; -x_17 = x_286; +x_296 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_296, 0, x_3); +lean_ctor_set(x_296, 1, x_25); +x_297 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_297, 0, x_296); +x_16 = x_297; +x_17 = x_289; goto block_24; } } else { -lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; +lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; +lean_dec(x_276); +lean_dec(x_275); lean_dec(x_273); -lean_dec(x_272); -lean_dec(x_270); -lean_dec(x_214); -lean_dec(x_211); +lean_dec(x_216); +lean_dec(x_213); lean_free_object(x_26); lean_dec(x_32); lean_free_object(x_25); @@ -52488,35 +52513,35 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_295 = lean_ctor_get(x_276, 0); -lean_inc(x_295); -x_296 = lean_ctor_get(x_276, 1); -lean_inc(x_296); -if (lean_is_exclusive(x_276)) { - lean_ctor_release(x_276, 0); - lean_ctor_release(x_276, 1); - x_297 = x_276; +x_298 = lean_ctor_get(x_279, 0); +lean_inc(x_298); +x_299 = lean_ctor_get(x_279, 1); +lean_inc(x_299); +if (lean_is_exclusive(x_279)) { + lean_ctor_release(x_279, 0); + lean_ctor_release(x_279, 1); + x_300 = x_279; } else { - lean_dec_ref(x_276); - x_297 = lean_box(0); + lean_dec_ref(x_279); + x_300 = lean_box(0); } -if (lean_is_scalar(x_297)) { - x_298 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_300)) { + x_301 = lean_alloc_ctor(1, 2, 0); } else { - x_298 = x_297; + x_301 = x_300; } -lean_ctor_set(x_298, 0, x_295); -lean_ctor_set(x_298, 1, x_296); -return x_298; +lean_ctor_set(x_301, 0, x_298); +lean_ctor_set(x_301, 1, x_299); +return x_301; } } } else { -lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; -lean_dec(x_218); -lean_dec(x_214); -lean_dec(x_211); +lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; +lean_dec(x_220); +lean_dec(x_216); +lean_dec(x_213); lean_free_object(x_26); lean_dec(x_32); lean_free_object(x_25); @@ -52529,33 +52554,33 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_299 = lean_ctor_get(x_229, 0); -lean_inc(x_299); -x_300 = lean_ctor_get(x_229, 1); -lean_inc(x_300); -if (lean_is_exclusive(x_229)) { - lean_ctor_release(x_229, 0); - lean_ctor_release(x_229, 1); - x_301 = x_229; +x_302 = lean_ctor_get(x_231, 0); +lean_inc(x_302); +x_303 = lean_ctor_get(x_231, 1); +lean_inc(x_303); +if (lean_is_exclusive(x_231)) { + lean_ctor_release(x_231, 0); + lean_ctor_release(x_231, 1); + x_304 = x_231; } else { - lean_dec_ref(x_229); - x_301 = lean_box(0); + lean_dec_ref(x_231); + x_304 = lean_box(0); } -if (lean_is_scalar(x_301)) { - x_302 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_304)) { + x_305 = lean_alloc_ctor(1, 2, 0); } else { - x_302 = x_301; + x_305 = x_304; } -lean_ctor_set(x_302, 0, x_299); -lean_ctor_set(x_302, 1, x_300); -return x_302; +lean_ctor_set(x_305, 0, x_302); +lean_ctor_set(x_305, 1, x_303); +return x_305; } } else { -lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; -lean_dec(x_212); -lean_dec(x_211); +lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; +lean_dec(x_214); +lean_dec(x_213); lean_free_object(x_26); lean_dec(x_32); lean_free_object(x_25); @@ -52568,152 +52593,154 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_303 = lean_ctor_get(x_213, 0); -lean_inc(x_303); -x_304 = lean_ctor_get(x_213, 1); -lean_inc(x_304); -if (lean_is_exclusive(x_213)) { - lean_ctor_release(x_213, 0); - lean_ctor_release(x_213, 1); - x_305 = x_213; +x_306 = lean_ctor_get(x_215, 0); +lean_inc(x_306); +x_307 = lean_ctor_get(x_215, 1); +lean_inc(x_307); +if (lean_is_exclusive(x_215)) { + lean_ctor_release(x_215, 0); + lean_ctor_release(x_215, 1); + x_308 = x_215; } else { - lean_dec_ref(x_213); - x_305 = lean_box(0); + lean_dec_ref(x_215); + x_308 = lean_box(0); } -if (lean_is_scalar(x_305)) { - x_306 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_308)) { + x_309 = lean_alloc_ctor(1, 2, 0); } else { - x_306 = x_305; + x_309 = x_308; } -lean_ctor_set(x_306, 0, x_303); -lean_ctor_set(x_306, 1, x_304); -return x_306; +lean_ctor_set(x_309, 0, x_306); +lean_ctor_set(x_309, 1, x_307); +return x_309; } } } else { -lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; -x_307 = lean_ctor_get(x_26, 0); -lean_inc(x_307); +lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; +x_310 = lean_ctor_get(x_26, 0); +lean_inc(x_310); lean_dec(x_26); -x_308 = lean_ctor_get(x_27, 0); -lean_inc(x_308); -x_309 = lean_ctor_get(x_27, 1); -lean_inc(x_309); +x_311 = lean_ctor_get(x_27, 0); +lean_inc(x_311); +x_312 = lean_ctor_get(x_27, 1); +lean_inc(x_312); if (lean_is_exclusive(x_27)) { lean_ctor_release(x_27, 0); lean_ctor_release(x_27, 1); - x_310 = x_27; + x_313 = x_27; } else { lean_dec_ref(x_27); - x_310 = lean_box(0); + x_313 = lean_box(0); } lean_inc(x_8); lean_inc(x_15); -x_311 = l_Lean_FVarId_getDecl(x_15, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_311) == 0) +x_314 = l_Lean_FVarId_getDecl(x_15, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_314) == 0) { -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; -x_312 = lean_ctor_get(x_311, 0); -lean_inc(x_312); -x_313 = lean_ctor_get(x_311, 1); -lean_inc(x_313); -lean_dec(x_311); -x_314 = l_Lean_LocalDecl_type(x_312); -x_315 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_314, x_8, x_9, x_10, x_11, x_313); -x_316 = lean_ctor_get(x_315, 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; 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_315 = lean_ctor_get(x_314, 0); +lean_inc(x_315); +x_316 = lean_ctor_get(x_314, 1); lean_inc(x_316); -x_317 = lean_ctor_get(x_315, 1); -lean_inc(x_317); -lean_dec(x_315); -x_318 = lean_ctor_get(x_1, 0); -lean_inc(x_318); -x_319 = lean_ctor_get(x_1, 1); +lean_dec(x_314); +x_317 = l_Lean_LocalDecl_type(x_315); +x_318 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_317, x_8, x_9, x_10, x_11, x_316); +x_319 = lean_ctor_get(x_318, 0); lean_inc(x_319); -x_320 = lean_ctor_get(x_1, 2); +x_320 = lean_ctor_get(x_318, 1); lean_inc(x_320); -x_321 = lean_ctor_get(x_1, 3); +lean_dec(x_318); +x_321 = lean_ctor_get(x_1, 0); lean_inc(x_321); -x_322 = lean_ctor_get(x_1, 4); +x_322 = lean_ctor_get(x_1, 1); lean_inc(x_322); -x_323 = l_Lean_LocalDecl_fvarId(x_312); -x_324 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_324, 0, x_323); -x_325 = l_Lean_Meta_SimpTheoremsArray_eraseTheorem(x_319, x_324); -x_326 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_326, 0, x_318); -lean_ctor_set(x_326, 1, x_325); -lean_ctor_set(x_326, 2, x_320); -lean_ctor_set(x_326, 3, x_321); -lean_ctor_set(x_326, 4, x_322); +x_323 = lean_ctor_get(x_1, 2); +lean_inc(x_323); +x_324 = lean_ctor_get(x_1, 3); +lean_inc(x_324); +x_325 = lean_ctor_get(x_1, 4); +lean_inc(x_325); +x_326 = l_Lean_LocalDecl_fvarId(x_315); +x_327 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_327, 0, x_326); +x_328 = l_Lean_Meta_SimpTheoremsArray_eraseTheorem(x_322, x_327); +x_329 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_329, 0, x_321); +lean_ctor_set(x_329, 1, x_328); +lean_ctor_set(x_329, 2, x_323); +lean_ctor_set(x_329, 3, x_324); +lean_ctor_set(x_329, 4, x_325); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_2); -lean_inc(x_316); -x_327 = l_Lean_Meta_simp(x_316, x_326, x_2, x_309, x_8, x_9, x_10, x_11, x_317); -if (lean_obj_tag(x_327) == 0) -{ -lean_object* x_328; lean_object* x_329; lean_object* x_330; -x_328 = lean_ctor_get(x_327, 0); -lean_inc(x_328); -x_329 = lean_ctor_get(x_328, 0); -lean_inc(x_329); -x_330 = lean_ctor_get(x_329, 1); -lean_inc(x_330); +lean_inc(x_8); +lean_inc(x_2); +lean_inc(x_319); +x_330 = l_Lean_Meta_simp(x_319, x_329, x_2, x_312, x_8, x_9, x_10, x_11, x_320); if (lean_obj_tag(x_330) == 0) { -lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; uint8_t x_336; -lean_dec(x_316); -lean_dec(x_312); -x_331 = lean_ctor_get(x_327, 1); +lean_object* x_331; lean_object* x_332; lean_object* x_333; +x_331 = lean_ctor_get(x_330, 0); lean_inc(x_331); -lean_dec(x_327); -x_332 = lean_ctor_get(x_328, 1); +x_332 = lean_ctor_get(x_331, 0); lean_inc(x_332); -if (lean_is_exclusive(x_328)) { - lean_ctor_release(x_328, 0); - lean_ctor_release(x_328, 1); - x_333 = x_328; +x_333 = lean_ctor_get(x_332, 1); +lean_inc(x_333); +if (lean_obj_tag(x_333) == 0) +{ +lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; uint8_t x_340; +lean_dec(x_319); +lean_dec(x_315); +x_334 = lean_ctor_get(x_330, 1); +lean_inc(x_334); +lean_dec(x_330); +x_335 = lean_ctor_get(x_331, 1); +lean_inc(x_335); +if (lean_is_exclusive(x_331)) { + lean_ctor_release(x_331, 0); + lean_ctor_release(x_331, 1); + x_336 = x_331; } else { - lean_dec_ref(x_328); - x_333 = lean_box(0); + lean_dec_ref(x_331); + x_336 = lean_box(0); } -x_334 = lean_ctor_get(x_329, 0); -lean_inc(x_334); -x_335 = l_Lean_Meta_Simp_isEqnThmHypothesis_go___closed__2; -x_336 = l_Lean_Expr_isConstOf(x_334, x_335); -lean_dec(x_334); -if (x_336 == 0) +x_337 = lean_ctor_get(x_332, 0); +lean_inc(x_337); +x_338 = l_Lean_Expr_consumeMData(x_337); +lean_dec(x_337); +x_339 = l_Lean_Meta_Simp_isEqnThmHypothesis_go___closed__2; +x_340 = l_Lean_Expr_isConstOf(x_338, x_339); +lean_dec(x_338); +if (x_340 == 0) { -lean_object* x_337; lean_object* x_338; -lean_dec(x_333); -lean_dec(x_310); +lean_object* x_341; lean_object* x_342; +lean_dec(x_336); +lean_dec(x_313); lean_free_object(x_25); -x_337 = lean_box(0); +x_341 = lean_box(0); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_3); -x_338 = l_Array_forInUnsafe_loop___at_Lean_Meta_simpGoal___spec__1___lambda__1(x_329, x_15, x_308, x_332, x_3, x_29, x_307, x_337, x_8, x_9, x_10, x_11, x_331); -if (lean_obj_tag(x_338) == 0) +x_342 = l_Array_forInUnsafe_loop___at_Lean_Meta_simpGoal___spec__1___lambda__1(x_332, x_15, x_311, x_335, x_3, x_29, x_310, x_341, x_8, x_9, x_10, x_11, x_334); +if (lean_obj_tag(x_342) == 0) { -lean_object* x_339; lean_object* x_340; -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_16 = x_339; -x_17 = x_340; +lean_object* x_343; lean_object* x_344; +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_16 = x_343; +x_17 = x_344; goto block_24; } else { -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_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -52721,100 +52748,100 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_341 = lean_ctor_get(x_338, 0); -lean_inc(x_341); -x_342 = lean_ctor_get(x_338, 1); -lean_inc(x_342); -if (lean_is_exclusive(x_338)) { - lean_ctor_release(x_338, 0); - lean_ctor_release(x_338, 1); - x_343 = x_338; +x_345 = lean_ctor_get(x_342, 0); +lean_inc(x_345); +x_346 = lean_ctor_get(x_342, 1); +lean_inc(x_346); +if (lean_is_exclusive(x_342)) { + lean_ctor_release(x_342, 0); + lean_ctor_release(x_342, 1); + x_347 = x_342; } else { - lean_dec_ref(x_338); - x_343 = lean_box(0); + lean_dec_ref(x_342); + x_347 = lean_box(0); } -if (lean_is_scalar(x_343)) { - x_344 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_347)) { + x_348 = lean_alloc_ctor(1, 2, 0); } else { - x_344 = x_343; + x_348 = x_347; } -lean_ctor_set(x_344, 0, x_341); -lean_ctor_set(x_344, 1, x_342); -return x_344; +lean_ctor_set(x_348, 0, x_345); +lean_ctor_set(x_348, 1, x_346); +return x_348; } } else { -lean_object* x_345; -lean_dec(x_329); +lean_object* x_349; +lean_dec(x_332); lean_inc(x_29); -x_345 = l_Lean_MVarId_getType(x_29, x_8, x_9, x_10, x_11, x_331); -if (lean_obj_tag(x_345) == 0) -{ -lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; -x_346 = lean_ctor_get(x_345, 0); -lean_inc(x_346); -x_347 = lean_ctor_get(x_345, 1); -lean_inc(x_347); -lean_dec(x_345); -x_348 = l_Lean_Expr_fvar___override(x_15); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -x_349 = l_Lean_Meta_mkFalseElim(x_346, x_348, x_8, x_9, x_10, x_11, x_347); +x_349 = l_Lean_MVarId_getType(x_29, x_8, x_9, x_10, x_11, x_334); if (lean_obj_tag(x_349) == 0) { -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_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; x_350 = lean_ctor_get(x_349, 0); lean_inc(x_350); x_351 = lean_ctor_get(x_349, 1); lean_inc(x_351); lean_dec(x_349); +x_352 = l_Lean_Expr_fvar___override(x_15); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_353 = l_Lean_Meta_mkFalseElim(x_350, x_352, x_8, x_9, x_10, x_11, x_351); +if (lean_obj_tag(x_353) == 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; +x_354 = lean_ctor_get(x_353, 0); +lean_inc(x_354); +x_355 = lean_ctor_get(x_353, 1); +lean_inc(x_355); +lean_dec(x_353); lean_inc(x_29); -x_352 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_29, x_350, x_8, x_9, x_10, x_11, x_351); -x_353 = lean_ctor_get(x_352, 1); -lean_inc(x_353); -lean_dec(x_352); -x_354 = lean_box(0); -lean_inc(x_332); -if (lean_is_scalar(x_333)) { - x_355 = lean_alloc_ctor(0, 2, 0); -} else { - x_355 = x_333; -} -lean_ctor_set(x_355, 0, x_354); -lean_ctor_set(x_355, 1, x_332); -x_356 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_356, 0, x_355); -if (lean_is_scalar(x_310)) { - x_357 = lean_alloc_ctor(0, 2, 0); +x_356 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_29, x_354, x_8, x_9, x_10, x_11, x_355); +x_357 = lean_ctor_get(x_356, 1); +lean_inc(x_357); +lean_dec(x_356); +x_358 = lean_box(0); +lean_inc(x_335); +if (lean_is_scalar(x_336)) { + x_359 = lean_alloc_ctor(0, 2, 0); } else { - x_357 = x_310; + x_359 = x_336; } -lean_ctor_set(x_357, 0, x_308); -lean_ctor_set(x_357, 1, x_332); -x_358 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_358, 0, x_307); -lean_ctor_set(x_358, 1, x_357); -lean_ctor_set(x_25, 1, x_358); -x_359 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_359, 0, x_356); -lean_ctor_set(x_359, 1, x_25); -x_360 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_359, 0, x_358); +lean_ctor_set(x_359, 1, x_335); +x_360 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_360, 0, x_359); -x_16 = x_360; -x_17 = x_353; +if (lean_is_scalar(x_313)) { + x_361 = lean_alloc_ctor(0, 2, 0); +} else { + x_361 = x_313; +} +lean_ctor_set(x_361, 0, x_311); +lean_ctor_set(x_361, 1, x_335); +x_362 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_362, 0, x_310); +lean_ctor_set(x_362, 1, x_361); +lean_ctor_set(x_25, 1, x_362); +x_363 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_363, 0, x_360); +lean_ctor_set(x_363, 1, x_25); +x_364 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_364, 0, x_363); +x_16 = x_364; +x_17 = x_357; goto block_24; } else { -lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; -lean_dec(x_333); -lean_dec(x_332); +lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; +lean_dec(x_336); +lean_dec(x_335); +lean_dec(x_313); +lean_dec(x_311); lean_dec(x_310); -lean_dec(x_308); -lean_dec(x_307); lean_free_object(x_25); lean_dec(x_29); lean_dec(x_11); @@ -52824,36 +52851,36 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_361 = lean_ctor_get(x_349, 0); -lean_inc(x_361); -x_362 = lean_ctor_get(x_349, 1); -lean_inc(x_362); -if (lean_is_exclusive(x_349)) { - lean_ctor_release(x_349, 0); - lean_ctor_release(x_349, 1); - x_363 = x_349; +x_365 = lean_ctor_get(x_353, 0); +lean_inc(x_365); +x_366 = lean_ctor_get(x_353, 1); +lean_inc(x_366); +if (lean_is_exclusive(x_353)) { + lean_ctor_release(x_353, 0); + lean_ctor_release(x_353, 1); + x_367 = x_353; } else { - lean_dec_ref(x_349); - x_363 = lean_box(0); + lean_dec_ref(x_353); + x_367 = lean_box(0); } -if (lean_is_scalar(x_363)) { - x_364 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_367)) { + x_368 = lean_alloc_ctor(1, 2, 0); } else { - x_364 = x_363; + x_368 = x_367; } -lean_ctor_set(x_364, 0, x_361); -lean_ctor_set(x_364, 1, x_362); -return x_364; +lean_ctor_set(x_368, 0, x_365); +lean_ctor_set(x_368, 1, x_366); +return x_368; } } else { -lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; -lean_dec(x_333); -lean_dec(x_332); +lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; +lean_dec(x_336); +lean_dec(x_335); +lean_dec(x_313); +lean_dec(x_311); lean_dec(x_310); -lean_dec(x_308); -lean_dec(x_307); lean_free_object(x_25); lean_dec(x_29); lean_dec(x_15); @@ -52864,206 +52891,165 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_365 = lean_ctor_get(x_345, 0); -lean_inc(x_365); -x_366 = lean_ctor_get(x_345, 1); -lean_inc(x_366); -if (lean_is_exclusive(x_345)) { - lean_ctor_release(x_345, 0); - lean_ctor_release(x_345, 1); - x_367 = x_345; +x_369 = lean_ctor_get(x_349, 0); +lean_inc(x_369); +x_370 = lean_ctor_get(x_349, 1); +lean_inc(x_370); +if (lean_is_exclusive(x_349)) { + lean_ctor_release(x_349, 0); + lean_ctor_release(x_349, 1); + x_371 = x_349; } else { - lean_dec_ref(x_345); - x_367 = lean_box(0); + lean_dec_ref(x_349); + x_371 = lean_box(0); } -if (lean_is_scalar(x_367)) { - x_368 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_371)) { + x_372 = lean_alloc_ctor(1, 2, 0); } else { - x_368 = x_367; + x_372 = x_371; } -lean_ctor_set(x_368, 0, x_365); -lean_ctor_set(x_368, 1, x_366); -return x_368; +lean_ctor_set(x_372, 0, x_369); +lean_ctor_set(x_372, 1, x_370); +return x_372; } } } else { -lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; uint8_t x_374; lean_object* x_375; -if (lean_is_exclusive(x_330)) { - lean_ctor_release(x_330, 0); - x_369 = x_330; +lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; uint8_t x_378; lean_object* x_379; +if (lean_is_exclusive(x_333)) { + lean_ctor_release(x_333, 0); + x_373 = x_333; } else { - lean_dec_ref(x_330); - x_369 = lean_box(0); + lean_dec_ref(x_333); + x_373 = lean_box(0); } -x_370 = lean_ctor_get(x_327, 1); -lean_inc(x_370); -lean_dec(x_327); -x_371 = lean_ctor_get(x_328, 1); -lean_inc(x_371); -if (lean_is_exclusive(x_328)) { - lean_ctor_release(x_328, 0); - lean_ctor_release(x_328, 1); - x_372 = x_328; -} else { - lean_dec_ref(x_328); - x_372 = lean_box(0); -} -x_373 = l_Lean_Expr_fvar___override(x_15); -x_374 = 1; +x_374 = lean_ctor_get(x_330, 1); +lean_inc(x_374); +lean_dec(x_330); +x_375 = lean_ctor_get(x_331, 1); +lean_inc(x_375); +if (lean_is_exclusive(x_331)) { + lean_ctor_release(x_331, 0); + lean_ctor_release(x_331, 1); + x_376 = x_331; +} else { + lean_dec_ref(x_331); + x_376 = lean_box(0); +} +x_377 = l_Lean_Expr_fvar___override(x_15); +x_378 = 1; lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_29); -x_375 = l_Lean_Meta_applySimpResultToProp(x_29, x_373, x_316, x_329, x_374, x_8, x_9, x_10, x_11, x_370); -lean_dec(x_316); -if (lean_obj_tag(x_375) == 0) +x_379 = l_Lean_Meta_applySimpResultToProp(x_29, x_377, x_319, x_332, x_378, x_8, x_9, x_10, x_11, x_374); +lean_dec(x_319); +if (lean_obj_tag(x_379) == 0) { -lean_object* x_376; -x_376 = lean_ctor_get(x_375, 0); -lean_inc(x_376); -if (lean_obj_tag(x_376) == 0) +lean_object* x_380; +x_380 = lean_ctor_get(x_379, 0); +lean_inc(x_380); +if (lean_obj_tag(x_380) == 0) { -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_dec(x_312); -x_377 = lean_ctor_get(x_375, 1); -lean_inc(x_377); -lean_dec(x_375); -x_378 = lean_box(0); -lean_inc(x_371); -if (lean_is_scalar(x_372)) { - x_379 = lean_alloc_ctor(0, 2, 0); -} else { - x_379 = x_372; -} -lean_ctor_set(x_379, 0, x_378); -lean_ctor_set(x_379, 1, x_371); -if (lean_is_scalar(x_369)) { - x_380 = lean_alloc_ctor(1, 1, 0); -} else { - x_380 = x_369; -} -lean_ctor_set(x_380, 0, x_379); -if (lean_is_scalar(x_310)) { - x_381 = lean_alloc_ctor(0, 2, 0); -} else { - x_381 = x_310; -} -lean_ctor_set(x_381, 0, x_308); -lean_ctor_set(x_381, 1, x_371); -x_382 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_382, 0, x_307); -lean_ctor_set(x_382, 1, x_381); -lean_ctor_set(x_25, 1, x_382); -x_383 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_383, 0, x_380); -lean_ctor_set(x_383, 1, x_25); -x_384 = lean_alloc_ctor(0, 1, 0); +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_315); +x_381 = lean_ctor_get(x_379, 1); +lean_inc(x_381); +lean_dec(x_379); +x_382 = lean_box(0); +lean_inc(x_375); +if (lean_is_scalar(x_376)) { + x_383 = lean_alloc_ctor(0, 2, 0); +} else { + x_383 = x_376; +} +lean_ctor_set(x_383, 0, x_382); +lean_ctor_set(x_383, 1, x_375); +if (lean_is_scalar(x_373)) { + x_384 = lean_alloc_ctor(1, 1, 0); +} else { + x_384 = x_373; +} lean_ctor_set(x_384, 0, x_383); -x_16 = x_384; -x_17 = x_377; +if (lean_is_scalar(x_313)) { + x_385 = lean_alloc_ctor(0, 2, 0); +} else { + x_385 = x_313; +} +lean_ctor_set(x_385, 0, x_311); +lean_ctor_set(x_385, 1, x_375); +x_386 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_386, 0, x_310); +lean_ctor_set(x_386, 1, x_385); +lean_ctor_set(x_25, 1, x_386); +x_387 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_387, 0, x_384); +lean_ctor_set(x_387, 1, x_25); +x_388 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_388, 0, x_387); +x_16 = x_388; +x_17 = x_381; goto block_24; } else { -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_dec(x_372); -lean_dec(x_369); -x_385 = lean_ctor_get(x_376, 0); -lean_inc(x_385); +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_dec(x_376); -x_386 = lean_ctor_get(x_375, 1); -lean_inc(x_386); -lean_dec(x_375); -x_387 = lean_ctor_get(x_385, 0); -lean_inc(x_387); -x_388 = lean_ctor_get(x_385, 1); -lean_inc(x_388); -lean_dec(x_385); -x_389 = l_Lean_LocalDecl_userName(x_312); -lean_dec(x_312); -x_390 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_390, 0, x_389); -lean_ctor_set(x_390, 1, x_388); -lean_ctor_set(x_390, 2, x_387); -x_391 = lean_array_push(x_308, x_390); -if (lean_is_scalar(x_310)) { - x_392 = lean_alloc_ctor(0, 2, 0); -} else { - x_392 = x_310; -} -lean_ctor_set(x_392, 0, x_391); -lean_ctor_set(x_392, 1, x_371); -x_393 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_393, 0, x_307); -lean_ctor_set(x_393, 1, x_392); -lean_ctor_set(x_25, 1, x_393); +lean_dec(x_373); +x_389 = lean_ctor_get(x_380, 0); +lean_inc(x_389); +lean_dec(x_380); +x_390 = lean_ctor_get(x_379, 1); +lean_inc(x_390); +lean_dec(x_379); +x_391 = lean_ctor_get(x_389, 0); +lean_inc(x_391); +x_392 = lean_ctor_get(x_389, 1); +lean_inc(x_392); +lean_dec(x_389); +x_393 = l_Lean_LocalDecl_userName(x_315); +lean_dec(x_315); +x_394 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_394, 0, x_393); +lean_ctor_set(x_394, 1, x_392); +lean_ctor_set(x_394, 2, x_391); +x_395 = lean_array_push(x_311, x_394); +if (lean_is_scalar(x_313)) { + x_396 = lean_alloc_ctor(0, 2, 0); +} else { + x_396 = x_313; +} +lean_ctor_set(x_396, 0, x_395); +lean_ctor_set(x_396, 1, x_375); +x_397 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_397, 0, x_310); +lean_ctor_set(x_397, 1, x_396); +lean_ctor_set(x_25, 1, x_397); lean_inc(x_3); -x_394 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_394, 0, x_3); -lean_ctor_set(x_394, 1, x_25); -x_395 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_395, 0, x_394); -x_16 = x_395; -x_17 = x_386; +x_398 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_398, 0, x_3); +lean_ctor_set(x_398, 1, x_25); +x_399 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_399, 0, x_398); +x_16 = x_399; +x_17 = x_390; goto block_24; } } else { -lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; -lean_dec(x_372); -lean_dec(x_371); -lean_dec(x_369); -lean_dec(x_312); -lean_dec(x_310); -lean_dec(x_308); -lean_dec(x_307); -lean_free_object(x_25); -lean_dec(x_29); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_396 = lean_ctor_get(x_375, 0); -lean_inc(x_396); -x_397 = lean_ctor_get(x_375, 1); -lean_inc(x_397); -if (lean_is_exclusive(x_375)) { - lean_ctor_release(x_375, 0); - lean_ctor_release(x_375, 1); - x_398 = x_375; -} else { - lean_dec_ref(x_375); - x_398 = lean_box(0); -} -if (lean_is_scalar(x_398)) { - x_399 = lean_alloc_ctor(1, 2, 0); -} else { - x_399 = x_398; -} -lean_ctor_set(x_399, 0, x_396); -lean_ctor_set(x_399, 1, x_397); -return x_399; -} -} -} -else -{ lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; -lean_dec(x_316); -lean_dec(x_312); +lean_dec(x_376); +lean_dec(x_375); +lean_dec(x_373); +lean_dec(x_315); +lean_dec(x_313); +lean_dec(x_311); lean_dec(x_310); -lean_dec(x_308); -lean_dec(x_307); lean_free_object(x_25); lean_dec(x_29); -lean_dec(x_15); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -53071,16 +53057,16 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_400 = lean_ctor_get(x_327, 0); +x_400 = lean_ctor_get(x_379, 0); lean_inc(x_400); -x_401 = lean_ctor_get(x_327, 1); +x_401 = lean_ctor_get(x_379, 1); lean_inc(x_401); -if (lean_is_exclusive(x_327)) { - lean_ctor_release(x_327, 0); - lean_ctor_release(x_327, 1); - x_402 = x_327; +if (lean_is_exclusive(x_379)) { + lean_ctor_release(x_379, 0); + lean_ctor_release(x_379, 1); + x_402 = x_379; } else { - lean_dec_ref(x_327); + lean_dec_ref(x_379); x_402 = lean_box(0); } if (lean_is_scalar(x_402)) { @@ -53093,13 +53079,15 @@ lean_ctor_set(x_403, 1, x_401); return x_403; } } +} else { lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; +lean_dec(x_319); +lean_dec(x_315); +lean_dec(x_313); +lean_dec(x_311); lean_dec(x_310); -lean_dec(x_309); -lean_dec(x_308); -lean_dec(x_307); lean_free_object(x_25); lean_dec(x_29); lean_dec(x_15); @@ -53110,16 +53098,16 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_404 = lean_ctor_get(x_311, 0); +x_404 = lean_ctor_get(x_330, 0); lean_inc(x_404); -x_405 = lean_ctor_get(x_311, 1); +x_405 = lean_ctor_get(x_330, 1); lean_inc(x_405); -if (lean_is_exclusive(x_311)) { - lean_ctor_release(x_311, 0); - lean_ctor_release(x_311, 1); - x_406 = x_311; +if (lean_is_exclusive(x_330)) { + lean_ctor_release(x_330, 0); + lean_ctor_release(x_330, 1); + x_406 = x_330; } else { - lean_dec_ref(x_311); + lean_dec_ref(x_330); x_406 = lean_box(0); } if (lean_is_scalar(x_406)) { @@ -53132,140 +53120,181 @@ lean_ctor_set(x_407, 1, x_405); return x_407; } } -} else { -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; -x_408 = lean_ctor_get(x_25, 0); +lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; +lean_dec(x_313); +lean_dec(x_312); +lean_dec(x_311); +lean_dec(x_310); +lean_free_object(x_25); +lean_dec(x_29); +lean_dec(x_15); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_408 = lean_ctor_get(x_314, 0); lean_inc(x_408); -lean_dec(x_25); -x_409 = lean_ctor_get(x_26, 0); +x_409 = lean_ctor_get(x_314, 1); lean_inc(x_409); +if (lean_is_exclusive(x_314)) { + lean_ctor_release(x_314, 0); + lean_ctor_release(x_314, 1); + x_410 = x_314; +} else { + lean_dec_ref(x_314); + x_410 = lean_box(0); +} +if (lean_is_scalar(x_410)) { + x_411 = lean_alloc_ctor(1, 2, 0); +} else { + x_411 = x_410; +} +lean_ctor_set(x_411, 0, x_408); +lean_ctor_set(x_411, 1, x_409); +return x_411; +} +} +} +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; +x_412 = lean_ctor_get(x_25, 0); +lean_inc(x_412); +lean_dec(x_25); +x_413 = lean_ctor_get(x_26, 0); +lean_inc(x_413); if (lean_is_exclusive(x_26)) { lean_ctor_release(x_26, 0); lean_ctor_release(x_26, 1); - x_410 = x_26; + x_414 = x_26; } else { lean_dec_ref(x_26); - x_410 = lean_box(0); + x_414 = lean_box(0); } -x_411 = lean_ctor_get(x_27, 0); -lean_inc(x_411); -x_412 = lean_ctor_get(x_27, 1); -lean_inc(x_412); +x_415 = lean_ctor_get(x_27, 0); +lean_inc(x_415); +x_416 = lean_ctor_get(x_27, 1); +lean_inc(x_416); if (lean_is_exclusive(x_27)) { lean_ctor_release(x_27, 0); lean_ctor_release(x_27, 1); - x_413 = x_27; + x_417 = x_27; } else { lean_dec_ref(x_27); - x_413 = lean_box(0); + x_417 = lean_box(0); } lean_inc(x_8); lean_inc(x_15); -x_414 = l_Lean_FVarId_getDecl(x_15, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_414) == 0) +x_418 = l_Lean_FVarId_getDecl(x_15, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_418) == 0) { -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; -x_415 = lean_ctor_get(x_414, 0); -lean_inc(x_415); -x_416 = lean_ctor_get(x_414, 1); -lean_inc(x_416); -lean_dec(x_414); -x_417 = l_Lean_LocalDecl_type(x_415); -x_418 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_417, x_8, x_9, x_10, x_11, x_416); +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_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_421 = lean_ctor_get(x_1, 0); -lean_inc(x_421); -x_422 = lean_ctor_get(x_1, 1); -lean_inc(x_422); -x_423 = lean_ctor_get(x_1, 2); +x_421 = l_Lean_LocalDecl_type(x_419); +x_422 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_421, x_8, x_9, x_10, x_11, x_420); +x_423 = lean_ctor_get(x_422, 0); lean_inc(x_423); -x_424 = lean_ctor_get(x_1, 3); +x_424 = lean_ctor_get(x_422, 1); lean_inc(x_424); -x_425 = lean_ctor_get(x_1, 4); +lean_dec(x_422); +x_425 = lean_ctor_get(x_1, 0); lean_inc(x_425); -x_426 = l_Lean_LocalDecl_fvarId(x_415); -x_427 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_427, 0, x_426); -x_428 = l_Lean_Meta_SimpTheoremsArray_eraseTheorem(x_422, x_427); -x_429 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_429, 0, x_421); -lean_ctor_set(x_429, 1, x_428); -lean_ctor_set(x_429, 2, x_423); -lean_ctor_set(x_429, 3, x_424); -lean_ctor_set(x_429, 4, x_425); +x_426 = lean_ctor_get(x_1, 1); +lean_inc(x_426); +x_427 = lean_ctor_get(x_1, 2); +lean_inc(x_427); +x_428 = lean_ctor_get(x_1, 3); +lean_inc(x_428); +x_429 = lean_ctor_get(x_1, 4); +lean_inc(x_429); +x_430 = l_Lean_LocalDecl_fvarId(x_419); +x_431 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_431, 0, x_430); +x_432 = l_Lean_Meta_SimpTheoremsArray_eraseTheorem(x_426, x_431); +x_433 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_433, 0, x_425); +lean_ctor_set(x_433, 1, x_432); +lean_ctor_set(x_433, 2, x_427); +lean_ctor_set(x_433, 3, x_428); +lean_ctor_set(x_433, 4, x_429); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_2); -lean_inc(x_419); -x_430 = l_Lean_Meta_simp(x_419, x_429, x_2, x_412, x_8, x_9, x_10, x_11, x_420); -if (lean_obj_tag(x_430) == 0) +lean_inc(x_423); +x_434 = l_Lean_Meta_simp(x_423, x_433, x_2, x_416, x_8, x_9, x_10, x_11, x_424); +if (lean_obj_tag(x_434) == 0) { -lean_object* x_431; lean_object* x_432; lean_object* x_433; -x_431 = lean_ctor_get(x_430, 0); -lean_inc(x_431); -x_432 = lean_ctor_get(x_431, 0); -lean_inc(x_432); -x_433 = lean_ctor_get(x_432, 1); -lean_inc(x_433); -if (lean_obj_tag(x_433) == 0) +lean_object* x_435; lean_object* x_436; lean_object* x_437; +x_435 = lean_ctor_get(x_434, 0); +lean_inc(x_435); +x_436 = lean_ctor_get(x_435, 0); +lean_inc(x_436); +x_437 = lean_ctor_get(x_436, 1); +lean_inc(x_437); +if (lean_obj_tag(x_437) == 0) { -lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; uint8_t x_439; +lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; uint8_t x_444; +lean_dec(x_423); lean_dec(x_419); -lean_dec(x_415); -x_434 = lean_ctor_get(x_430, 1); -lean_inc(x_434); -lean_dec(x_430); -x_435 = lean_ctor_get(x_431, 1); -lean_inc(x_435); -if (lean_is_exclusive(x_431)) { - lean_ctor_release(x_431, 0); - lean_ctor_release(x_431, 1); - x_436 = x_431; +x_438 = lean_ctor_get(x_434, 1); +lean_inc(x_438); +lean_dec(x_434); +x_439 = lean_ctor_get(x_435, 1); +lean_inc(x_439); +if (lean_is_exclusive(x_435)) { + lean_ctor_release(x_435, 0); + lean_ctor_release(x_435, 1); + x_440 = x_435; } else { - lean_dec_ref(x_431); - x_436 = lean_box(0); + lean_dec_ref(x_435); + x_440 = lean_box(0); } -x_437 = lean_ctor_get(x_432, 0); -lean_inc(x_437); -x_438 = l_Lean_Meta_Simp_isEqnThmHypothesis_go___closed__2; -x_439 = l_Lean_Expr_isConstOf(x_437, x_438); -lean_dec(x_437); -if (x_439 == 0) +x_441 = lean_ctor_get(x_436, 0); +lean_inc(x_441); +x_442 = l_Lean_Expr_consumeMData(x_441); +lean_dec(x_441); +x_443 = l_Lean_Meta_Simp_isEqnThmHypothesis_go___closed__2; +x_444 = l_Lean_Expr_isConstOf(x_442, x_443); +lean_dec(x_442); +if (x_444 == 0) { -lean_object* x_440; lean_object* x_441; -lean_dec(x_436); -lean_dec(x_413); -lean_dec(x_410); -x_440 = lean_box(0); +lean_object* x_445; lean_object* x_446; +lean_dec(x_440); +lean_dec(x_417); +lean_dec(x_414); +x_445 = lean_box(0); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_3); -x_441 = l_Array_forInUnsafe_loop___at_Lean_Meta_simpGoal___spec__1___lambda__1(x_432, x_15, x_411, x_435, x_3, x_408, x_409, x_440, x_8, x_9, x_10, x_11, x_434); -if (lean_obj_tag(x_441) == 0) +x_446 = l_Array_forInUnsafe_loop___at_Lean_Meta_simpGoal___spec__1___lambda__1(x_436, x_15, x_415, x_439, x_3, x_412, x_413, x_445, x_8, x_9, x_10, x_11, x_438); +if (lean_obj_tag(x_446) == 0) { -lean_object* x_442; lean_object* x_443; -x_442 = lean_ctor_get(x_441, 0); -lean_inc(x_442); -x_443 = lean_ctor_get(x_441, 1); -lean_inc(x_443); -lean_dec(x_441); -x_16 = x_442; -x_17 = x_443; +lean_object* x_447; lean_object* x_448; +x_447 = lean_ctor_get(x_446, 0); +lean_inc(x_447); +x_448 = lean_ctor_get(x_446, 1); +lean_inc(x_448); +lean_dec(x_446); +x_16 = x_447; +x_17 = x_448; goto block_24; } else { -lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; +lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -53273,108 +53302,108 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_444 = lean_ctor_get(x_441, 0); -lean_inc(x_444); -x_445 = lean_ctor_get(x_441, 1); -lean_inc(x_445); -if (lean_is_exclusive(x_441)) { - lean_ctor_release(x_441, 0); - lean_ctor_release(x_441, 1); - x_446 = x_441; +x_449 = lean_ctor_get(x_446, 0); +lean_inc(x_449); +x_450 = lean_ctor_get(x_446, 1); +lean_inc(x_450); +if (lean_is_exclusive(x_446)) { + lean_ctor_release(x_446, 0); + lean_ctor_release(x_446, 1); + x_451 = x_446; } else { - lean_dec_ref(x_441); - x_446 = lean_box(0); + lean_dec_ref(x_446); + x_451 = lean_box(0); } -if (lean_is_scalar(x_446)) { - x_447 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_451)) { + x_452 = lean_alloc_ctor(1, 2, 0); } else { - x_447 = x_446; + x_452 = x_451; } -lean_ctor_set(x_447, 0, x_444); -lean_ctor_set(x_447, 1, x_445); -return x_447; +lean_ctor_set(x_452, 0, x_449); +lean_ctor_set(x_452, 1, x_450); +return x_452; } } else { -lean_object* x_448; -lean_dec(x_432); -lean_inc(x_408); -x_448 = l_Lean_MVarId_getType(x_408, x_8, x_9, x_10, x_11, x_434); -if (lean_obj_tag(x_448) == 0) +lean_object* x_453; +lean_dec(x_436); +lean_inc(x_412); +x_453 = l_Lean_MVarId_getType(x_412, x_8, x_9, x_10, x_11, x_438); +if (lean_obj_tag(x_453) == 0) { -lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; -x_449 = lean_ctor_get(x_448, 0); -lean_inc(x_449); -x_450 = lean_ctor_get(x_448, 1); -lean_inc(x_450); -lean_dec(x_448); -x_451 = l_Lean_Expr_fvar___override(x_15); +lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; +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 = l_Lean_Expr_fvar___override(x_15); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -x_452 = l_Lean_Meta_mkFalseElim(x_449, x_451, x_8, x_9, x_10, x_11, x_450); -if (lean_obj_tag(x_452) == 0) +x_457 = l_Lean_Meta_mkFalseElim(x_454, x_456, x_8, x_9, x_10, x_11, x_455); +if (lean_obj_tag(x_457) == 0) { -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; -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); -lean_inc(x_408); -x_455 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_408, x_453, x_8, x_9, x_10, x_11, x_454); -x_456 = lean_ctor_get(x_455, 1); -lean_inc(x_456); -lean_dec(x_455); -x_457 = lean_box(0); -lean_inc(x_435); -if (lean_is_scalar(x_436)) { - x_458 = lean_alloc_ctor(0, 2, 0); -} else { - x_458 = x_436; -} -lean_ctor_set(x_458, 0, x_457); -lean_ctor_set(x_458, 1, x_435); -x_459 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_459, 0, x_458); -if (lean_is_scalar(x_413)) { - x_460 = lean_alloc_ctor(0, 2, 0); +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; +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); +lean_inc(x_412); +x_460 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_412, x_458, x_8, x_9, x_10, x_11, x_459); +x_461 = lean_ctor_get(x_460, 1); +lean_inc(x_461); +lean_dec(x_460); +x_462 = lean_box(0); +lean_inc(x_439); +if (lean_is_scalar(x_440)) { + x_463 = lean_alloc_ctor(0, 2, 0); } else { - x_460 = x_413; + x_463 = x_440; } -lean_ctor_set(x_460, 0, x_411); -lean_ctor_set(x_460, 1, x_435); -if (lean_is_scalar(x_410)) { - x_461 = lean_alloc_ctor(0, 2, 0); -} else { - x_461 = x_410; -} -lean_ctor_set(x_461, 0, x_409); -lean_ctor_set(x_461, 1, x_460); -x_462 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_462, 0, x_408); -lean_ctor_set(x_462, 1, x_461); -x_463 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_463, 0, x_459); -lean_ctor_set(x_463, 1, x_462); -x_464 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_463, 0, x_462); +lean_ctor_set(x_463, 1, x_439); +x_464 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_464, 0, x_463); -x_16 = x_464; -x_17 = x_456; +if (lean_is_scalar(x_417)) { + x_465 = lean_alloc_ctor(0, 2, 0); +} else { + x_465 = x_417; +} +lean_ctor_set(x_465, 0, x_415); +lean_ctor_set(x_465, 1, x_439); +if (lean_is_scalar(x_414)) { + x_466 = lean_alloc_ctor(0, 2, 0); +} else { + x_466 = x_414; +} +lean_ctor_set(x_466, 0, x_413); +lean_ctor_set(x_466, 1, x_465); +x_467 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_467, 0, x_412); +lean_ctor_set(x_467, 1, x_466); +x_468 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_468, 0, x_464); +lean_ctor_set(x_468, 1, x_467); +x_469 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_469, 0, x_468); +x_16 = x_469; +x_17 = x_461; goto block_24; } else { -lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; -lean_dec(x_436); -lean_dec(x_435); +lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; +lean_dec(x_440); +lean_dec(x_439); +lean_dec(x_417); +lean_dec(x_415); +lean_dec(x_414); lean_dec(x_413); -lean_dec(x_411); -lean_dec(x_410); -lean_dec(x_409); -lean_dec(x_408); +lean_dec(x_412); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -53382,38 +53411,38 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_465 = lean_ctor_get(x_452, 0); -lean_inc(x_465); -x_466 = lean_ctor_get(x_452, 1); -lean_inc(x_466); -if (lean_is_exclusive(x_452)) { - lean_ctor_release(x_452, 0); - lean_ctor_release(x_452, 1); - x_467 = x_452; +x_470 = lean_ctor_get(x_457, 0); +lean_inc(x_470); +x_471 = lean_ctor_get(x_457, 1); +lean_inc(x_471); +if (lean_is_exclusive(x_457)) { + lean_ctor_release(x_457, 0); + lean_ctor_release(x_457, 1); + x_472 = x_457; } else { - lean_dec_ref(x_452); - x_467 = lean_box(0); + lean_dec_ref(x_457); + x_472 = lean_box(0); } -if (lean_is_scalar(x_467)) { - x_468 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_472)) { + x_473 = lean_alloc_ctor(1, 2, 0); } else { - x_468 = x_467; + x_473 = x_472; } -lean_ctor_set(x_468, 0, x_465); -lean_ctor_set(x_468, 1, x_466); -return x_468; +lean_ctor_set(x_473, 0, x_470); +lean_ctor_set(x_473, 1, x_471); +return x_473; } } else { -lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; -lean_dec(x_436); -lean_dec(x_435); +lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; +lean_dec(x_440); +lean_dec(x_439); +lean_dec(x_417); +lean_dec(x_415); +lean_dec(x_414); lean_dec(x_413); -lean_dec(x_411); -lean_dec(x_410); -lean_dec(x_409); -lean_dec(x_408); +lean_dec(x_412); lean_dec(x_15); lean_dec(x_11); lean_dec(x_10); @@ -53422,177 +53451,177 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_469 = lean_ctor_get(x_448, 0); -lean_inc(x_469); -x_470 = lean_ctor_get(x_448, 1); -lean_inc(x_470); -if (lean_is_exclusive(x_448)) { - lean_ctor_release(x_448, 0); - lean_ctor_release(x_448, 1); - x_471 = x_448; +x_474 = lean_ctor_get(x_453, 0); +lean_inc(x_474); +x_475 = lean_ctor_get(x_453, 1); +lean_inc(x_475); +if (lean_is_exclusive(x_453)) { + lean_ctor_release(x_453, 0); + lean_ctor_release(x_453, 1); + x_476 = x_453; } else { - lean_dec_ref(x_448); - x_471 = lean_box(0); + lean_dec_ref(x_453); + x_476 = lean_box(0); } -if (lean_is_scalar(x_471)) { - x_472 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_476)) { + x_477 = lean_alloc_ctor(1, 2, 0); } else { - x_472 = x_471; + x_477 = x_476; } -lean_ctor_set(x_472, 0, x_469); -lean_ctor_set(x_472, 1, x_470); -return x_472; +lean_ctor_set(x_477, 0, x_474); +lean_ctor_set(x_477, 1, x_475); +return x_477; } } } else { -lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; uint8_t x_478; lean_object* x_479; -if (lean_is_exclusive(x_433)) { - lean_ctor_release(x_433, 0); - x_473 = x_433; +lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; lean_object* x_482; uint8_t x_483; lean_object* x_484; +if (lean_is_exclusive(x_437)) { + lean_ctor_release(x_437, 0); + x_478 = x_437; } else { - lean_dec_ref(x_433); - x_473 = lean_box(0); + lean_dec_ref(x_437); + x_478 = lean_box(0); } -x_474 = lean_ctor_get(x_430, 1); -lean_inc(x_474); -lean_dec(x_430); -x_475 = lean_ctor_get(x_431, 1); -lean_inc(x_475); -if (lean_is_exclusive(x_431)) { - lean_ctor_release(x_431, 0); - lean_ctor_release(x_431, 1); - x_476 = x_431; +x_479 = lean_ctor_get(x_434, 1); +lean_inc(x_479); +lean_dec(x_434); +x_480 = lean_ctor_get(x_435, 1); +lean_inc(x_480); +if (lean_is_exclusive(x_435)) { + lean_ctor_release(x_435, 0); + lean_ctor_release(x_435, 1); + x_481 = x_435; } else { - lean_dec_ref(x_431); - x_476 = lean_box(0); + lean_dec_ref(x_435); + x_481 = lean_box(0); } -x_477 = l_Lean_Expr_fvar___override(x_15); -x_478 = 1; +x_482 = l_Lean_Expr_fvar___override(x_15); +x_483 = 1; lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -lean_inc(x_408); -x_479 = l_Lean_Meta_applySimpResultToProp(x_408, x_477, x_419, x_432, x_478, x_8, x_9, x_10, x_11, x_474); -lean_dec(x_419); -if (lean_obj_tag(x_479) == 0) +lean_inc(x_412); +x_484 = l_Lean_Meta_applySimpResultToProp(x_412, x_482, x_423, x_436, x_483, x_8, x_9, x_10, x_11, x_479); +lean_dec(x_423); +if (lean_obj_tag(x_484) == 0) { -lean_object* x_480; -x_480 = lean_ctor_get(x_479, 0); -lean_inc(x_480); -if (lean_obj_tag(x_480) == 0) +lean_object* x_485; +x_485 = lean_ctor_get(x_484, 0); +lean_inc(x_485); +if (lean_obj_tag(x_485) == 0) { -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_dec(x_415); -x_481 = lean_ctor_get(x_479, 1); -lean_inc(x_481); -lean_dec(x_479); -x_482 = lean_box(0); -lean_inc(x_475); -if (lean_is_scalar(x_476)) { - x_483 = lean_alloc_ctor(0, 2, 0); +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_dec(x_419); +x_486 = lean_ctor_get(x_484, 1); +lean_inc(x_486); +lean_dec(x_484); +x_487 = lean_box(0); +lean_inc(x_480); +if (lean_is_scalar(x_481)) { + x_488 = lean_alloc_ctor(0, 2, 0); } else { - x_483 = x_476; + x_488 = x_481; } -lean_ctor_set(x_483, 0, x_482); -lean_ctor_set(x_483, 1, x_475); -if (lean_is_scalar(x_473)) { - x_484 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_488, 0, x_487); +lean_ctor_set(x_488, 1, x_480); +if (lean_is_scalar(x_478)) { + x_489 = lean_alloc_ctor(1, 1, 0); } else { - x_484 = x_473; + x_489 = x_478; } -lean_ctor_set(x_484, 0, x_483); -if (lean_is_scalar(x_413)) { - x_485 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_489, 0, x_488); +if (lean_is_scalar(x_417)) { + x_490 = lean_alloc_ctor(0, 2, 0); } else { - x_485 = x_413; + x_490 = x_417; } -lean_ctor_set(x_485, 0, x_411); -lean_ctor_set(x_485, 1, x_475); -if (lean_is_scalar(x_410)) { - x_486 = lean_alloc_ctor(0, 2, 0); -} else { - x_486 = x_410; -} -lean_ctor_set(x_486, 0, x_409); -lean_ctor_set(x_486, 1, x_485); -x_487 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_487, 0, x_408); -lean_ctor_set(x_487, 1, x_486); -x_488 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_488, 0, x_484); -lean_ctor_set(x_488, 1, x_487); -x_489 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_489, 0, x_488); -x_16 = x_489; -x_17 = x_481; +lean_ctor_set(x_490, 0, x_415); +lean_ctor_set(x_490, 1, x_480); +if (lean_is_scalar(x_414)) { + x_491 = lean_alloc_ctor(0, 2, 0); +} else { + x_491 = x_414; +} +lean_ctor_set(x_491, 0, x_413); +lean_ctor_set(x_491, 1, x_490); +x_492 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_492, 0, x_412); +lean_ctor_set(x_492, 1, x_491); +x_493 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_493, 0, x_489); +lean_ctor_set(x_493, 1, x_492); +x_494 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_494, 0, x_493); +x_16 = x_494; +x_17 = x_486; goto block_24; } 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_dec(x_476); -lean_dec(x_473); -x_490 = lean_ctor_get(x_480, 0); -lean_inc(x_490); -lean_dec(x_480); -x_491 = lean_ctor_get(x_479, 1); -lean_inc(x_491); -lean_dec(x_479); -x_492 = lean_ctor_get(x_490, 0); -lean_inc(x_492); -x_493 = lean_ctor_get(x_490, 1); -lean_inc(x_493); -lean_dec(x_490); -x_494 = l_Lean_LocalDecl_userName(x_415); -lean_dec(x_415); -x_495 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_495, 0, x_494); -lean_ctor_set(x_495, 1, x_493); -lean_ctor_set(x_495, 2, x_492); -x_496 = lean_array_push(x_411, x_495); -if (lean_is_scalar(x_413)) { - x_497 = lean_alloc_ctor(0, 2, 0); -} else { - x_497 = x_413; -} -lean_ctor_set(x_497, 0, x_496); -lean_ctor_set(x_497, 1, x_475); -if (lean_is_scalar(x_410)) { - x_498 = lean_alloc_ctor(0, 2, 0); -} else { - x_498 = x_410; -} -lean_ctor_set(x_498, 0, x_409); -lean_ctor_set(x_498, 1, x_497); -x_499 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_499, 0, x_408); -lean_ctor_set(x_499, 1, x_498); +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_dec(x_481); +lean_dec(x_478); +x_495 = lean_ctor_get(x_485, 0); +lean_inc(x_495); +lean_dec(x_485); +x_496 = lean_ctor_get(x_484, 1); +lean_inc(x_496); +lean_dec(x_484); +x_497 = lean_ctor_get(x_495, 0); +lean_inc(x_497); +x_498 = lean_ctor_get(x_495, 1); +lean_inc(x_498); +lean_dec(x_495); +x_499 = l_Lean_LocalDecl_userName(x_419); +lean_dec(x_419); +x_500 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_500, 0, x_499); +lean_ctor_set(x_500, 1, x_498); +lean_ctor_set(x_500, 2, x_497); +x_501 = lean_array_push(x_415, x_500); +if (lean_is_scalar(x_417)) { + x_502 = lean_alloc_ctor(0, 2, 0); +} else { + x_502 = x_417; +} +lean_ctor_set(x_502, 0, x_501); +lean_ctor_set(x_502, 1, x_480); +if (lean_is_scalar(x_414)) { + x_503 = lean_alloc_ctor(0, 2, 0); +} else { + x_503 = x_414; +} +lean_ctor_set(x_503, 0, x_413); +lean_ctor_set(x_503, 1, x_502); +x_504 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_504, 0, x_412); +lean_ctor_set(x_504, 1, x_503); lean_inc(x_3); -x_500 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_500, 0, x_3); -lean_ctor_set(x_500, 1, x_499); -x_501 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_501, 0, x_500); -x_16 = x_501; -x_17 = x_491; +x_505 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_505, 0, x_3); +lean_ctor_set(x_505, 1, x_504); +x_506 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_506, 0, x_505); +x_16 = x_506; +x_17 = x_496; goto block_24; } } else { -lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; -lean_dec(x_476); -lean_dec(x_475); -lean_dec(x_473); +lean_object* x_507; lean_object* x_508; lean_object* x_509; lean_object* x_510; +lean_dec(x_481); +lean_dec(x_480); +lean_dec(x_478); +lean_dec(x_419); +lean_dec(x_417); lean_dec(x_415); +lean_dec(x_414); lean_dec(x_413); -lean_dec(x_411); -lean_dec(x_410); -lean_dec(x_409); -lean_dec(x_408); +lean_dec(x_412); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -53600,39 +53629,39 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_502 = lean_ctor_get(x_479, 0); -lean_inc(x_502); -x_503 = lean_ctor_get(x_479, 1); -lean_inc(x_503); -if (lean_is_exclusive(x_479)) { - lean_ctor_release(x_479, 0); - lean_ctor_release(x_479, 1); - x_504 = x_479; +x_507 = lean_ctor_get(x_484, 0); +lean_inc(x_507); +x_508 = lean_ctor_get(x_484, 1); +lean_inc(x_508); +if (lean_is_exclusive(x_484)) { + lean_ctor_release(x_484, 0); + lean_ctor_release(x_484, 1); + x_509 = x_484; } else { - lean_dec_ref(x_479); - x_504 = lean_box(0); + lean_dec_ref(x_484); + x_509 = lean_box(0); } -if (lean_is_scalar(x_504)) { - x_505 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_509)) { + x_510 = lean_alloc_ctor(1, 2, 0); } else { - x_505 = x_504; + x_510 = x_509; } -lean_ctor_set(x_505, 0, x_502); -lean_ctor_set(x_505, 1, x_503); -return x_505; +lean_ctor_set(x_510, 0, x_507); +lean_ctor_set(x_510, 1, x_508); +return x_510; } } } else { -lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; +lean_object* x_511; lean_object* x_512; lean_object* x_513; lean_object* x_514; +lean_dec(x_423); lean_dec(x_419); +lean_dec(x_417); lean_dec(x_415); +lean_dec(x_414); lean_dec(x_413); -lean_dec(x_411); -lean_dec(x_410); -lean_dec(x_409); -lean_dec(x_408); +lean_dec(x_412); lean_dec(x_15); lean_dec(x_11); lean_dec(x_10); @@ -53641,37 +53670,37 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_506 = lean_ctor_get(x_430, 0); -lean_inc(x_506); -x_507 = lean_ctor_get(x_430, 1); -lean_inc(x_507); -if (lean_is_exclusive(x_430)) { - lean_ctor_release(x_430, 0); - lean_ctor_release(x_430, 1); - x_508 = x_430; +x_511 = lean_ctor_get(x_434, 0); +lean_inc(x_511); +x_512 = lean_ctor_get(x_434, 1); +lean_inc(x_512); +if (lean_is_exclusive(x_434)) { + lean_ctor_release(x_434, 0); + lean_ctor_release(x_434, 1); + x_513 = x_434; } else { - lean_dec_ref(x_430); - x_508 = lean_box(0); + lean_dec_ref(x_434); + x_513 = lean_box(0); } -if (lean_is_scalar(x_508)) { - x_509 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_513)) { + x_514 = lean_alloc_ctor(1, 2, 0); } else { - x_509 = x_508; + x_514 = x_513; } -lean_ctor_set(x_509, 0, x_506); -lean_ctor_set(x_509, 1, x_507); -return x_509; +lean_ctor_set(x_514, 0, x_511); +lean_ctor_set(x_514, 1, x_512); +return x_514; } } else { -lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; +lean_object* x_515; lean_object* x_516; lean_object* x_517; lean_object* x_518; +lean_dec(x_417); +lean_dec(x_416); +lean_dec(x_415); +lean_dec(x_414); lean_dec(x_413); lean_dec(x_412); -lean_dec(x_411); -lean_dec(x_410); -lean_dec(x_409); -lean_dec(x_408); lean_dec(x_15); lean_dec(x_11); lean_dec(x_10); @@ -53680,26 +53709,26 @@ lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_510 = lean_ctor_get(x_414, 0); -lean_inc(x_510); -x_511 = lean_ctor_get(x_414, 1); -lean_inc(x_511); -if (lean_is_exclusive(x_414)) { - lean_ctor_release(x_414, 0); - lean_ctor_release(x_414, 1); - x_512 = x_414; +x_515 = lean_ctor_get(x_418, 0); +lean_inc(x_515); +x_516 = lean_ctor_get(x_418, 1); +lean_inc(x_516); +if (lean_is_exclusive(x_418)) { + lean_ctor_release(x_418, 0); + lean_ctor_release(x_418, 1); + x_517 = x_418; } else { - lean_dec_ref(x_414); - x_512 = lean_box(0); + lean_dec_ref(x_418); + x_517 = lean_box(0); } -if (lean_is_scalar(x_512)) { - x_513 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_517)) { + x_518 = lean_alloc_ctor(1, 2, 0); } else { - x_513 = x_512; + x_518 = x_517; } -lean_ctor_set(x_513, 0, x_510); -lean_ctor_set(x_513, 1, x_511); -return x_513; +lean_ctor_set(x_518, 0, x_515); +lean_ctor_set(x_518, 1, x_516); +return x_518; } } block_24: @@ -55586,7 +55615,7 @@ lean_inc(x_29); x_30 = lean_ctor_get(x_28, 1); lean_inc(x_30); lean_dec(x_28); -x_31 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_29, x_7, x_8, x_9, x_10, x_30); +x_31 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_29, x_7, x_8, x_9, x_10, x_30); x_32 = lean_ctor_get(x_31, 0); lean_inc(x_32); x_33 = lean_ctor_get(x_31, 1); @@ -55611,117 +55640,119 @@ lean_dec(x_35); x_38 = !lean_is_exclusive(x_36); if (x_38 == 0) { -lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; +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_36, 0); x_40 = lean_ctor_get(x_36, 1); -x_41 = l_Lean_Meta_Simp_isEqnThmHypothesis_go___closed__2; -x_42 = l_Lean_Expr_isConstOf(x_39, x_41); -if (x_42 == 0) +x_41 = l_Lean_Expr_consumeMData(x_39); +x_42 = l_Lean_Meta_Simp_isEqnThmHypothesis_go___closed__2; +x_43 = l_Lean_Expr_isConstOf(x_41, x_42); +lean_dec(x_41); +if (x_43 == 0) { -lean_object* x_43; lean_object* x_44; +lean_object* x_44; lean_object* x_45; lean_free_object(x_36); lean_free_object(x_24); -x_43 = lean_box(0); +x_44 = lean_box(0); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_2); -x_44 = l_Array_forInUnsafe_loop___at_Lean_Meta_dsimpGoal___spec__1___lambda__1(x_39, x_32, x_14, x_40, x_2, x_26, x_43, x_7, x_8, x_9, x_10, x_37); +x_45 = l_Array_forInUnsafe_loop___at_Lean_Meta_dsimpGoal___spec__1___lambda__1(x_39, x_32, x_14, x_40, x_2, x_26, x_44, x_7, x_8, x_9, x_10, x_37); lean_dec(x_32); -if (lean_obj_tag(x_44) == 0) +if (lean_obj_tag(x_45) == 0) { -lean_object* x_45; lean_object* x_46; -x_45 = lean_ctor_get(x_44, 0); -lean_inc(x_45); -x_46 = lean_ctor_get(x_44, 1); +lean_object* x_46; lean_object* x_47; +x_46 = lean_ctor_get(x_45, 0); lean_inc(x_46); -lean_dec(x_44); -x_15 = x_45; -x_16 = x_46; +x_47 = lean_ctor_get(x_45, 1); +lean_inc(x_47); +lean_dec(x_45); +x_15 = x_46; +x_16 = x_47; goto block_23; } else { -uint8_t x_47; +uint8_t x_48; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_2); lean_dec(x_1); -x_47 = !lean_is_exclusive(x_44); -if (x_47 == 0) +x_48 = !lean_is_exclusive(x_45); +if (x_48 == 0) { -return x_44; +return x_45; } else { -lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_48 = lean_ctor_get(x_44, 0); -x_49 = lean_ctor_get(x_44, 1); +lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_49 = lean_ctor_get(x_45, 0); +x_50 = lean_ctor_get(x_45, 1); +lean_inc(x_50); lean_inc(x_49); -lean_inc(x_48); -lean_dec(x_44); -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; +lean_dec(x_45); +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_51; +lean_object* x_52; lean_dec(x_39); lean_dec(x_32); lean_inc(x_26); -x_51 = l_Lean_MVarId_getType(x_26, x_7, x_8, x_9, x_10, x_37); -if (lean_obj_tag(x_51) == 0) +x_52 = l_Lean_MVarId_getType(x_26, x_7, x_8, x_9, x_10, x_37); +if (lean_obj_tag(x_52) == 0) { -lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; -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; +x_53 = lean_ctor_get(x_52, 0); lean_inc(x_53); -lean_dec(x_51); -x_54 = l_Lean_Expr_fvar___override(x_14); +x_54 = lean_ctor_get(x_52, 1); +lean_inc(x_54); +lean_dec(x_52); +x_55 = l_Lean_Expr_fvar___override(x_14); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); -x_55 = l_Lean_Meta_mkFalseElim(x_52, x_54, x_7, x_8, x_9, x_10, x_53); -if (lean_obj_tag(x_55) == 0) +x_56 = l_Lean_Meta_mkFalseElim(x_53, x_55, x_7, x_8, x_9, x_10, x_54); +if (lean_obj_tag(x_56) == 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; lean_object* x_62; lean_object* x_63; -x_56 = lean_ctor_get(x_55, 0); -lean_inc(x_56); -x_57 = lean_ctor_get(x_55, 1); +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); -lean_dec(x_55); +x_58 = lean_ctor_get(x_56, 1); +lean_inc(x_58); +lean_dec(x_56); lean_inc(x_26); -x_58 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_26, x_56, x_7, x_8, x_9, x_10, x_57); -x_59 = lean_ctor_get(x_58, 1); -lean_inc(x_59); -lean_dec(x_58); -x_60 = lean_box(0); +x_59 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_26, x_57, x_7, x_8, x_9, x_10, x_58); +x_60 = lean_ctor_get(x_59, 1); +lean_inc(x_60); +lean_dec(x_59); +x_61 = lean_box(0); lean_inc(x_40); -lean_ctor_set(x_36, 0, x_60); -x_61 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_61, 0, x_36); +lean_ctor_set(x_36, 0, x_61); +x_62 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_62, 0, x_36); lean_ctor_set(x_24, 1, x_40); -x_62 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_62, 0, x_61); -lean_ctor_set(x_62, 1, x_24); -x_63 = lean_alloc_ctor(0, 1, 0); +x_63 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_63, 0, x_62); -x_15 = x_63; -x_16 = x_59; +lean_ctor_set(x_63, 1, x_24); +x_64 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_64, 0, x_63); +x_15 = x_64; +x_16 = x_60; goto block_23; } else { -uint8_t x_64; +uint8_t x_65; lean_free_object(x_36); lean_dec(x_40); lean_free_object(x_24); @@ -55732,29 +55763,29 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_2); lean_dec(x_1); -x_64 = !lean_is_exclusive(x_55); -if (x_64 == 0) +x_65 = !lean_is_exclusive(x_56); +if (x_65 == 0) { -return x_55; +return x_56; } else { -lean_object* x_65; lean_object* x_66; lean_object* x_67; -x_65 = lean_ctor_get(x_55, 0); -x_66 = lean_ctor_get(x_55, 1); +lean_object* x_66; lean_object* x_67; lean_object* x_68; +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_inc(x_65); -lean_dec(x_55); -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; +lean_dec(x_56); +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 { -uint8_t x_68; +uint8_t x_69; lean_free_object(x_36); lean_dec(x_40); lean_free_object(x_24); @@ -55766,147 +55797,149 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_2); lean_dec(x_1); -x_68 = !lean_is_exclusive(x_51); -if (x_68 == 0) +x_69 = !lean_is_exclusive(x_52); +if (x_69 == 0) { -return x_51; +return x_52; } else { -lean_object* x_69; lean_object* x_70; lean_object* x_71; -x_69 = lean_ctor_get(x_51, 0); -x_70 = lean_ctor_get(x_51, 1); +lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_70 = lean_ctor_get(x_52, 0); +x_71 = lean_ctor_get(x_52, 1); +lean_inc(x_71); lean_inc(x_70); -lean_inc(x_69); -lean_dec(x_51); -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; +lean_dec(x_52); +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_72; lean_object* x_73; lean_object* x_74; uint8_t x_75; -x_72 = lean_ctor_get(x_36, 0); -x_73 = lean_ctor_get(x_36, 1); +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; uint8_t x_77; +x_73 = lean_ctor_get(x_36, 0); +x_74 = lean_ctor_get(x_36, 1); +lean_inc(x_74); lean_inc(x_73); -lean_inc(x_72); lean_dec(x_36); -x_74 = l_Lean_Meta_Simp_isEqnThmHypothesis_go___closed__2; -x_75 = l_Lean_Expr_isConstOf(x_72, x_74); -if (x_75 == 0) +x_75 = l_Lean_Expr_consumeMData(x_73); +x_76 = l_Lean_Meta_Simp_isEqnThmHypothesis_go___closed__2; +x_77 = l_Lean_Expr_isConstOf(x_75, x_76); +lean_dec(x_75); +if (x_77 == 0) { -lean_object* x_76; lean_object* x_77; +lean_object* x_78; lean_object* x_79; lean_free_object(x_24); -x_76 = lean_box(0); +x_78 = lean_box(0); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_2); -x_77 = l_Array_forInUnsafe_loop___at_Lean_Meta_dsimpGoal___spec__1___lambda__1(x_72, x_32, x_14, x_73, x_2, x_26, x_76, x_7, x_8, x_9, x_10, x_37); +x_79 = l_Array_forInUnsafe_loop___at_Lean_Meta_dsimpGoal___spec__1___lambda__1(x_73, x_32, x_14, x_74, x_2, x_26, x_78, x_7, x_8, x_9, x_10, x_37); lean_dec(x_32); -if (lean_obj_tag(x_77) == 0) +if (lean_obj_tag(x_79) == 0) { -lean_object* x_78; lean_object* x_79; -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_15 = x_78; -x_16 = x_79; +lean_object* x_80; lean_object* x_81; +x_80 = lean_ctor_get(x_79, 0); +lean_inc(x_80); +x_81 = lean_ctor_get(x_79, 1); +lean_inc(x_81); +lean_dec(x_79); +x_15 = x_80; +x_16 = x_81; goto block_23; } else { -lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_2); lean_dec(x_1); -x_80 = lean_ctor_get(x_77, 0); -lean_inc(x_80); -x_81 = lean_ctor_get(x_77, 1); -lean_inc(x_81); -if (lean_is_exclusive(x_77)) { - lean_ctor_release(x_77, 0); - lean_ctor_release(x_77, 1); - x_82 = x_77; +x_82 = lean_ctor_get(x_79, 0); +lean_inc(x_82); +x_83 = lean_ctor_get(x_79, 1); +lean_inc(x_83); +if (lean_is_exclusive(x_79)) { + lean_ctor_release(x_79, 0); + lean_ctor_release(x_79, 1); + x_84 = x_79; } else { - lean_dec_ref(x_77); - x_82 = lean_box(0); + lean_dec_ref(x_79); + x_84 = lean_box(0); } -if (lean_is_scalar(x_82)) { - x_83 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_84)) { + x_85 = lean_alloc_ctor(1, 2, 0); } else { - x_83 = x_82; + x_85 = x_84; } -lean_ctor_set(x_83, 0, x_80); -lean_ctor_set(x_83, 1, x_81); -return x_83; +lean_ctor_set(x_85, 0, x_82); +lean_ctor_set(x_85, 1, x_83); +return x_85; } } else { -lean_object* x_84; -lean_dec(x_72); +lean_object* x_86; +lean_dec(x_73); lean_dec(x_32); lean_inc(x_26); -x_84 = l_Lean_MVarId_getType(x_26, x_7, x_8, x_9, x_10, x_37); -if (lean_obj_tag(x_84) == 0) +x_86 = l_Lean_MVarId_getType(x_26, x_7, x_8, x_9, x_10, x_37); +if (lean_obj_tag(x_86) == 0) { -lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; -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 = l_Lean_Expr_fvar___override(x_14); +lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; +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 = l_Lean_Expr_fvar___override(x_14); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); -x_88 = l_Lean_Meta_mkFalseElim(x_85, x_87, x_7, x_8, x_9, x_10, x_86); -if (lean_obj_tag(x_88) == 0) +x_90 = l_Lean_Meta_mkFalseElim(x_87, x_89, x_7, x_8, x_9, x_10, x_88); +if (lean_obj_tag(x_90) == 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; lean_object* x_96; lean_object* x_97; -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); -lean_inc(x_26); -x_91 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_26, x_89, x_7, x_8, x_9, x_10, x_90); -x_92 = lean_ctor_get(x_91, 1); +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_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_91); -x_93 = lean_box(0); -lean_inc(x_73); -x_94 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_94, 0, x_93); -lean_ctor_set(x_94, 1, x_73); -x_95 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_95, 0, x_94); -lean_ctor_set(x_24, 1, x_73); +lean_dec(x_90); +lean_inc(x_26); +x_93 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_26, x_91, x_7, x_8, x_9, x_10, x_92); +x_94 = lean_ctor_get(x_93, 1); +lean_inc(x_94); +lean_dec(x_93); +x_95 = lean_box(0); +lean_inc(x_74); x_96 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_96, 0, x_95); -lean_ctor_set(x_96, 1, x_24); -x_97 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_96, 1, x_74); +x_97 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_97, 0, x_96); -x_15 = x_97; -x_16 = x_92; +lean_ctor_set(x_24, 1, x_74); +x_98 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_98, 0, x_97); +lean_ctor_set(x_98, 1, x_24); +x_99 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_99, 0, x_98); +x_15 = x_99; +x_16 = x_94; goto block_23; } else { -lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; -lean_dec(x_73); +lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; +lean_dec(x_74); lean_free_object(x_24); lean_dec(x_26); lean_dec(x_10); @@ -55915,32 +55948,32 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_2); lean_dec(x_1); -x_98 = lean_ctor_get(x_88, 0); -lean_inc(x_98); -x_99 = lean_ctor_get(x_88, 1); -lean_inc(x_99); -if (lean_is_exclusive(x_88)) { - lean_ctor_release(x_88, 0); - lean_ctor_release(x_88, 1); - x_100 = x_88; +x_100 = lean_ctor_get(x_90, 0); +lean_inc(x_100); +x_101 = lean_ctor_get(x_90, 1); +lean_inc(x_101); +if (lean_is_exclusive(x_90)) { + lean_ctor_release(x_90, 0); + lean_ctor_release(x_90, 1); + x_102 = x_90; } else { - lean_dec_ref(x_88); - x_100 = lean_box(0); + lean_dec_ref(x_90); + x_102 = lean_box(0); } -if (lean_is_scalar(x_100)) { - x_101 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_102)) { + x_103 = lean_alloc_ctor(1, 2, 0); } else { - x_101 = x_100; + x_103 = x_102; } -lean_ctor_set(x_101, 0, x_98); -lean_ctor_set(x_101, 1, x_99); -return x_101; +lean_ctor_set(x_103, 0, x_100); +lean_ctor_set(x_103, 1, x_101); +return x_103; } } else { -lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; -lean_dec(x_73); +lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; +lean_dec(x_74); lean_free_object(x_24); lean_dec(x_26); lean_dec(x_14); @@ -55950,33 +55983,33 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_2); lean_dec(x_1); -x_102 = lean_ctor_get(x_84, 0); -lean_inc(x_102); -x_103 = lean_ctor_get(x_84, 1); -lean_inc(x_103); -if (lean_is_exclusive(x_84)) { - lean_ctor_release(x_84, 0); - lean_ctor_release(x_84, 1); - x_104 = x_84; +x_104 = lean_ctor_get(x_86, 0); +lean_inc(x_104); +x_105 = lean_ctor_get(x_86, 1); +lean_inc(x_105); +if (lean_is_exclusive(x_86)) { + lean_ctor_release(x_86, 0); + lean_ctor_release(x_86, 1); + x_106 = x_86; } else { - lean_dec_ref(x_84); - x_104 = lean_box(0); + lean_dec_ref(x_86); + x_106 = lean_box(0); } -if (lean_is_scalar(x_104)) { - x_105 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_106)) { + x_107 = lean_alloc_ctor(1, 2, 0); } else { - x_105 = x_104; + x_107 = x_106; } -lean_ctor_set(x_105, 0, x_102); -lean_ctor_set(x_105, 1, x_103); -return x_105; +lean_ctor_set(x_107, 0, x_104); +lean_ctor_set(x_107, 1, x_105); +return x_107; } } } } else { -uint8_t x_106; +uint8_t x_108; lean_dec(x_32); lean_free_object(x_24); lean_dec(x_26); @@ -55987,29 +56020,29 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_2); lean_dec(x_1); -x_106 = !lean_is_exclusive(x_35); -if (x_106 == 0) +x_108 = !lean_is_exclusive(x_35); +if (x_108 == 0) { return x_35; } else { -lean_object* x_107; lean_object* x_108; lean_object* x_109; -x_107 = lean_ctor_get(x_35, 0); -x_108 = lean_ctor_get(x_35, 1); -lean_inc(x_108); -lean_inc(x_107); +lean_object* x_109; lean_object* x_110; lean_object* x_111; +x_109 = lean_ctor_get(x_35, 0); +x_110 = lean_ctor_get(x_35, 1); +lean_inc(x_110); +lean_inc(x_109); lean_dec(x_35); -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; +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 { -uint8_t x_110; +uint8_t x_112; lean_free_object(x_24); lean_dec(x_26); lean_dec(x_14); @@ -56019,231 +56052,233 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_2); lean_dec(x_1); -x_110 = !lean_is_exclusive(x_28); -if (x_110 == 0) +x_112 = !lean_is_exclusive(x_28); +if (x_112 == 0) { return x_28; } else { -lean_object* x_111; lean_object* x_112; lean_object* x_113; -x_111 = lean_ctor_get(x_28, 0); -x_112 = lean_ctor_get(x_28, 1); -lean_inc(x_112); -lean_inc(x_111); +lean_object* x_113; lean_object* x_114; lean_object* x_115; +x_113 = lean_ctor_get(x_28, 0); +x_114 = lean_ctor_get(x_28, 1); +lean_inc(x_114); +lean_inc(x_113); lean_dec(x_28); -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; +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; } } } else { -lean_object* x_114; lean_object* x_115; -x_114 = lean_ctor_get(x_24, 0); -lean_inc(x_114); +lean_object* x_116; lean_object* x_117; +x_116 = lean_ctor_get(x_24, 0); +lean_inc(x_116); lean_dec(x_24); lean_inc(x_7); lean_inc(x_14); -x_115 = l_Lean_FVarId_getType(x_14, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_115) == 0) +x_117 = l_Lean_FVarId_getType(x_14, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_117) == 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; -x_116 = lean_ctor_get(x_115, 0); -lean_inc(x_116); -x_117 = lean_ctor_get(x_115, 1); -lean_inc(x_117); -lean_dec(x_115); -x_118 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_116, x_7, x_8, x_9, x_10, x_117); -x_119 = lean_ctor_get(x_118, 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; +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_Meta_transform___at___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimp___spec__1___closed__2; +lean_dec(x_117); +x_120 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_118, x_7, x_8, x_9, x_10, x_119); +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 = l_Lean_Meta_transform___at___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimp___spec__1___closed__2; lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_1); -lean_inc(x_119); -x_122 = l_Lean_Meta_dsimp(x_119, x_1, x_121, x_7, x_8, x_9, x_10, x_120); -if (lean_obj_tag(x_122) == 0) +lean_inc(x_121); +x_124 = l_Lean_Meta_dsimp(x_121, x_1, x_123, x_7, x_8, x_9, x_10, x_122); +if (lean_obj_tag(x_124) == 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, 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_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; +x_125 = lean_ctor_get(x_124, 0); lean_inc(x_125); -x_126 = lean_ctor_get(x_123, 1); +x_126 = lean_ctor_get(x_124, 1); lean_inc(x_126); -if (lean_is_exclusive(x_123)) { - lean_ctor_release(x_123, 0); - lean_ctor_release(x_123, 1); - x_127 = x_123; +lean_dec(x_124); +x_127 = lean_ctor_get(x_125, 0); +lean_inc(x_127); +x_128 = lean_ctor_get(x_125, 1); +lean_inc(x_128); +if (lean_is_exclusive(x_125)) { + lean_ctor_release(x_125, 0); + lean_ctor_release(x_125, 1); + x_129 = x_125; } else { - lean_dec_ref(x_123); - x_127 = lean_box(0); + lean_dec_ref(x_125); + x_129 = lean_box(0); } -x_128 = l_Lean_Meta_Simp_isEqnThmHypothesis_go___closed__2; -x_129 = l_Lean_Expr_isConstOf(x_125, x_128); -if (x_129 == 0) +x_130 = l_Lean_Expr_consumeMData(x_127); +x_131 = l_Lean_Meta_Simp_isEqnThmHypothesis_go___closed__2; +x_132 = l_Lean_Expr_isConstOf(x_130, x_131); +lean_dec(x_130); +if (x_132 == 0) { -lean_object* x_130; lean_object* x_131; -lean_dec(x_127); -x_130 = lean_box(0); +lean_object* x_133; lean_object* x_134; +lean_dec(x_129); +x_133 = lean_box(0); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_2); -x_131 = l_Array_forInUnsafe_loop___at_Lean_Meta_dsimpGoal___spec__1___lambda__1(x_125, x_119, x_14, x_126, x_2, x_114, x_130, x_7, x_8, x_9, x_10, x_124); -lean_dec(x_119); -if (lean_obj_tag(x_131) == 0) +x_134 = l_Array_forInUnsafe_loop___at_Lean_Meta_dsimpGoal___spec__1___lambda__1(x_127, x_121, x_14, x_128, x_2, x_116, x_133, x_7, x_8, x_9, x_10, x_126); +lean_dec(x_121); +if (lean_obj_tag(x_134) == 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_15 = x_132; -x_16 = x_133; +lean_object* x_135; lean_object* x_136; +x_135 = lean_ctor_get(x_134, 0); +lean_inc(x_135); +x_136 = lean_ctor_get(x_134, 1); +lean_inc(x_136); +lean_dec(x_134); +x_15 = x_135; +x_16 = x_136; goto block_23; } else { -lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; +lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_2); lean_dec(x_1); -x_134 = lean_ctor_get(x_131, 0); -lean_inc(x_134); -x_135 = lean_ctor_get(x_131, 1); -lean_inc(x_135); -if (lean_is_exclusive(x_131)) { - lean_ctor_release(x_131, 0); - lean_ctor_release(x_131, 1); - x_136 = x_131; +x_137 = lean_ctor_get(x_134, 0); +lean_inc(x_137); +x_138 = lean_ctor_get(x_134, 1); +lean_inc(x_138); +if (lean_is_exclusive(x_134)) { + lean_ctor_release(x_134, 0); + lean_ctor_release(x_134, 1); + x_139 = x_134; } else { - lean_dec_ref(x_131); - x_136 = lean_box(0); + lean_dec_ref(x_134); + x_139 = lean_box(0); } -if (lean_is_scalar(x_136)) { - x_137 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_139)) { + x_140 = lean_alloc_ctor(1, 2, 0); } else { - x_137 = x_136; + x_140 = x_139; } -lean_ctor_set(x_137, 0, x_134); -lean_ctor_set(x_137, 1, x_135); -return x_137; +lean_ctor_set(x_140, 0, x_137); +lean_ctor_set(x_140, 1, x_138); +return x_140; } } else { -lean_object* x_138; -lean_dec(x_125); -lean_dec(x_119); -lean_inc(x_114); -x_138 = l_Lean_MVarId_getType(x_114, x_7, x_8, x_9, x_10, x_124); -if (lean_obj_tag(x_138) == 0) +lean_object* x_141; +lean_dec(x_127); +lean_dec(x_121); +lean_inc(x_116); +x_141 = l_Lean_MVarId_getType(x_116, x_7, x_8, x_9, x_10, x_126); +if (lean_obj_tag(x_141) == 0) { -lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; -x_139 = lean_ctor_get(x_138, 0); -lean_inc(x_139); -x_140 = lean_ctor_get(x_138, 1); -lean_inc(x_140); -lean_dec(x_138); -x_141 = l_Lean_Expr_fvar___override(x_14); +lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; +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 = l_Lean_Expr_fvar___override(x_14); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); -x_142 = l_Lean_Meta_mkFalseElim(x_139, x_141, x_7, x_8, x_9, x_10, x_140); -if (lean_obj_tag(x_142) == 0) +x_145 = l_Lean_Meta_mkFalseElim(x_142, x_144, x_7, x_8, x_9, x_10, x_143); +if (lean_obj_tag(x_145) == 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; -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); -lean_inc(x_114); -x_145 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_114, x_143, x_7, x_8, x_9, x_10, x_144); -x_146 = lean_ctor_get(x_145, 1); +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 = 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_147 = lean_box(0); -lean_inc(x_126); -if (lean_is_scalar(x_127)) { - x_148 = lean_alloc_ctor(0, 2, 0); +lean_inc(x_116); +x_148 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_116, x_146, x_7, x_8, x_9, x_10, x_147); +x_149 = lean_ctor_get(x_148, 1); +lean_inc(x_149); +lean_dec(x_148); +x_150 = lean_box(0); +lean_inc(x_128); +if (lean_is_scalar(x_129)) { + x_151 = lean_alloc_ctor(0, 2, 0); } else { - x_148 = x_127; + x_151 = x_129; } -lean_ctor_set(x_148, 0, x_147); -lean_ctor_set(x_148, 1, x_126); -x_149 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_149, 0, x_148); -x_150 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_150, 0, x_114); -lean_ctor_set(x_150, 1, x_126); -x_151 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_151, 0, x_149); -lean_ctor_set(x_151, 1, x_150); -x_152 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_151, 0, x_150); +lean_ctor_set(x_151, 1, x_128); +x_152 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_152, 0, x_151); -x_15 = x_152; -x_16 = x_146; +x_153 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_153, 0, x_116); +lean_ctor_set(x_153, 1, x_128); +x_154 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_154, 0, x_152); +lean_ctor_set(x_154, 1, x_153); +x_155 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_155, 0, x_154); +x_15 = x_155; +x_16 = x_149; goto block_23; } else { -lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; -lean_dec(x_127); -lean_dec(x_126); -lean_dec(x_114); +lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; +lean_dec(x_129); +lean_dec(x_128); +lean_dec(x_116); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_2); lean_dec(x_1); -x_153 = lean_ctor_get(x_142, 0); -lean_inc(x_153); -x_154 = lean_ctor_get(x_142, 1); -lean_inc(x_154); -if (lean_is_exclusive(x_142)) { - lean_ctor_release(x_142, 0); - lean_ctor_release(x_142, 1); - x_155 = x_142; +x_156 = lean_ctor_get(x_145, 0); +lean_inc(x_156); +x_157 = lean_ctor_get(x_145, 1); +lean_inc(x_157); +if (lean_is_exclusive(x_145)) { + lean_ctor_release(x_145, 0); + lean_ctor_release(x_145, 1); + x_158 = x_145; } else { - lean_dec_ref(x_142); - x_155 = lean_box(0); + lean_dec_ref(x_145); + x_158 = lean_box(0); } -if (lean_is_scalar(x_155)) { - x_156 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_158)) { + x_159 = lean_alloc_ctor(1, 2, 0); } else { - x_156 = x_155; + x_159 = x_158; } -lean_ctor_set(x_156, 0, x_153); -lean_ctor_set(x_156, 1, x_154); -return x_156; +lean_ctor_set(x_159, 0, x_156); +lean_ctor_set(x_159, 1, x_157); +return x_159; } } else { -lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; -lean_dec(x_127); -lean_dec(x_126); -lean_dec(x_114); +lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; +lean_dec(x_129); +lean_dec(x_128); +lean_dec(x_116); lean_dec(x_14); lean_dec(x_10); lean_dec(x_9); @@ -56251,34 +56286,34 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_2); lean_dec(x_1); -x_157 = lean_ctor_get(x_138, 0); -lean_inc(x_157); -x_158 = lean_ctor_get(x_138, 1); -lean_inc(x_158); -if (lean_is_exclusive(x_138)) { - lean_ctor_release(x_138, 0); - lean_ctor_release(x_138, 1); - x_159 = x_138; +x_160 = lean_ctor_get(x_141, 0); +lean_inc(x_160); +x_161 = lean_ctor_get(x_141, 1); +lean_inc(x_161); +if (lean_is_exclusive(x_141)) { + lean_ctor_release(x_141, 0); + lean_ctor_release(x_141, 1); + x_162 = x_141; } else { - lean_dec_ref(x_138); - x_159 = lean_box(0); + lean_dec_ref(x_141); + x_162 = lean_box(0); } -if (lean_is_scalar(x_159)) { - x_160 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_162)) { + x_163 = lean_alloc_ctor(1, 2, 0); } else { - x_160 = x_159; + x_163 = x_162; } -lean_ctor_set(x_160, 0, x_157); -lean_ctor_set(x_160, 1, x_158); -return x_160; +lean_ctor_set(x_163, 0, x_160); +lean_ctor_set(x_163, 1, x_161); +return x_163; } } } else { -lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; -lean_dec(x_119); -lean_dec(x_114); +lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; +lean_dec(x_121); +lean_dec(x_116); lean_dec(x_14); lean_dec(x_10); lean_dec(x_9); @@ -56286,32 +56321,32 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_2); lean_dec(x_1); -x_161 = lean_ctor_get(x_122, 0); -lean_inc(x_161); -x_162 = lean_ctor_get(x_122, 1); -lean_inc(x_162); -if (lean_is_exclusive(x_122)) { - lean_ctor_release(x_122, 0); - lean_ctor_release(x_122, 1); - x_163 = x_122; +x_164 = lean_ctor_get(x_124, 0); +lean_inc(x_164); +x_165 = lean_ctor_get(x_124, 1); +lean_inc(x_165); +if (lean_is_exclusive(x_124)) { + lean_ctor_release(x_124, 0); + lean_ctor_release(x_124, 1); + x_166 = x_124; } else { - lean_dec_ref(x_122); - x_163 = lean_box(0); + lean_dec_ref(x_124); + x_166 = lean_box(0); } -if (lean_is_scalar(x_163)) { - x_164 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_166)) { + x_167 = lean_alloc_ctor(1, 2, 0); } else { - x_164 = x_163; + x_167 = x_166; } -lean_ctor_set(x_164, 0, x_161); -lean_ctor_set(x_164, 1, x_162); -return x_164; +lean_ctor_set(x_167, 0, x_164); +lean_ctor_set(x_167, 1, x_165); +return x_167; } } else { -lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; -lean_dec(x_114); +lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; +lean_dec(x_116); lean_dec(x_14); lean_dec(x_10); lean_dec(x_9); @@ -56319,26 +56354,26 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_2); lean_dec(x_1); -x_165 = lean_ctor_get(x_115, 0); -lean_inc(x_165); -x_166 = lean_ctor_get(x_115, 1); -lean_inc(x_166); -if (lean_is_exclusive(x_115)) { - lean_ctor_release(x_115, 0); - lean_ctor_release(x_115, 1); - x_167 = x_115; +x_168 = lean_ctor_get(x_117, 0); +lean_inc(x_168); +x_169 = lean_ctor_get(x_117, 1); +lean_inc(x_169); +if (lean_is_exclusive(x_117)) { + lean_ctor_release(x_117, 0); + lean_ctor_release(x_117, 1); + x_170 = x_117; } else { - lean_dec_ref(x_115); - x_167 = lean_box(0); + lean_dec_ref(x_117); + x_170 = lean_box(0); } -if (lean_is_scalar(x_167)) { - x_168 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_170)) { + x_171 = lean_alloc_ctor(1, 2, 0); } else { - x_168 = x_167; + x_171 = x_170; } -lean_ctor_set(x_168, 0, x_165); -lean_ctor_set(x_168, 1, x_166); -return x_168; +lean_ctor_set(x_171, 0, x_168); +lean_ctor_set(x_171, 1, x_169); +return x_171; } } block_23: @@ -56888,111 +56923,115 @@ lean_dec(x_18); x_21 = !lean_is_exclusive(x_19); if (x_21 == 0) { -lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; x_22 = lean_ctor_get(x_19, 0); x_23 = lean_ctor_get(x_19, 1); -x_24 = l_Lean_Meta_Simp_DefaultMethods_discharge_x3f___lambda__1___closed__6; -x_25 = l_Lean_Expr_isConstOf(x_22, x_24); -if (x_25 == 0) +x_24 = l_Lean_Expr_consumeMData(x_22); +x_25 = l_Lean_Meta_Simp_DefaultMethods_discharge_x3f___lambda__1___closed__6; +x_26 = l_Lean_Expr_isConstOf(x_24, x_25); +lean_dec(x_24); +if (x_26 == 0) { -lean_object* x_26; lean_object* x_27; +lean_object* x_27; lean_object* x_28; lean_free_object(x_19); -x_26 = lean_box(0); -x_27 = l_Lean_Meta_dsimpGoal___lambda__4(x_12, x_23, x_16, x_22, x_3, x_4, x_26, x_7, x_8, x_9, x_10, x_20); +x_27 = lean_box(0); +x_28 = l_Lean_Meta_dsimpGoal___lambda__4(x_12, x_23, x_16, x_22, x_3, x_4, x_27, x_7, x_8, x_9, x_10, x_20); lean_dec(x_16); -return x_27; +return x_28; } else { -lean_object* x_28; lean_object* x_29; uint8_t x_30; +lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_dec(x_22); lean_dec(x_16); -x_28 = l_Lean_Meta_simpTargetCore___closed__3; -x_29 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_4, x_28, x_7, x_8, x_9, x_10, x_20); +x_29 = l_Lean_Meta_simpTargetCore___closed__3; +x_30 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_4, x_29, x_7, x_8, x_9, x_10, x_20); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -x_30 = !lean_is_exclusive(x_29); -if (x_30 == 0) +x_31 = !lean_is_exclusive(x_30); +if (x_31 == 0) { -lean_object* x_31; -x_31 = lean_ctor_get(x_29, 0); -lean_dec(x_31); +lean_object* x_32; +x_32 = lean_ctor_get(x_30, 0); +lean_dec(x_32); lean_ctor_set(x_19, 0, x_3); -lean_ctor_set(x_29, 0, x_19); -return x_29; +lean_ctor_set(x_30, 0, x_19); +return x_30; } else { -lean_object* x_32; lean_object* x_33; -x_32 = lean_ctor_get(x_29, 1); -lean_inc(x_32); -lean_dec(x_29); +lean_object* x_33; lean_object* x_34; +x_33 = lean_ctor_get(x_30, 1); +lean_inc(x_33); +lean_dec(x_30); lean_ctor_set(x_19, 0, x_3); -x_33 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_33, 0, x_19); -lean_ctor_set(x_33, 1, x_32); -return x_33; +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_19); +lean_ctor_set(x_34, 1, x_33); +return x_34; } } } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; -x_34 = lean_ctor_get(x_19, 0); -x_35 = lean_ctor_get(x_19, 1); +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; +x_35 = lean_ctor_get(x_19, 0); +x_36 = lean_ctor_get(x_19, 1); +lean_inc(x_36); lean_inc(x_35); -lean_inc(x_34); lean_dec(x_19); -x_36 = l_Lean_Meta_Simp_DefaultMethods_discharge_x3f___lambda__1___closed__6; -x_37 = l_Lean_Expr_isConstOf(x_34, x_36); -if (x_37 == 0) +x_37 = l_Lean_Expr_consumeMData(x_35); +x_38 = l_Lean_Meta_Simp_DefaultMethods_discharge_x3f___lambda__1___closed__6; +x_39 = l_Lean_Expr_isConstOf(x_37, x_38); +lean_dec(x_37); +if (x_39 == 0) { -lean_object* x_38; lean_object* x_39; -x_38 = lean_box(0); -x_39 = l_Lean_Meta_dsimpGoal___lambda__4(x_12, x_35, x_16, x_34, x_3, x_4, x_38, x_7, x_8, x_9, x_10, x_20); +lean_object* x_40; lean_object* x_41; +x_40 = lean_box(0); +x_41 = l_Lean_Meta_dsimpGoal___lambda__4(x_12, x_36, x_16, x_35, x_3, x_4, x_40, x_7, x_8, x_9, x_10, x_20); lean_dec(x_16); -return x_39; +return x_41; } 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_dec(x_34); +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_35); lean_dec(x_16); -x_40 = l_Lean_Meta_simpTargetCore___closed__3; -x_41 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_4, x_40, x_7, x_8, x_9, x_10, x_20); +x_42 = l_Lean_Meta_simpTargetCore___closed__3; +x_43 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_4, x_42, x_7, x_8, x_9, x_10, x_20); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -x_42 = lean_ctor_get(x_41, 1); -lean_inc(x_42); -if (lean_is_exclusive(x_41)) { - lean_ctor_release(x_41, 0); - lean_ctor_release(x_41, 1); - x_43 = x_41; +x_44 = lean_ctor_get(x_43, 1); +lean_inc(x_44); +if (lean_is_exclusive(x_43)) { + lean_ctor_release(x_43, 0); + lean_ctor_release(x_43, 1); + x_45 = x_43; } else { - lean_dec_ref(x_41); - x_43 = lean_box(0); + lean_dec_ref(x_43); + x_45 = lean_box(0); } -x_44 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_44, 0, x_3); -lean_ctor_set(x_44, 1, x_35); -if (lean_is_scalar(x_43)) { - x_45 = lean_alloc_ctor(0, 2, 0); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_3); +lean_ctor_set(x_46, 1, x_36); +if (lean_is_scalar(x_45)) { + x_47 = lean_alloc_ctor(0, 2, 0); } else { - x_45 = x_43; + x_47 = x_45; } -lean_ctor_set(x_45, 0, x_44); -lean_ctor_set(x_45, 1, x_42); -return x_45; +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_44); +return x_47; } } } else { -uint8_t x_46; +uint8_t x_48; lean_dec(x_16); lean_dec(x_10); lean_dec(x_9); @@ -57000,29 +57039,29 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_4); lean_dec(x_3); -x_46 = !lean_is_exclusive(x_18); -if (x_46 == 0) +x_48 = !lean_is_exclusive(x_18); +if (x_48 == 0) { return x_18; } else { -lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_47 = lean_ctor_get(x_18, 0); -x_48 = lean_ctor_get(x_18, 1); -lean_inc(x_48); -lean_inc(x_47); +lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_49 = lean_ctor_get(x_18, 0); +x_50 = lean_ctor_get(x_18, 1); +lean_inc(x_50); +lean_inc(x_49); lean_dec(x_18); -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; +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 { -uint8_t x_50; +uint8_t x_52; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -57031,23 +57070,23 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_50 = !lean_is_exclusive(x_15); -if (x_50 == 0) +x_52 = !lean_is_exclusive(x_15); +if (x_52 == 0) { return x_15; } else { -lean_object* x_51; lean_object* x_52; lean_object* x_53; -x_51 = lean_ctor_get(x_15, 0); -x_52 = lean_ctor_get(x_15, 1); -lean_inc(x_52); -lean_inc(x_51); +lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_53 = lean_ctor_get(x_15, 0); +x_54 = lean_ctor_get(x_15, 1); +lean_inc(x_54); +lean_inc(x_53); lean_dec(x_15); -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; +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; } } } diff --git a/stage0/stdlib/Lean/Meta/Tactic/Simp/Rewrite.c b/stage0/stdlib/Lean/Meta/Tactic/Simp/Rewrite.c index 5787ab7c8032..71a59cbd2652 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Simp/Rewrite.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Simp/Rewrite.c @@ -311,6 +311,7 @@ static lean_object* l_Lean_Meta_Simp_synthesizeArgs_synthesizeInstance___closed_ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_rewritePost___spec__1(lean_object*, 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_object*); static lean_object* l_Lean_Meta_ppSimpTheorem___at___private_Lean_Meta_Tactic_Simp_Rewrite_0__Lean_Meta_Simp_tryTheoremCore_go___spec__1___closed__3; static lean_object* l_Lean_Meta_Simp_rewrite_x3f___closed__2; +lean_object* l_Lean_Expr_consumeMData(lean_object*); lean_object* l_Lean_MessageData_ofName(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_simpMatchCore_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_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_Rewrite_0__Lean_Meta_Simp_tryTheoremCore(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -12128,1344 +12129,1347 @@ uint8_t x_8; x_8 = l_Lean_Expr_hasMVar(x_1); if (x_8 == 0) { -lean_object* x_9; uint8_t x_10; -x_9 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__2; -x_10 = l_Lean_Expr_isConstOf(x_1, x_9); -if (x_10 == 0) -{ -lean_object* x_11; uint8_t x_12; -x_11 = l_Lean_Meta_Simp_rewriteCtorEq_x3f___lambda__1___closed__2; -x_12 = l_Lean_Expr_isConstOf(x_1, x_11); -if (x_12 == 0) +lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_9 = l_Lean_Expr_consumeMData(x_1); +x_10 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__2; +x_11 = l_Lean_Expr_isConstOf(x_9, x_10); +if (x_11 == 0) { -uint8_t x_13; -x_13 = !lean_is_exclusive(x_2); +lean_object* x_12; uint8_t x_13; +x_12 = l_Lean_Meta_Simp_rewriteCtorEq_x3f___lambda__1___closed__2; +x_13 = l_Lean_Expr_isConstOf(x_9, x_12); +lean_dec(x_9); if (x_13 == 0) { -lean_object* x_14; uint8_t x_15; -x_14 = lean_ctor_get(x_2, 0); -x_15 = !lean_is_exclusive(x_14); -if (x_15 == 0) +uint8_t x_14; +x_14 = !lean_is_exclusive(x_2); +if (x_14 == 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; 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; uint8_t x_31; uint8_t x_32; uint8_t x_33; lean_object* x_34; -x_16 = lean_ctor_get(x_2, 1); -x_17 = lean_ctor_get(x_2, 2); -x_18 = lean_ctor_get(x_2, 3); -x_19 = lean_ctor_get(x_2, 4); -x_20 = lean_ctor_get(x_2, 5); -x_21 = lean_ctor_get_uint8(x_14, 0); -x_22 = lean_ctor_get_uint8(x_14, 1); -x_23 = lean_ctor_get_uint8(x_14, 2); -x_24 = lean_ctor_get_uint8(x_14, 3); -x_25 = lean_ctor_get_uint8(x_14, 4); -x_26 = lean_ctor_get_uint8(x_14, 6); -x_27 = lean_ctor_get_uint8(x_14, 7); -x_28 = lean_ctor_get_uint8(x_14, 8); -x_29 = lean_ctor_get_uint8(x_14, 9); -x_30 = lean_ctor_get_uint8(x_14, 10); -x_31 = lean_ctor_get_uint8(x_14, 11); -x_32 = lean_ctor_get_uint8(x_14, 12); -x_33 = 3; -lean_ctor_set_uint8(x_14, 5, x_33); +lean_object* x_15; uint8_t x_16; +x_15 = lean_ctor_get(x_2, 0); +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; 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; uint8_t x_31; uint8_t x_32; uint8_t x_33; uint8_t x_34; lean_object* x_35; +x_17 = lean_ctor_get(x_2, 1); +x_18 = lean_ctor_get(x_2, 2); +x_19 = lean_ctor_get(x_2, 3); +x_20 = lean_ctor_get(x_2, 4); +x_21 = lean_ctor_get(x_2, 5); +x_22 = lean_ctor_get_uint8(x_15, 0); +x_23 = lean_ctor_get_uint8(x_15, 1); +x_24 = lean_ctor_get_uint8(x_15, 2); +x_25 = lean_ctor_get_uint8(x_15, 3); +x_26 = lean_ctor_get_uint8(x_15, 4); +x_27 = lean_ctor_get_uint8(x_15, 6); +x_28 = lean_ctor_get_uint8(x_15, 7); +x_29 = lean_ctor_get_uint8(x_15, 8); +x_30 = lean_ctor_get_uint8(x_15, 9); +x_31 = lean_ctor_get_uint8(x_15, 10); +x_32 = lean_ctor_get_uint8(x_15, 11); +x_33 = lean_ctor_get_uint8(x_15, 12); +x_34 = 3; +lean_ctor_set_uint8(x_15, 5, x_34); +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_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); lean_inc(x_1); -x_34 = l_Lean_Meta_mkDecide(x_1, x_2, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_34) == 0) +x_35 = l_Lean_Meta_mkDecide(x_1, x_2, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_35) == 0) { -lean_object* x_35; lean_object* x_36; uint8_t x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_35 = lean_ctor_get(x_34, 0); -lean_inc(x_35); -x_36 = lean_ctor_get(x_34, 1); +lean_object* x_36; lean_object* x_37; uint8_t x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_36 = lean_ctor_get(x_35, 0); lean_inc(x_36); -lean_dec(x_34); -x_37 = 1; -x_38 = lean_alloc_ctor(0, 0, 13); -lean_ctor_set_uint8(x_38, 0, x_21); -lean_ctor_set_uint8(x_38, 1, x_22); -lean_ctor_set_uint8(x_38, 2, x_23); -lean_ctor_set_uint8(x_38, 3, x_24); -lean_ctor_set_uint8(x_38, 4, x_25); -lean_ctor_set_uint8(x_38, 5, x_37); -lean_ctor_set_uint8(x_38, 6, x_26); -lean_ctor_set_uint8(x_38, 7, x_27); -lean_ctor_set_uint8(x_38, 8, x_28); -lean_ctor_set_uint8(x_38, 9, x_29); -lean_ctor_set_uint8(x_38, 10, x_30); -lean_ctor_set_uint8(x_38, 11, x_31); -lean_ctor_set_uint8(x_38, 12, x_32); -x_39 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_39, 1, x_16); -lean_ctor_set(x_39, 2, x_17); -lean_ctor_set(x_39, 3, x_18); -lean_ctor_set(x_39, 4, x_19); -lean_ctor_set(x_39, 5, x_20); +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +lean_dec(x_35); +x_38 = 1; +x_39 = lean_alloc_ctor(0, 0, 13); +lean_ctor_set_uint8(x_39, 0, x_22); +lean_ctor_set_uint8(x_39, 1, x_23); +lean_ctor_set_uint8(x_39, 2, x_24); +lean_ctor_set_uint8(x_39, 3, x_25); +lean_ctor_set_uint8(x_39, 4, x_26); +lean_ctor_set_uint8(x_39, 5, x_38); +lean_ctor_set_uint8(x_39, 6, x_27); +lean_ctor_set_uint8(x_39, 7, x_28); +lean_ctor_set_uint8(x_39, 8, x_29); +lean_ctor_set_uint8(x_39, 9, x_30); +lean_ctor_set_uint8(x_39, 10, x_31); +lean_ctor_set_uint8(x_39, 11, x_32); +lean_ctor_set_uint8(x_39, 12, x_33); +x_40 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_17); +lean_ctor_set(x_40, 2, x_18); +lean_ctor_set(x_40, 3, x_19); +lean_ctor_set(x_40, 4, x_20); +lean_ctor_set(x_40, 5, x_21); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -lean_inc(x_35); -x_40 = lean_whnf(x_35, x_39, x_3, x_4, x_5, x_36); -if (lean_obj_tag(x_40) == 0) +lean_inc(x_36); +x_41 = lean_whnf(x_36, x_40, x_3, x_4, x_5, x_37); +if (lean_obj_tag(x_41) == 0) { -uint8_t x_41; -x_41 = !lean_is_exclusive(x_40); -if (x_41 == 0) +uint8_t x_42; +x_42 = !lean_is_exclusive(x_41); +if (x_42 == 0) { -lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; -x_42 = lean_ctor_get(x_40, 0); -x_43 = lean_ctor_get(x_40, 1); -x_44 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__5; -x_45 = l_Lean_Expr_isConstOf(x_42, x_44); -if (x_45 == 0) +lean_object* x_43; lean_object* x_44; lean_object* x_45; uint8_t x_46; +x_43 = lean_ctor_get(x_41, 0); +x_44 = lean_ctor_get(x_41, 1); +x_45 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__5; +x_46 = l_Lean_Expr_isConstOf(x_43, x_45); +if (x_46 == 0) { -lean_object* x_46; uint8_t x_47; -x_46 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__7; -x_47 = l_Lean_Expr_isConstOf(x_42, x_46); -lean_dec(x_42); -if (x_47 == 0) +lean_object* x_47; uint8_t x_48; +x_47 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__7; +x_48 = l_Lean_Expr_isConstOf(x_43, x_47); +lean_dec(x_43); +if (x_48 == 0) { -lean_object* x_48; -lean_dec(x_35); +lean_object* x_49; +lean_dec(x_36); lean_dec(x_2); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_48 = lean_box(0); -lean_ctor_set(x_40, 0, x_48); -return x_40; +x_49 = lean_box(0); +lean_ctor_set(x_41, 0, x_49); +return x_41; } else { -lean_object* x_49; lean_object* x_50; -lean_free_object(x_40); -x_49 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__8; -x_50 = l_Lean_Meta_mkEqRefl(x_49, x_2, x_3, x_4, x_5, x_43); -if (lean_obj_tag(x_50) == 0) +lean_object* x_50; lean_object* x_51; +lean_free_object(x_41); +x_50 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__8; +x_51 = l_Lean_Meta_mkEqRefl(x_50, x_2, x_3, x_4, x_5, x_44); +if (lean_obj_tag(x_51) == 0) { -uint8_t x_51; -x_51 = !lean_is_exclusive(x_50); -if (x_51 == 0) +uint8_t x_52; +x_52 = !lean_is_exclusive(x_51); +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; 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_52 = lean_ctor_get(x_50, 0); -x_53 = l_Lean_Expr_appArg_x21(x_35); -lean_dec(x_35); -x_54 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__12; -x_55 = lean_array_push(x_54, x_1); -x_56 = lean_array_push(x_55, x_53); -x_57 = lean_array_push(x_56, x_52); -x_58 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__11; -x_59 = l_Lean_mkAppN(x_58, x_57); -x_60 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_60, 0, x_59); -x_61 = l_Lean_Meta_Simp_rewriteCtorEq_x3f___lambda__1___closed__3; -x_62 = lean_unsigned_to_nat(0u); -x_63 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_63, 0, x_61); -lean_ctor_set(x_63, 1, x_60); -lean_ctor_set(x_63, 2, x_62); -x_64 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_64, 0, x_63); -lean_ctor_set(x_50, 0, x_64); -return x_50; +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_53 = lean_ctor_get(x_51, 0); +x_54 = l_Lean_Expr_appArg_x21(x_36); +lean_dec(x_36); +x_55 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__12; +x_56 = lean_array_push(x_55, x_1); +x_57 = lean_array_push(x_56, x_54); +x_58 = lean_array_push(x_57, x_53); +x_59 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__11; +x_60 = l_Lean_mkAppN(x_59, x_58); +x_61 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_61, 0, x_60); +x_62 = l_Lean_Meta_Simp_rewriteCtorEq_x3f___lambda__1___closed__3; +x_63 = lean_unsigned_to_nat(0u); +x_64 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_64, 0, x_62); +lean_ctor_set(x_64, 1, x_61); +lean_ctor_set(x_64, 2, x_63); +x_65 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_65, 0, x_64); +lean_ctor_set(x_51, 0, x_65); +return x_51; } 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; 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_65 = lean_ctor_get(x_50, 0); -x_66 = lean_ctor_get(x_50, 1); +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_66 = lean_ctor_get(x_51, 0); +x_67 = lean_ctor_get(x_51, 1); +lean_inc(x_67); lean_inc(x_66); -lean_inc(x_65); -lean_dec(x_50); -x_67 = l_Lean_Expr_appArg_x21(x_35); -lean_dec(x_35); -x_68 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__12; -x_69 = lean_array_push(x_68, x_1); -x_70 = lean_array_push(x_69, x_67); -x_71 = lean_array_push(x_70, x_65); -x_72 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__11; -x_73 = l_Lean_mkAppN(x_72, x_71); -x_74 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_74, 0, x_73); -x_75 = l_Lean_Meta_Simp_rewriteCtorEq_x3f___lambda__1___closed__3; -x_76 = lean_unsigned_to_nat(0u); -x_77 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_77, 0, x_75); -lean_ctor_set(x_77, 1, x_74); -lean_ctor_set(x_77, 2, x_76); -x_78 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_78, 0, x_77); -x_79 = lean_alloc_ctor(0, 2, 0); +lean_dec(x_51); +x_68 = l_Lean_Expr_appArg_x21(x_36); +lean_dec(x_36); +x_69 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__12; +x_70 = lean_array_push(x_69, x_1); +x_71 = lean_array_push(x_70, x_68); +x_72 = lean_array_push(x_71, x_66); +x_73 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__11; +x_74 = l_Lean_mkAppN(x_73, x_72); +x_75 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_75, 0, x_74); +x_76 = l_Lean_Meta_Simp_rewriteCtorEq_x3f___lambda__1___closed__3; +x_77 = lean_unsigned_to_nat(0u); +x_78 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_78, 0, x_76); +lean_ctor_set(x_78, 1, x_75); +lean_ctor_set(x_78, 2, x_77); +x_79 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_79, 0, x_78); -lean_ctor_set(x_79, 1, x_66); -return x_79; +x_80 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_80, 0, x_79); +lean_ctor_set(x_80, 1, x_67); +return x_80; } } else { -uint8_t x_80; -lean_dec(x_35); +uint8_t x_81; +lean_dec(x_36); lean_dec(x_1); -x_80 = !lean_is_exclusive(x_50); -if (x_80 == 0) +x_81 = !lean_is_exclusive(x_51); +if (x_81 == 0) { -lean_object* x_81; lean_object* x_82; -x_81 = lean_ctor_get(x_50, 0); -lean_dec(x_81); -x_82 = lean_box(0); -lean_ctor_set_tag(x_50, 0); -lean_ctor_set(x_50, 0, x_82); -return x_50; +lean_object* x_82; lean_object* x_83; +x_82 = lean_ctor_get(x_51, 0); +lean_dec(x_82); +x_83 = lean_box(0); +lean_ctor_set_tag(x_51, 0); +lean_ctor_set(x_51, 0, x_83); +return x_51; } else { -lean_object* x_83; lean_object* x_84; lean_object* x_85; -x_83 = lean_ctor_get(x_50, 1); -lean_inc(x_83); -lean_dec(x_50); -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_83); -return x_85; +lean_object* x_84; lean_object* x_85; lean_object* x_86; +x_84 = lean_ctor_get(x_51, 1); +lean_inc(x_84); +lean_dec(x_51); +x_85 = lean_box(0); +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_free_object(x_40); -lean_dec(x_42); -x_86 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__13; -x_87 = l_Lean_Meta_mkEqRefl(x_86, x_2, x_3, x_4, x_5, x_43); -if (lean_obj_tag(x_87) == 0) +lean_object* x_87; lean_object* x_88; +lean_free_object(x_41); +lean_dec(x_43); +x_87 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__13; +x_88 = l_Lean_Meta_mkEqRefl(x_87, x_2, x_3, x_4, x_5, x_44); +if (lean_obj_tag(x_88) == 0) { -uint8_t x_88; -x_88 = !lean_is_exclusive(x_87); -if (x_88 == 0) +uint8_t x_89; +x_89 = !lean_is_exclusive(x_88); +if (x_89 == 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; 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_89 = lean_ctor_get(x_87, 0); -x_90 = l_Lean_Expr_appArg_x21(x_35); -lean_dec(x_35); -x_91 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__12; -x_92 = lean_array_push(x_91, x_1); -x_93 = lean_array_push(x_92, x_90); -x_94 = lean_array_push(x_93, x_89); -x_95 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__17; -x_96 = l_Lean_mkAppN(x_95, x_94); -x_97 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_97, 0, x_96); -x_98 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__14; -x_99 = lean_unsigned_to_nat(0u); -x_100 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_100, 0, x_98); -lean_ctor_set(x_100, 1, x_97); -lean_ctor_set(x_100, 2, x_99); -x_101 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_101, 0, x_100); -lean_ctor_set(x_87, 0, x_101); -return x_87; +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_90 = lean_ctor_get(x_88, 0); +x_91 = l_Lean_Expr_appArg_x21(x_36); +lean_dec(x_36); +x_92 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__12; +x_93 = lean_array_push(x_92, x_1); +x_94 = lean_array_push(x_93, x_91); +x_95 = lean_array_push(x_94, x_90); +x_96 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__17; +x_97 = l_Lean_mkAppN(x_96, x_95); +x_98 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_98, 0, x_97); +x_99 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__14; +x_100 = lean_unsigned_to_nat(0u); +x_101 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_101, 0, x_99); +lean_ctor_set(x_101, 1, x_98); +lean_ctor_set(x_101, 2, x_100); +x_102 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_102, 0, x_101); +lean_ctor_set(x_88, 0, x_102); +return x_88; } else { -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_102 = lean_ctor_get(x_87, 0); -x_103 = lean_ctor_get(x_87, 1); +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_103 = lean_ctor_get(x_88, 0); +x_104 = lean_ctor_get(x_88, 1); +lean_inc(x_104); lean_inc(x_103); -lean_inc(x_102); -lean_dec(x_87); -x_104 = l_Lean_Expr_appArg_x21(x_35); -lean_dec(x_35); -x_105 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__12; -x_106 = lean_array_push(x_105, x_1); -x_107 = lean_array_push(x_106, x_104); -x_108 = lean_array_push(x_107, x_102); -x_109 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__17; -x_110 = l_Lean_mkAppN(x_109, x_108); -x_111 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_111, 0, x_110); -x_112 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__14; -x_113 = lean_unsigned_to_nat(0u); -x_114 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_114, 0, x_112); -lean_ctor_set(x_114, 1, x_111); -lean_ctor_set(x_114, 2, x_113); -x_115 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_115, 0, x_114); -x_116 = lean_alloc_ctor(0, 2, 0); +lean_dec(x_88); +x_105 = l_Lean_Expr_appArg_x21(x_36); +lean_dec(x_36); +x_106 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__12; +x_107 = lean_array_push(x_106, x_1); +x_108 = lean_array_push(x_107, x_105); +x_109 = lean_array_push(x_108, x_103); +x_110 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__17; +x_111 = l_Lean_mkAppN(x_110, x_109); +x_112 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_112, 0, x_111); +x_113 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__14; +x_114 = lean_unsigned_to_nat(0u); +x_115 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_115, 0, x_113); +lean_ctor_set(x_115, 1, x_112); +lean_ctor_set(x_115, 2, x_114); +x_116 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_116, 0, x_115); -lean_ctor_set(x_116, 1, x_103); -return x_116; +x_117 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_117, 0, x_116); +lean_ctor_set(x_117, 1, x_104); +return x_117; } } else { -uint8_t x_117; -lean_dec(x_35); +uint8_t x_118; +lean_dec(x_36); lean_dec(x_1); -x_117 = !lean_is_exclusive(x_87); -if (x_117 == 0) +x_118 = !lean_is_exclusive(x_88); +if (x_118 == 0) { -lean_object* x_118; lean_object* x_119; -x_118 = lean_ctor_get(x_87, 0); -lean_dec(x_118); -x_119 = lean_box(0); -lean_ctor_set_tag(x_87, 0); -lean_ctor_set(x_87, 0, x_119); -return x_87; +lean_object* x_119; lean_object* x_120; +x_119 = lean_ctor_get(x_88, 0); +lean_dec(x_119); +x_120 = lean_box(0); +lean_ctor_set_tag(x_88, 0); +lean_ctor_set(x_88, 0, x_120); +return x_88; } else { -lean_object* x_120; lean_object* x_121; lean_object* x_122; -x_120 = lean_ctor_get(x_87, 1); -lean_inc(x_120); -lean_dec(x_87); -x_121 = lean_box(0); -x_122 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_122, 0, x_121); -lean_ctor_set(x_122, 1, x_120); -return x_122; +lean_object* x_121; lean_object* x_122; lean_object* x_123; +x_121 = lean_ctor_get(x_88, 1); +lean_inc(x_121); +lean_dec(x_88); +x_122 = lean_box(0); +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 { -lean_object* x_123; lean_object* x_124; lean_object* x_125; uint8_t x_126; -x_123 = lean_ctor_get(x_40, 0); -x_124 = lean_ctor_get(x_40, 1); +lean_object* x_124; lean_object* x_125; lean_object* x_126; uint8_t x_127; +x_124 = lean_ctor_get(x_41, 0); +x_125 = lean_ctor_get(x_41, 1); +lean_inc(x_125); lean_inc(x_124); -lean_inc(x_123); -lean_dec(x_40); -x_125 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__5; -x_126 = l_Lean_Expr_isConstOf(x_123, x_125); -if (x_126 == 0) +lean_dec(x_41); +x_126 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__5; +x_127 = l_Lean_Expr_isConstOf(x_124, x_126); +if (x_127 == 0) { -lean_object* x_127; uint8_t x_128; -x_127 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__7; -x_128 = l_Lean_Expr_isConstOf(x_123, x_127); -lean_dec(x_123); -if (x_128 == 0) +lean_object* x_128; uint8_t x_129; +x_128 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__7; +x_129 = l_Lean_Expr_isConstOf(x_124, x_128); +lean_dec(x_124); +if (x_129 == 0) { -lean_object* x_129; lean_object* x_130; -lean_dec(x_35); +lean_object* x_130; lean_object* x_131; +lean_dec(x_36); lean_dec(x_2); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_129 = lean_box(0); -x_130 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_130, 0, x_129); -lean_ctor_set(x_130, 1, x_124); -return x_130; +x_130 = lean_box(0); +x_131 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_131, 0, x_130); +lean_ctor_set(x_131, 1, x_125); +return x_131; } else { -lean_object* x_131; lean_object* x_132; -x_131 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__8; -x_132 = l_Lean_Meta_mkEqRefl(x_131, x_2, x_3, x_4, x_5, x_124); -if (lean_obj_tag(x_132) == 0) +lean_object* x_132; lean_object* x_133; +x_132 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__8; +x_133 = l_Lean_Meta_mkEqRefl(x_132, x_2, x_3, x_4, x_5, x_125); +if (lean_obj_tag(x_133) == 0) { -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_133 = lean_ctor_get(x_132, 0); -lean_inc(x_133); -x_134 = lean_ctor_get(x_132, 1); +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_134 = lean_ctor_get(x_133, 0); 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; +x_135 = lean_ctor_get(x_133, 1); +lean_inc(x_135); +if (lean_is_exclusive(x_133)) { + lean_ctor_release(x_133, 0); + lean_ctor_release(x_133, 1); + x_136 = x_133; } else { - lean_dec_ref(x_132); - x_135 = lean_box(0); + lean_dec_ref(x_133); + x_136 = lean_box(0); } -x_136 = l_Lean_Expr_appArg_x21(x_35); -lean_dec(x_35); -x_137 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__12; -x_138 = lean_array_push(x_137, x_1); -x_139 = lean_array_push(x_138, x_136); -x_140 = lean_array_push(x_139, x_133); -x_141 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__11; -x_142 = l_Lean_mkAppN(x_141, x_140); -x_143 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_143, 0, x_142); -x_144 = l_Lean_Meta_Simp_rewriteCtorEq_x3f___lambda__1___closed__3; -x_145 = lean_unsigned_to_nat(0u); -x_146 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_146, 0, x_144); -lean_ctor_set(x_146, 1, x_143); -lean_ctor_set(x_146, 2, x_145); -x_147 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_147, 0, x_146); -if (lean_is_scalar(x_135)) { - x_148 = lean_alloc_ctor(0, 2, 0); +x_137 = l_Lean_Expr_appArg_x21(x_36); +lean_dec(x_36); +x_138 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__12; +x_139 = lean_array_push(x_138, x_1); +x_140 = lean_array_push(x_139, x_137); +x_141 = lean_array_push(x_140, x_134); +x_142 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__11; +x_143 = l_Lean_mkAppN(x_142, x_141); +x_144 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_144, 0, x_143); +x_145 = l_Lean_Meta_Simp_rewriteCtorEq_x3f___lambda__1___closed__3; +x_146 = lean_unsigned_to_nat(0u); +x_147 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_147, 0, x_145); +lean_ctor_set(x_147, 1, x_144); +lean_ctor_set(x_147, 2, x_146); +x_148 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_148, 0, x_147); +if (lean_is_scalar(x_136)) { + x_149 = lean_alloc_ctor(0, 2, 0); } else { - x_148 = x_135; + x_149 = x_136; } -lean_ctor_set(x_148, 0, x_147); -lean_ctor_set(x_148, 1, x_134); -return x_148; +lean_ctor_set(x_149, 0, x_148); +lean_ctor_set(x_149, 1, x_135); +return x_149; } else { -lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; -lean_dec(x_35); +lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; +lean_dec(x_36); lean_dec(x_1); -x_149 = lean_ctor_get(x_132, 1); -lean_inc(x_149); -if (lean_is_exclusive(x_132)) { - lean_ctor_release(x_132, 0); - lean_ctor_release(x_132, 1); - x_150 = x_132; +x_150 = lean_ctor_get(x_133, 1); +lean_inc(x_150); +if (lean_is_exclusive(x_133)) { + lean_ctor_release(x_133, 0); + lean_ctor_release(x_133, 1); + x_151 = x_133; } else { - lean_dec_ref(x_132); - x_150 = lean_box(0); + lean_dec_ref(x_133); + x_151 = lean_box(0); } -x_151 = lean_box(0); -if (lean_is_scalar(x_150)) { - x_152 = 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_152 = x_150; - lean_ctor_set_tag(x_152, 0); + x_153 = x_151; + lean_ctor_set_tag(x_153, 0); } -lean_ctor_set(x_152, 0, x_151); -lean_ctor_set(x_152, 1, x_149); -return x_152; +lean_ctor_set(x_153, 0, x_152); +lean_ctor_set(x_153, 1, x_150); +return x_153; } } } else { -lean_object* x_153; lean_object* x_154; -lean_dec(x_123); -x_153 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__13; -x_154 = l_Lean_Meta_mkEqRefl(x_153, x_2, x_3, x_4, x_5, x_124); -if (lean_obj_tag(x_154) == 0) +lean_object* x_154; lean_object* x_155; +lean_dec(x_124); +x_154 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__13; +x_155 = l_Lean_Meta_mkEqRefl(x_154, x_2, x_3, x_4, x_5, x_125); +if (lean_obj_tag(x_155) == 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; 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_155 = lean_ctor_get(x_154, 0); -lean_inc(x_155); -x_156 = lean_ctor_get(x_154, 1); +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_156 = lean_ctor_get(x_155, 0); lean_inc(x_156); -if (lean_is_exclusive(x_154)) { - lean_ctor_release(x_154, 0); - lean_ctor_release(x_154, 1); - x_157 = x_154; +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_154); - x_157 = lean_box(0); + lean_dec_ref(x_155); + x_158 = lean_box(0); } -x_158 = l_Lean_Expr_appArg_x21(x_35); -lean_dec(x_35); -x_159 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__12; -x_160 = lean_array_push(x_159, x_1); -x_161 = lean_array_push(x_160, x_158); -x_162 = lean_array_push(x_161, x_155); -x_163 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__17; -x_164 = l_Lean_mkAppN(x_163, x_162); -x_165 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_165, 0, x_164); -x_166 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__14; -x_167 = lean_unsigned_to_nat(0u); -x_168 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_168, 0, x_166); -lean_ctor_set(x_168, 1, x_165); -lean_ctor_set(x_168, 2, x_167); -x_169 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_169, 0, x_168); -if (lean_is_scalar(x_157)) { - x_170 = lean_alloc_ctor(0, 2, 0); +x_159 = l_Lean_Expr_appArg_x21(x_36); +lean_dec(x_36); +x_160 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__12; +x_161 = lean_array_push(x_160, x_1); +x_162 = lean_array_push(x_161, x_159); +x_163 = lean_array_push(x_162, x_156); +x_164 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__17; +x_165 = l_Lean_mkAppN(x_164, x_163); +x_166 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_166, 0, x_165); +x_167 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__14; +x_168 = lean_unsigned_to_nat(0u); +x_169 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_169, 0, x_167); +lean_ctor_set(x_169, 1, x_166); +lean_ctor_set(x_169, 2, x_168); +x_170 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_170, 0, x_169); +if (lean_is_scalar(x_158)) { + x_171 = lean_alloc_ctor(0, 2, 0); } else { - x_170 = x_157; + x_171 = x_158; } -lean_ctor_set(x_170, 0, x_169); -lean_ctor_set(x_170, 1, x_156); -return x_170; +lean_ctor_set(x_171, 0, x_170); +lean_ctor_set(x_171, 1, x_157); +return x_171; } else { -lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; -lean_dec(x_35); +lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; +lean_dec(x_36); lean_dec(x_1); -x_171 = lean_ctor_get(x_154, 1); -lean_inc(x_171); -if (lean_is_exclusive(x_154)) { - lean_ctor_release(x_154, 0); - lean_ctor_release(x_154, 1); - x_172 = x_154; +x_172 = lean_ctor_get(x_155, 1); +lean_inc(x_172); +if (lean_is_exclusive(x_155)) { + lean_ctor_release(x_155, 0); + lean_ctor_release(x_155, 1); + x_173 = x_155; } else { - lean_dec_ref(x_154); - x_172 = lean_box(0); + lean_dec_ref(x_155); + x_173 = lean_box(0); } -x_173 = lean_box(0); -if (lean_is_scalar(x_172)) { - x_174 = lean_alloc_ctor(0, 2, 0); +x_174 = lean_box(0); +if (lean_is_scalar(x_173)) { + x_175 = lean_alloc_ctor(0, 2, 0); } else { - x_174 = x_172; - lean_ctor_set_tag(x_174, 0); + x_175 = x_173; + lean_ctor_set_tag(x_175, 0); } -lean_ctor_set(x_174, 0, x_173); -lean_ctor_set(x_174, 1, x_171); -return x_174; +lean_ctor_set(x_175, 0, x_174); +lean_ctor_set(x_175, 1, x_172); +return x_175; } } } } else { -uint8_t x_175; -lean_dec(x_35); +uint8_t x_176; +lean_dec(x_36); lean_dec(x_2); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_175 = !lean_is_exclusive(x_40); -if (x_175 == 0) -{ -lean_object* x_176; lean_object* x_177; -x_176 = lean_ctor_get(x_40, 0); -lean_dec(x_176); -x_177 = lean_box(0); -lean_ctor_set_tag(x_40, 0); -lean_ctor_set(x_40, 0, x_177); -return x_40; +x_176 = !lean_is_exclusive(x_41); +if (x_176 == 0) +{ +lean_object* x_177; lean_object* x_178; +x_177 = lean_ctor_get(x_41, 0); +lean_dec(x_177); +x_178 = lean_box(0); +lean_ctor_set_tag(x_41, 0); +lean_ctor_set(x_41, 0, x_178); +return x_41; } else { -lean_object* x_178; lean_object* x_179; lean_object* x_180; -x_178 = lean_ctor_get(x_40, 1); -lean_inc(x_178); -lean_dec(x_40); -x_179 = lean_box(0); -x_180 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_180, 0, x_179); -lean_ctor_set(x_180, 1, x_178); -return x_180; +lean_object* x_179; lean_object* x_180; lean_object* x_181; +x_179 = lean_ctor_get(x_41, 1); +lean_inc(x_179); +lean_dec(x_41); +x_180 = lean_box(0); +x_181 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_181, 0, x_180); +lean_ctor_set(x_181, 1, x_179); +return x_181; } } } else { -uint8_t x_181; +uint8_t x_182; lean_dec(x_2); +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_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_181 = !lean_is_exclusive(x_34); -if (x_181 == 0) +x_182 = !lean_is_exclusive(x_35); +if (x_182 == 0) { -lean_object* x_182; lean_object* x_183; -x_182 = lean_ctor_get(x_34, 0); -lean_dec(x_182); -x_183 = lean_box(0); -lean_ctor_set_tag(x_34, 0); -lean_ctor_set(x_34, 0, x_183); -return x_34; -} -else -{ -lean_object* x_184; lean_object* x_185; lean_object* x_186; -x_184 = lean_ctor_get(x_34, 1); -lean_inc(x_184); -lean_dec(x_34); -x_185 = lean_box(0); -x_186 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_186, 0, x_185); -lean_ctor_set(x_186, 1, x_184); -return x_186; -} -} +lean_object* x_183; lean_object* x_184; +x_183 = lean_ctor_get(x_35, 0); +lean_dec(x_183); +x_184 = lean_box(0); +lean_ctor_set_tag(x_35, 0); +lean_ctor_set(x_35, 0, x_184); +return x_35; } 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; uint8_t x_193; uint8_t x_194; uint8_t x_195; uint8_t x_196; uint8_t x_197; uint8_t x_198; uint8_t x_199; uint8_t x_200; uint8_t x_201; uint8_t x_202; uint8_t x_203; uint8_t x_204; lean_object* x_205; lean_object* x_206; -x_187 = lean_ctor_get(x_2, 1); -x_188 = lean_ctor_get(x_2, 2); -x_189 = lean_ctor_get(x_2, 3); -x_190 = lean_ctor_get(x_2, 4); -x_191 = lean_ctor_get(x_2, 5); -x_192 = lean_ctor_get_uint8(x_14, 0); -x_193 = lean_ctor_get_uint8(x_14, 1); -x_194 = lean_ctor_get_uint8(x_14, 2); -x_195 = lean_ctor_get_uint8(x_14, 3); -x_196 = lean_ctor_get_uint8(x_14, 4); -x_197 = lean_ctor_get_uint8(x_14, 6); -x_198 = lean_ctor_get_uint8(x_14, 7); -x_199 = lean_ctor_get_uint8(x_14, 8); -x_200 = lean_ctor_get_uint8(x_14, 9); -x_201 = lean_ctor_get_uint8(x_14, 10); -x_202 = lean_ctor_get_uint8(x_14, 11); -x_203 = lean_ctor_get_uint8(x_14, 12); -lean_dec(x_14); -x_204 = 3; -x_205 = lean_alloc_ctor(0, 0, 13); -lean_ctor_set_uint8(x_205, 0, x_192); -lean_ctor_set_uint8(x_205, 1, x_193); -lean_ctor_set_uint8(x_205, 2, x_194); -lean_ctor_set_uint8(x_205, 3, x_195); -lean_ctor_set_uint8(x_205, 4, x_196); -lean_ctor_set_uint8(x_205, 5, x_204); -lean_ctor_set_uint8(x_205, 6, x_197); -lean_ctor_set_uint8(x_205, 7, x_198); -lean_ctor_set_uint8(x_205, 8, x_199); -lean_ctor_set_uint8(x_205, 9, x_200); -lean_ctor_set_uint8(x_205, 10, x_201); -lean_ctor_set_uint8(x_205, 11, x_202); -lean_ctor_set_uint8(x_205, 12, x_203); +lean_object* x_185; lean_object* x_186; lean_object* x_187; +x_185 = lean_ctor_get(x_35, 1); +lean_inc(x_185); +lean_dec(x_35); +x_186 = lean_box(0); +x_187 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_187, 0, x_186); +lean_ctor_set(x_187, 1, x_185); +return x_187; +} +} +} +else +{ +lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; uint8_t x_193; uint8_t x_194; uint8_t x_195; uint8_t x_196; uint8_t x_197; uint8_t x_198; uint8_t x_199; uint8_t x_200; uint8_t x_201; uint8_t x_202; uint8_t x_203; uint8_t x_204; uint8_t x_205; lean_object* x_206; lean_object* x_207; +x_188 = lean_ctor_get(x_2, 1); +x_189 = lean_ctor_get(x_2, 2); +x_190 = lean_ctor_get(x_2, 3); +x_191 = lean_ctor_get(x_2, 4); +x_192 = lean_ctor_get(x_2, 5); +x_193 = lean_ctor_get_uint8(x_15, 0); +x_194 = lean_ctor_get_uint8(x_15, 1); +x_195 = lean_ctor_get_uint8(x_15, 2); +x_196 = lean_ctor_get_uint8(x_15, 3); +x_197 = lean_ctor_get_uint8(x_15, 4); +x_198 = lean_ctor_get_uint8(x_15, 6); +x_199 = lean_ctor_get_uint8(x_15, 7); +x_200 = lean_ctor_get_uint8(x_15, 8); +x_201 = lean_ctor_get_uint8(x_15, 9); +x_202 = lean_ctor_get_uint8(x_15, 10); +x_203 = lean_ctor_get_uint8(x_15, 11); +x_204 = lean_ctor_get_uint8(x_15, 12); +lean_dec(x_15); +x_205 = 3; +x_206 = lean_alloc_ctor(0, 0, 13); +lean_ctor_set_uint8(x_206, 0, x_193); +lean_ctor_set_uint8(x_206, 1, x_194); +lean_ctor_set_uint8(x_206, 2, x_195); +lean_ctor_set_uint8(x_206, 3, x_196); +lean_ctor_set_uint8(x_206, 4, x_197); +lean_ctor_set_uint8(x_206, 5, x_205); +lean_ctor_set_uint8(x_206, 6, x_198); +lean_ctor_set_uint8(x_206, 7, x_199); +lean_ctor_set_uint8(x_206, 8, x_200); +lean_ctor_set_uint8(x_206, 9, x_201); +lean_ctor_set_uint8(x_206, 10, x_202); +lean_ctor_set_uint8(x_206, 11, x_203); +lean_ctor_set_uint8(x_206, 12, x_204); +lean_inc(x_192); lean_inc(x_191); lean_inc(x_190); lean_inc(x_189); lean_inc(x_188); -lean_inc(x_187); -lean_ctor_set(x_2, 0, x_205); +lean_ctor_set(x_2, 0, x_206); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); lean_inc(x_1); -x_206 = l_Lean_Meta_mkDecide(x_1, x_2, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_206) == 0) +x_207 = l_Lean_Meta_mkDecide(x_1, x_2, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_207) == 0) { -lean_object* x_207; lean_object* x_208; uint8_t x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; -x_207 = lean_ctor_get(x_206, 0); -lean_inc(x_207); -x_208 = lean_ctor_get(x_206, 1); +lean_object* x_208; lean_object* x_209; uint8_t x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; +x_208 = lean_ctor_get(x_207, 0); lean_inc(x_208); -lean_dec(x_206); -x_209 = 1; -x_210 = lean_alloc_ctor(0, 0, 13); -lean_ctor_set_uint8(x_210, 0, x_192); -lean_ctor_set_uint8(x_210, 1, x_193); -lean_ctor_set_uint8(x_210, 2, x_194); -lean_ctor_set_uint8(x_210, 3, x_195); -lean_ctor_set_uint8(x_210, 4, x_196); -lean_ctor_set_uint8(x_210, 5, x_209); -lean_ctor_set_uint8(x_210, 6, x_197); -lean_ctor_set_uint8(x_210, 7, x_198); -lean_ctor_set_uint8(x_210, 8, x_199); -lean_ctor_set_uint8(x_210, 9, x_200); -lean_ctor_set_uint8(x_210, 10, x_201); -lean_ctor_set_uint8(x_210, 11, x_202); -lean_ctor_set_uint8(x_210, 12, x_203); -x_211 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_211, 0, x_210); -lean_ctor_set(x_211, 1, x_187); -lean_ctor_set(x_211, 2, x_188); -lean_ctor_set(x_211, 3, x_189); -lean_ctor_set(x_211, 4, x_190); -lean_ctor_set(x_211, 5, x_191); +x_209 = lean_ctor_get(x_207, 1); +lean_inc(x_209); +lean_dec(x_207); +x_210 = 1; +x_211 = lean_alloc_ctor(0, 0, 13); +lean_ctor_set_uint8(x_211, 0, x_193); +lean_ctor_set_uint8(x_211, 1, x_194); +lean_ctor_set_uint8(x_211, 2, x_195); +lean_ctor_set_uint8(x_211, 3, x_196); +lean_ctor_set_uint8(x_211, 4, x_197); +lean_ctor_set_uint8(x_211, 5, x_210); +lean_ctor_set_uint8(x_211, 6, x_198); +lean_ctor_set_uint8(x_211, 7, x_199); +lean_ctor_set_uint8(x_211, 8, x_200); +lean_ctor_set_uint8(x_211, 9, x_201); +lean_ctor_set_uint8(x_211, 10, x_202); +lean_ctor_set_uint8(x_211, 11, x_203); +lean_ctor_set_uint8(x_211, 12, x_204); +x_212 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_212, 0, x_211); +lean_ctor_set(x_212, 1, x_188); +lean_ctor_set(x_212, 2, x_189); +lean_ctor_set(x_212, 3, x_190); +lean_ctor_set(x_212, 4, x_191); +lean_ctor_set(x_212, 5, x_192); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -lean_inc(x_207); -x_212 = lean_whnf(x_207, x_211, x_3, x_4, x_5, x_208); -if (lean_obj_tag(x_212) == 0) +lean_inc(x_208); +x_213 = lean_whnf(x_208, x_212, x_3, x_4, x_5, x_209); +if (lean_obj_tag(x_213) == 0) { -lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; uint8_t x_217; -x_213 = lean_ctor_get(x_212, 0); -lean_inc(x_213); -x_214 = lean_ctor_get(x_212, 1); +lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; uint8_t x_218; +x_214 = lean_ctor_get(x_213, 0); lean_inc(x_214); -if (lean_is_exclusive(x_212)) { - lean_ctor_release(x_212, 0); - lean_ctor_release(x_212, 1); - x_215 = x_212; +x_215 = lean_ctor_get(x_213, 1); +lean_inc(x_215); +if (lean_is_exclusive(x_213)) { + lean_ctor_release(x_213, 0); + lean_ctor_release(x_213, 1); + x_216 = x_213; } else { - lean_dec_ref(x_212); - x_215 = lean_box(0); + lean_dec_ref(x_213); + x_216 = lean_box(0); } -x_216 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__5; -x_217 = l_Lean_Expr_isConstOf(x_213, x_216); -if (x_217 == 0) +x_217 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__5; +x_218 = l_Lean_Expr_isConstOf(x_214, x_217); +if (x_218 == 0) { -lean_object* x_218; uint8_t x_219; -x_218 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__7; -x_219 = l_Lean_Expr_isConstOf(x_213, x_218); -lean_dec(x_213); -if (x_219 == 0) +lean_object* x_219; uint8_t x_220; +x_219 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__7; +x_220 = l_Lean_Expr_isConstOf(x_214, x_219); +lean_dec(x_214); +if (x_220 == 0) { -lean_object* x_220; lean_object* x_221; -lean_dec(x_207); +lean_object* x_221; lean_object* x_222; +lean_dec(x_208); lean_dec(x_2); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_220 = lean_box(0); -if (lean_is_scalar(x_215)) { - x_221 = lean_alloc_ctor(0, 2, 0); +x_221 = lean_box(0); +if (lean_is_scalar(x_216)) { + x_222 = lean_alloc_ctor(0, 2, 0); } else { - x_221 = x_215; + x_222 = x_216; } -lean_ctor_set(x_221, 0, x_220); -lean_ctor_set(x_221, 1, x_214); -return x_221; +lean_ctor_set(x_222, 0, x_221); +lean_ctor_set(x_222, 1, x_215); +return x_222; } else { -lean_object* x_222; lean_object* x_223; -lean_dec(x_215); -x_222 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__8; -x_223 = l_Lean_Meta_mkEqRefl(x_222, x_2, x_3, x_4, x_5, x_214); -if (lean_obj_tag(x_223) == 0) +lean_object* x_223; lean_object* x_224; +lean_dec(x_216); +x_223 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__8; +x_224 = l_Lean_Meta_mkEqRefl(x_223, x_2, x_3, x_4, x_5, x_215); +if (lean_obj_tag(x_224) == 0) { -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_224 = lean_ctor_get(x_223, 0); -lean_inc(x_224); -x_225 = lean_ctor_get(x_223, 1); +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_225 = lean_ctor_get(x_224, 0); lean_inc(x_225); -if (lean_is_exclusive(x_223)) { - lean_ctor_release(x_223, 0); - lean_ctor_release(x_223, 1); - x_226 = x_223; +x_226 = lean_ctor_get(x_224, 1); +lean_inc(x_226); +if (lean_is_exclusive(x_224)) { + lean_ctor_release(x_224, 0); + lean_ctor_release(x_224, 1); + x_227 = x_224; } else { - lean_dec_ref(x_223); - x_226 = lean_box(0); + lean_dec_ref(x_224); + x_227 = lean_box(0); } -x_227 = l_Lean_Expr_appArg_x21(x_207); -lean_dec(x_207); -x_228 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__12; -x_229 = lean_array_push(x_228, x_1); -x_230 = lean_array_push(x_229, x_227); -x_231 = lean_array_push(x_230, x_224); -x_232 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__11; -x_233 = l_Lean_mkAppN(x_232, x_231); -x_234 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_234, 0, x_233); -x_235 = l_Lean_Meta_Simp_rewriteCtorEq_x3f___lambda__1___closed__3; -x_236 = lean_unsigned_to_nat(0u); -x_237 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_237, 0, x_235); -lean_ctor_set(x_237, 1, x_234); -lean_ctor_set(x_237, 2, x_236); -x_238 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_238, 0, x_237); -if (lean_is_scalar(x_226)) { - x_239 = lean_alloc_ctor(0, 2, 0); +x_228 = l_Lean_Expr_appArg_x21(x_208); +lean_dec(x_208); +x_229 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__12; +x_230 = lean_array_push(x_229, x_1); +x_231 = lean_array_push(x_230, x_228); +x_232 = lean_array_push(x_231, x_225); +x_233 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__11; +x_234 = l_Lean_mkAppN(x_233, x_232); +x_235 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_235, 0, x_234); +x_236 = l_Lean_Meta_Simp_rewriteCtorEq_x3f___lambda__1___closed__3; +x_237 = lean_unsigned_to_nat(0u); +x_238 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_238, 0, x_236); +lean_ctor_set(x_238, 1, x_235); +lean_ctor_set(x_238, 2, x_237); +x_239 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_239, 0, x_238); +if (lean_is_scalar(x_227)) { + x_240 = lean_alloc_ctor(0, 2, 0); } else { - x_239 = x_226; + x_240 = x_227; } -lean_ctor_set(x_239, 0, x_238); -lean_ctor_set(x_239, 1, x_225); -return x_239; +lean_ctor_set(x_240, 0, x_239); +lean_ctor_set(x_240, 1, x_226); +return x_240; } else { -lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; -lean_dec(x_207); +lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; +lean_dec(x_208); lean_dec(x_1); -x_240 = lean_ctor_get(x_223, 1); -lean_inc(x_240); -if (lean_is_exclusive(x_223)) { - lean_ctor_release(x_223, 0); - lean_ctor_release(x_223, 1); - x_241 = x_223; +x_241 = lean_ctor_get(x_224, 1); +lean_inc(x_241); +if (lean_is_exclusive(x_224)) { + lean_ctor_release(x_224, 0); + lean_ctor_release(x_224, 1); + x_242 = x_224; } else { - lean_dec_ref(x_223); - x_241 = lean_box(0); + lean_dec_ref(x_224); + x_242 = lean_box(0); } -x_242 = lean_box(0); -if (lean_is_scalar(x_241)) { - x_243 = lean_alloc_ctor(0, 2, 0); +x_243 = lean_box(0); +if (lean_is_scalar(x_242)) { + x_244 = lean_alloc_ctor(0, 2, 0); } else { - x_243 = x_241; - lean_ctor_set_tag(x_243, 0); + x_244 = x_242; + lean_ctor_set_tag(x_244, 0); } -lean_ctor_set(x_243, 0, x_242); -lean_ctor_set(x_243, 1, x_240); -return x_243; +lean_ctor_set(x_244, 0, x_243); +lean_ctor_set(x_244, 1, x_241); +return x_244; } } } else { -lean_object* x_244; lean_object* x_245; -lean_dec(x_215); -lean_dec(x_213); -x_244 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__13; -x_245 = l_Lean_Meta_mkEqRefl(x_244, x_2, x_3, x_4, x_5, x_214); -if (lean_obj_tag(x_245) == 0) -{ -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_246 = lean_ctor_get(x_245, 0); -lean_inc(x_246); -x_247 = lean_ctor_get(x_245, 1); +lean_object* x_245; lean_object* x_246; +lean_dec(x_216); +lean_dec(x_214); +x_245 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__13; +x_246 = l_Lean_Meta_mkEqRefl(x_245, x_2, x_3, x_4, x_5, x_215); +if (lean_obj_tag(x_246) == 0) +{ +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_247 = lean_ctor_get(x_246, 0); lean_inc(x_247); -if (lean_is_exclusive(x_245)) { - lean_ctor_release(x_245, 0); - lean_ctor_release(x_245, 1); - x_248 = x_245; +x_248 = lean_ctor_get(x_246, 1); +lean_inc(x_248); +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_245); - x_248 = lean_box(0); + lean_dec_ref(x_246); + x_249 = lean_box(0); } -x_249 = l_Lean_Expr_appArg_x21(x_207); -lean_dec(x_207); -x_250 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__12; -x_251 = lean_array_push(x_250, x_1); -x_252 = lean_array_push(x_251, x_249); -x_253 = lean_array_push(x_252, x_246); -x_254 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__17; -x_255 = l_Lean_mkAppN(x_254, x_253); -x_256 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_256, 0, x_255); -x_257 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__14; -x_258 = lean_unsigned_to_nat(0u); -x_259 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_259, 0, x_257); -lean_ctor_set(x_259, 1, x_256); -lean_ctor_set(x_259, 2, x_258); -x_260 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_260, 0, x_259); -if (lean_is_scalar(x_248)) { - x_261 = lean_alloc_ctor(0, 2, 0); +x_250 = l_Lean_Expr_appArg_x21(x_208); +lean_dec(x_208); +x_251 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__12; +x_252 = lean_array_push(x_251, x_1); +x_253 = lean_array_push(x_252, x_250); +x_254 = lean_array_push(x_253, x_247); +x_255 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__17; +x_256 = l_Lean_mkAppN(x_255, x_254); +x_257 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_257, 0, x_256); +x_258 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__14; +x_259 = lean_unsigned_to_nat(0u); +x_260 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_260, 0, x_258); +lean_ctor_set(x_260, 1, x_257); +lean_ctor_set(x_260, 2, x_259); +x_261 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_261, 0, x_260); +if (lean_is_scalar(x_249)) { + x_262 = lean_alloc_ctor(0, 2, 0); } else { - x_261 = x_248; + x_262 = x_249; } -lean_ctor_set(x_261, 0, x_260); -lean_ctor_set(x_261, 1, x_247); -return x_261; +lean_ctor_set(x_262, 0, x_261); +lean_ctor_set(x_262, 1, x_248); +return x_262; } else { -lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; -lean_dec(x_207); +lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; +lean_dec(x_208); lean_dec(x_1); -x_262 = lean_ctor_get(x_245, 1); -lean_inc(x_262); -if (lean_is_exclusive(x_245)) { - lean_ctor_release(x_245, 0); - lean_ctor_release(x_245, 1); - x_263 = x_245; +x_263 = lean_ctor_get(x_246, 1); +lean_inc(x_263); +if (lean_is_exclusive(x_246)) { + lean_ctor_release(x_246, 0); + lean_ctor_release(x_246, 1); + x_264 = x_246; } else { - lean_dec_ref(x_245); - x_263 = lean_box(0); + lean_dec_ref(x_246); + x_264 = lean_box(0); } -x_264 = lean_box(0); -if (lean_is_scalar(x_263)) { - x_265 = lean_alloc_ctor(0, 2, 0); +x_265 = lean_box(0); +if (lean_is_scalar(x_264)) { + x_266 = lean_alloc_ctor(0, 2, 0); } else { - x_265 = x_263; - lean_ctor_set_tag(x_265, 0); + x_266 = x_264; + lean_ctor_set_tag(x_266, 0); } -lean_ctor_set(x_265, 0, x_264); -lean_ctor_set(x_265, 1, x_262); -return x_265; +lean_ctor_set(x_266, 0, x_265); +lean_ctor_set(x_266, 1, x_263); +return x_266; } } } else { -lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; -lean_dec(x_207); +lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; +lean_dec(x_208); lean_dec(x_2); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_266 = lean_ctor_get(x_212, 1); -lean_inc(x_266); -if (lean_is_exclusive(x_212)) { - lean_ctor_release(x_212, 0); - lean_ctor_release(x_212, 1); - x_267 = x_212; +x_267 = lean_ctor_get(x_213, 1); +lean_inc(x_267); +if (lean_is_exclusive(x_213)) { + lean_ctor_release(x_213, 0); + lean_ctor_release(x_213, 1); + x_268 = x_213; } else { - lean_dec_ref(x_212); - x_267 = lean_box(0); + lean_dec_ref(x_213); + x_268 = lean_box(0); } -x_268 = lean_box(0); -if (lean_is_scalar(x_267)) { - x_269 = lean_alloc_ctor(0, 2, 0); +x_269 = lean_box(0); +if (lean_is_scalar(x_268)) { + x_270 = lean_alloc_ctor(0, 2, 0); } else { - x_269 = x_267; - lean_ctor_set_tag(x_269, 0); + x_270 = x_268; + lean_ctor_set_tag(x_270, 0); } -lean_ctor_set(x_269, 0, x_268); -lean_ctor_set(x_269, 1, x_266); -return x_269; +lean_ctor_set(x_270, 0, x_269); +lean_ctor_set(x_270, 1, x_267); +return x_270; } } else { -lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; +lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_dec(x_2); +lean_dec(x_192); lean_dec(x_191); lean_dec(x_190); lean_dec(x_189); lean_dec(x_188); -lean_dec(x_187); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_270 = lean_ctor_get(x_206, 1); -lean_inc(x_270); -if (lean_is_exclusive(x_206)) { - lean_ctor_release(x_206, 0); - lean_ctor_release(x_206, 1); - x_271 = x_206; +x_271 = lean_ctor_get(x_207, 1); +lean_inc(x_271); +if (lean_is_exclusive(x_207)) { + lean_ctor_release(x_207, 0); + lean_ctor_release(x_207, 1); + x_272 = x_207; } else { - lean_dec_ref(x_206); - x_271 = lean_box(0); + lean_dec_ref(x_207); + x_272 = lean_box(0); } -x_272 = lean_box(0); -if (lean_is_scalar(x_271)) { - x_273 = lean_alloc_ctor(0, 2, 0); +x_273 = lean_box(0); +if (lean_is_scalar(x_272)) { + x_274 = lean_alloc_ctor(0, 2, 0); } else { - x_273 = x_271; - lean_ctor_set_tag(x_273, 0); + x_274 = x_272; + lean_ctor_set_tag(x_274, 0); } -lean_ctor_set(x_273, 0, x_272); -lean_ctor_set(x_273, 1, x_270); -return x_273; +lean_ctor_set(x_274, 0, x_273); +lean_ctor_set(x_274, 1, x_271); +return x_274; } } } else { -lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; uint8_t x_280; uint8_t x_281; uint8_t x_282; uint8_t x_283; uint8_t x_284; uint8_t x_285; uint8_t x_286; uint8_t x_287; uint8_t x_288; uint8_t x_289; uint8_t x_290; uint8_t x_291; lean_object* x_292; uint8_t x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; -x_274 = lean_ctor_get(x_2, 0); -x_275 = lean_ctor_get(x_2, 1); -x_276 = lean_ctor_get(x_2, 2); -x_277 = lean_ctor_get(x_2, 3); -x_278 = lean_ctor_get(x_2, 4); -x_279 = lean_ctor_get(x_2, 5); +lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; uint8_t x_281; uint8_t x_282; uint8_t x_283; uint8_t x_284; uint8_t x_285; uint8_t x_286; uint8_t x_287; uint8_t x_288; uint8_t x_289; uint8_t x_290; uint8_t x_291; uint8_t x_292; lean_object* x_293; uint8_t x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; +x_275 = lean_ctor_get(x_2, 0); +x_276 = lean_ctor_get(x_2, 1); +x_277 = lean_ctor_get(x_2, 2); +x_278 = lean_ctor_get(x_2, 3); +x_279 = lean_ctor_get(x_2, 4); +x_280 = lean_ctor_get(x_2, 5); +lean_inc(x_280); lean_inc(x_279); lean_inc(x_278); lean_inc(x_277); lean_inc(x_276); lean_inc(x_275); -lean_inc(x_274); lean_dec(x_2); -x_280 = lean_ctor_get_uint8(x_274, 0); -x_281 = lean_ctor_get_uint8(x_274, 1); -x_282 = lean_ctor_get_uint8(x_274, 2); -x_283 = lean_ctor_get_uint8(x_274, 3); -x_284 = lean_ctor_get_uint8(x_274, 4); -x_285 = lean_ctor_get_uint8(x_274, 6); -x_286 = lean_ctor_get_uint8(x_274, 7); -x_287 = lean_ctor_get_uint8(x_274, 8); -x_288 = lean_ctor_get_uint8(x_274, 9); -x_289 = lean_ctor_get_uint8(x_274, 10); -x_290 = lean_ctor_get_uint8(x_274, 11); -x_291 = lean_ctor_get_uint8(x_274, 12); -if (lean_is_exclusive(x_274)) { - x_292 = x_274; +x_281 = lean_ctor_get_uint8(x_275, 0); +x_282 = lean_ctor_get_uint8(x_275, 1); +x_283 = lean_ctor_get_uint8(x_275, 2); +x_284 = lean_ctor_get_uint8(x_275, 3); +x_285 = lean_ctor_get_uint8(x_275, 4); +x_286 = lean_ctor_get_uint8(x_275, 6); +x_287 = lean_ctor_get_uint8(x_275, 7); +x_288 = lean_ctor_get_uint8(x_275, 8); +x_289 = lean_ctor_get_uint8(x_275, 9); +x_290 = lean_ctor_get_uint8(x_275, 10); +x_291 = lean_ctor_get_uint8(x_275, 11); +x_292 = lean_ctor_get_uint8(x_275, 12); +if (lean_is_exclusive(x_275)) { + x_293 = x_275; } else { - lean_dec_ref(x_274); - x_292 = lean_box(0); + lean_dec_ref(x_275); + x_293 = lean_box(0); } -x_293 = 3; -if (lean_is_scalar(x_292)) { - x_294 = lean_alloc_ctor(0, 0, 13); +x_294 = 3; +if (lean_is_scalar(x_293)) { + x_295 = lean_alloc_ctor(0, 0, 13); } else { - x_294 = x_292; -} -lean_ctor_set_uint8(x_294, 0, x_280); -lean_ctor_set_uint8(x_294, 1, x_281); -lean_ctor_set_uint8(x_294, 2, x_282); -lean_ctor_set_uint8(x_294, 3, x_283); -lean_ctor_set_uint8(x_294, 4, x_284); -lean_ctor_set_uint8(x_294, 5, x_293); -lean_ctor_set_uint8(x_294, 6, x_285); -lean_ctor_set_uint8(x_294, 7, x_286); -lean_ctor_set_uint8(x_294, 8, x_287); -lean_ctor_set_uint8(x_294, 9, x_288); -lean_ctor_set_uint8(x_294, 10, x_289); -lean_ctor_set_uint8(x_294, 11, x_290); -lean_ctor_set_uint8(x_294, 12, x_291); + x_295 = x_293; +} +lean_ctor_set_uint8(x_295, 0, x_281); +lean_ctor_set_uint8(x_295, 1, x_282); +lean_ctor_set_uint8(x_295, 2, x_283); +lean_ctor_set_uint8(x_295, 3, x_284); +lean_ctor_set_uint8(x_295, 4, x_285); +lean_ctor_set_uint8(x_295, 5, x_294); +lean_ctor_set_uint8(x_295, 6, x_286); +lean_ctor_set_uint8(x_295, 7, x_287); +lean_ctor_set_uint8(x_295, 8, x_288); +lean_ctor_set_uint8(x_295, 9, x_289); +lean_ctor_set_uint8(x_295, 10, x_290); +lean_ctor_set_uint8(x_295, 11, x_291); +lean_ctor_set_uint8(x_295, 12, x_292); +lean_inc(x_280); lean_inc(x_279); lean_inc(x_278); lean_inc(x_277); lean_inc(x_276); -lean_inc(x_275); -x_295 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_295, 0, x_294); -lean_ctor_set(x_295, 1, x_275); -lean_ctor_set(x_295, 2, x_276); -lean_ctor_set(x_295, 3, x_277); -lean_ctor_set(x_295, 4, x_278); -lean_ctor_set(x_295, 5, x_279); +x_296 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_296, 0, x_295); +lean_ctor_set(x_296, 1, x_276); +lean_ctor_set(x_296, 2, x_277); +lean_ctor_set(x_296, 3, x_278); +lean_ctor_set(x_296, 4, x_279); +lean_ctor_set(x_296, 5, x_280); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -lean_inc(x_295); +lean_inc(x_296); lean_inc(x_1); -x_296 = l_Lean_Meta_mkDecide(x_1, x_295, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_296) == 0) +x_297 = l_Lean_Meta_mkDecide(x_1, x_296, x_3, x_4, x_5, x_6); +if (lean_obj_tag(x_297) == 0) { -lean_object* x_297; lean_object* x_298; uint8_t x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; -x_297 = lean_ctor_get(x_296, 0); -lean_inc(x_297); -x_298 = lean_ctor_get(x_296, 1); +lean_object* x_298; lean_object* x_299; uint8_t x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; +x_298 = lean_ctor_get(x_297, 0); lean_inc(x_298); -lean_dec(x_296); -x_299 = 1; -x_300 = lean_alloc_ctor(0, 0, 13); -lean_ctor_set_uint8(x_300, 0, x_280); -lean_ctor_set_uint8(x_300, 1, x_281); -lean_ctor_set_uint8(x_300, 2, x_282); -lean_ctor_set_uint8(x_300, 3, x_283); -lean_ctor_set_uint8(x_300, 4, x_284); -lean_ctor_set_uint8(x_300, 5, x_299); -lean_ctor_set_uint8(x_300, 6, x_285); -lean_ctor_set_uint8(x_300, 7, x_286); -lean_ctor_set_uint8(x_300, 8, x_287); -lean_ctor_set_uint8(x_300, 9, x_288); -lean_ctor_set_uint8(x_300, 10, x_289); -lean_ctor_set_uint8(x_300, 11, x_290); -lean_ctor_set_uint8(x_300, 12, x_291); -x_301 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_301, 0, x_300); -lean_ctor_set(x_301, 1, x_275); -lean_ctor_set(x_301, 2, x_276); -lean_ctor_set(x_301, 3, x_277); -lean_ctor_set(x_301, 4, x_278); -lean_ctor_set(x_301, 5, x_279); +x_299 = lean_ctor_get(x_297, 1); +lean_inc(x_299); +lean_dec(x_297); +x_300 = 1; +x_301 = lean_alloc_ctor(0, 0, 13); +lean_ctor_set_uint8(x_301, 0, x_281); +lean_ctor_set_uint8(x_301, 1, x_282); +lean_ctor_set_uint8(x_301, 2, x_283); +lean_ctor_set_uint8(x_301, 3, x_284); +lean_ctor_set_uint8(x_301, 4, x_285); +lean_ctor_set_uint8(x_301, 5, x_300); +lean_ctor_set_uint8(x_301, 6, x_286); +lean_ctor_set_uint8(x_301, 7, x_287); +lean_ctor_set_uint8(x_301, 8, x_288); +lean_ctor_set_uint8(x_301, 9, x_289); +lean_ctor_set_uint8(x_301, 10, x_290); +lean_ctor_set_uint8(x_301, 11, x_291); +lean_ctor_set_uint8(x_301, 12, x_292); +x_302 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_302, 0, x_301); +lean_ctor_set(x_302, 1, x_276); +lean_ctor_set(x_302, 2, x_277); +lean_ctor_set(x_302, 3, x_278); +lean_ctor_set(x_302, 4, x_279); +lean_ctor_set(x_302, 5, x_280); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -lean_inc(x_297); -x_302 = lean_whnf(x_297, x_301, x_3, x_4, x_5, x_298); -if (lean_obj_tag(x_302) == 0) -{ -lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; uint8_t x_307; -x_303 = lean_ctor_get(x_302, 0); -lean_inc(x_303); -x_304 = lean_ctor_get(x_302, 1); +lean_inc(x_298); +x_303 = lean_whnf(x_298, x_302, x_3, x_4, x_5, x_299); +if (lean_obj_tag(x_303) == 0) +{ +lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; uint8_t x_308; +x_304 = lean_ctor_get(x_303, 0); lean_inc(x_304); -if (lean_is_exclusive(x_302)) { - lean_ctor_release(x_302, 0); - lean_ctor_release(x_302, 1); - x_305 = x_302; +x_305 = lean_ctor_get(x_303, 1); +lean_inc(x_305); +if (lean_is_exclusive(x_303)) { + lean_ctor_release(x_303, 0); + lean_ctor_release(x_303, 1); + x_306 = x_303; } else { - lean_dec_ref(x_302); - x_305 = lean_box(0); + lean_dec_ref(x_303); + x_306 = lean_box(0); } -x_306 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__5; -x_307 = l_Lean_Expr_isConstOf(x_303, x_306); -if (x_307 == 0) +x_307 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__5; +x_308 = l_Lean_Expr_isConstOf(x_304, x_307); +if (x_308 == 0) { -lean_object* x_308; uint8_t x_309; -x_308 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__7; -x_309 = l_Lean_Expr_isConstOf(x_303, x_308); -lean_dec(x_303); -if (x_309 == 0) +lean_object* x_309; uint8_t x_310; +x_309 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__7; +x_310 = l_Lean_Expr_isConstOf(x_304, x_309); +lean_dec(x_304); +if (x_310 == 0) { -lean_object* x_310; lean_object* x_311; -lean_dec(x_297); -lean_dec(x_295); +lean_object* x_311; lean_object* x_312; +lean_dec(x_298); +lean_dec(x_296); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_310 = lean_box(0); -if (lean_is_scalar(x_305)) { - x_311 = lean_alloc_ctor(0, 2, 0); +x_311 = lean_box(0); +if (lean_is_scalar(x_306)) { + x_312 = lean_alloc_ctor(0, 2, 0); } else { - x_311 = x_305; + x_312 = x_306; } -lean_ctor_set(x_311, 0, x_310); -lean_ctor_set(x_311, 1, x_304); -return x_311; +lean_ctor_set(x_312, 0, x_311); +lean_ctor_set(x_312, 1, x_305); +return x_312; } else { -lean_object* x_312; lean_object* x_313; -lean_dec(x_305); -x_312 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__8; -x_313 = l_Lean_Meta_mkEqRefl(x_312, x_295, x_3, x_4, x_5, x_304); -if (lean_obj_tag(x_313) == 0) +lean_object* x_313; lean_object* x_314; +lean_dec(x_306); +x_313 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__8; +x_314 = l_Lean_Meta_mkEqRefl(x_313, x_296, x_3, x_4, x_5, x_305); +if (lean_obj_tag(x_314) == 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; 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_314 = lean_ctor_get(x_313, 0); -lean_inc(x_314); -x_315 = lean_ctor_get(x_313, 1); +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; +x_315 = lean_ctor_get(x_314, 0); lean_inc(x_315); -if (lean_is_exclusive(x_313)) { - lean_ctor_release(x_313, 0); - lean_ctor_release(x_313, 1); - x_316 = x_313; +x_316 = lean_ctor_get(x_314, 1); +lean_inc(x_316); +if (lean_is_exclusive(x_314)) { + lean_ctor_release(x_314, 0); + lean_ctor_release(x_314, 1); + x_317 = x_314; } else { - lean_dec_ref(x_313); - x_316 = lean_box(0); -} -x_317 = l_Lean_Expr_appArg_x21(x_297); -lean_dec(x_297); -x_318 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__12; -x_319 = lean_array_push(x_318, x_1); -x_320 = lean_array_push(x_319, x_317); -x_321 = lean_array_push(x_320, x_314); -x_322 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__11; -x_323 = l_Lean_mkAppN(x_322, x_321); -x_324 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_324, 0, x_323); -x_325 = l_Lean_Meta_Simp_rewriteCtorEq_x3f___lambda__1___closed__3; -x_326 = lean_unsigned_to_nat(0u); -x_327 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_327, 0, x_325); -lean_ctor_set(x_327, 1, x_324); -lean_ctor_set(x_327, 2, x_326); -x_328 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_328, 0, x_327); -if (lean_is_scalar(x_316)) { - x_329 = lean_alloc_ctor(0, 2, 0); + lean_dec_ref(x_314); + x_317 = lean_box(0); +} +x_318 = l_Lean_Expr_appArg_x21(x_298); +lean_dec(x_298); +x_319 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__12; +x_320 = lean_array_push(x_319, x_1); +x_321 = lean_array_push(x_320, x_318); +x_322 = lean_array_push(x_321, x_315); +x_323 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__11; +x_324 = l_Lean_mkAppN(x_323, x_322); +x_325 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_325, 0, x_324); +x_326 = l_Lean_Meta_Simp_rewriteCtorEq_x3f___lambda__1___closed__3; +x_327 = lean_unsigned_to_nat(0u); +x_328 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_328, 0, x_326); +lean_ctor_set(x_328, 1, x_325); +lean_ctor_set(x_328, 2, x_327); +x_329 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_329, 0, x_328); +if (lean_is_scalar(x_317)) { + x_330 = lean_alloc_ctor(0, 2, 0); } else { - x_329 = x_316; + x_330 = x_317; } -lean_ctor_set(x_329, 0, x_328); -lean_ctor_set(x_329, 1, x_315); -return x_329; +lean_ctor_set(x_330, 0, x_329); +lean_ctor_set(x_330, 1, x_316); +return x_330; } else { -lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; -lean_dec(x_297); +lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; +lean_dec(x_298); lean_dec(x_1); -x_330 = lean_ctor_get(x_313, 1); -lean_inc(x_330); -if (lean_is_exclusive(x_313)) { - lean_ctor_release(x_313, 0); - lean_ctor_release(x_313, 1); - x_331 = x_313; +x_331 = lean_ctor_get(x_314, 1); +lean_inc(x_331); +if (lean_is_exclusive(x_314)) { + lean_ctor_release(x_314, 0); + lean_ctor_release(x_314, 1); + x_332 = x_314; } else { - lean_dec_ref(x_313); - x_331 = lean_box(0); + lean_dec_ref(x_314); + x_332 = lean_box(0); } -x_332 = lean_box(0); -if (lean_is_scalar(x_331)) { - x_333 = lean_alloc_ctor(0, 2, 0); +x_333 = lean_box(0); +if (lean_is_scalar(x_332)) { + x_334 = lean_alloc_ctor(0, 2, 0); } else { - x_333 = x_331; - lean_ctor_set_tag(x_333, 0); + x_334 = x_332; + lean_ctor_set_tag(x_334, 0); } -lean_ctor_set(x_333, 0, x_332); -lean_ctor_set(x_333, 1, x_330); -return x_333; +lean_ctor_set(x_334, 0, x_333); +lean_ctor_set(x_334, 1, x_331); +return x_334; } } } else { -lean_object* x_334; lean_object* x_335; -lean_dec(x_305); -lean_dec(x_303); -x_334 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__13; -x_335 = l_Lean_Meta_mkEqRefl(x_334, x_295, x_3, x_4, x_5, x_304); -if (lean_obj_tag(x_335) == 0) +lean_object* x_335; lean_object* x_336; +lean_dec(x_306); +lean_dec(x_304); +x_335 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__13; +x_336 = l_Lean_Meta_mkEqRefl(x_335, x_296, x_3, x_4, x_5, x_305); +if (lean_obj_tag(x_336) == 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; 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_336 = lean_ctor_get(x_335, 0); -lean_inc(x_336); -x_337 = lean_ctor_get(x_335, 1); +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_337 = lean_ctor_get(x_336, 0); lean_inc(x_337); -if (lean_is_exclusive(x_335)) { - lean_ctor_release(x_335, 0); - lean_ctor_release(x_335, 1); - x_338 = x_335; +x_338 = lean_ctor_get(x_336, 1); +lean_inc(x_338); +if (lean_is_exclusive(x_336)) { + lean_ctor_release(x_336, 0); + lean_ctor_release(x_336, 1); + x_339 = x_336; } else { - lean_dec_ref(x_335); - x_338 = lean_box(0); -} -x_339 = l_Lean_Expr_appArg_x21(x_297); -lean_dec(x_297); -x_340 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__12; -x_341 = lean_array_push(x_340, x_1); -x_342 = lean_array_push(x_341, x_339); -x_343 = lean_array_push(x_342, x_336); -x_344 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__17; -x_345 = l_Lean_mkAppN(x_344, x_343); -x_346 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_346, 0, x_345); -x_347 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__14; -x_348 = lean_unsigned_to_nat(0u); -x_349 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_349, 0, x_347); -lean_ctor_set(x_349, 1, x_346); -lean_ctor_set(x_349, 2, x_348); -x_350 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_350, 0, x_349); -if (lean_is_scalar(x_338)) { - x_351 = lean_alloc_ctor(0, 2, 0); + lean_dec_ref(x_336); + x_339 = lean_box(0); +} +x_340 = l_Lean_Expr_appArg_x21(x_298); +lean_dec(x_298); +x_341 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__12; +x_342 = lean_array_push(x_341, x_1); +x_343 = lean_array_push(x_342, x_340); +x_344 = lean_array_push(x_343, x_337); +x_345 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__17; +x_346 = l_Lean_mkAppN(x_345, x_344); +x_347 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_347, 0, x_346); +x_348 = l_Lean_Meta_Simp_rewriteUsingDecide_x3f___closed__14; +x_349 = lean_unsigned_to_nat(0u); +x_350 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_350, 0, x_348); +lean_ctor_set(x_350, 1, x_347); +lean_ctor_set(x_350, 2, x_349); +x_351 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_351, 0, x_350); +if (lean_is_scalar(x_339)) { + x_352 = lean_alloc_ctor(0, 2, 0); } else { - x_351 = x_338; + x_352 = x_339; } -lean_ctor_set(x_351, 0, x_350); -lean_ctor_set(x_351, 1, x_337); -return x_351; +lean_ctor_set(x_352, 0, x_351); +lean_ctor_set(x_352, 1, x_338); +return x_352; } else { -lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; -lean_dec(x_297); +lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; +lean_dec(x_298); lean_dec(x_1); -x_352 = lean_ctor_get(x_335, 1); -lean_inc(x_352); -if (lean_is_exclusive(x_335)) { - lean_ctor_release(x_335, 0); - lean_ctor_release(x_335, 1); - x_353 = x_335; +x_353 = lean_ctor_get(x_336, 1); +lean_inc(x_353); +if (lean_is_exclusive(x_336)) { + lean_ctor_release(x_336, 0); + lean_ctor_release(x_336, 1); + x_354 = x_336; } else { - lean_dec_ref(x_335); - x_353 = lean_box(0); + lean_dec_ref(x_336); + x_354 = lean_box(0); } -x_354 = lean_box(0); -if (lean_is_scalar(x_353)) { - x_355 = lean_alloc_ctor(0, 2, 0); +x_355 = lean_box(0); +if (lean_is_scalar(x_354)) { + x_356 = lean_alloc_ctor(0, 2, 0); } else { - x_355 = x_353; - lean_ctor_set_tag(x_355, 0); + x_356 = x_354; + lean_ctor_set_tag(x_356, 0); } -lean_ctor_set(x_355, 0, x_354); -lean_ctor_set(x_355, 1, x_352); -return x_355; +lean_ctor_set(x_356, 0, x_355); +lean_ctor_set(x_356, 1, x_353); +return x_356; } } } else { -lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; -lean_dec(x_297); -lean_dec(x_295); +lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; +lean_dec(x_298); +lean_dec(x_296); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_356 = lean_ctor_get(x_302, 1); -lean_inc(x_356); -if (lean_is_exclusive(x_302)) { - lean_ctor_release(x_302, 0); - lean_ctor_release(x_302, 1); - x_357 = x_302; +x_357 = lean_ctor_get(x_303, 1); +lean_inc(x_357); +if (lean_is_exclusive(x_303)) { + lean_ctor_release(x_303, 0); + lean_ctor_release(x_303, 1); + x_358 = x_303; } else { - lean_dec_ref(x_302); - x_357 = lean_box(0); + lean_dec_ref(x_303); + x_358 = lean_box(0); } -x_358 = lean_box(0); -if (lean_is_scalar(x_357)) { - x_359 = lean_alloc_ctor(0, 2, 0); +x_359 = lean_box(0); +if (lean_is_scalar(x_358)) { + x_360 = lean_alloc_ctor(0, 2, 0); } else { - x_359 = x_357; - lean_ctor_set_tag(x_359, 0); + x_360 = x_358; + lean_ctor_set_tag(x_360, 0); } -lean_ctor_set(x_359, 0, x_358); -lean_ctor_set(x_359, 1, x_356); -return x_359; +lean_ctor_set(x_360, 0, x_359); +lean_ctor_set(x_360, 1, x_357); +return x_360; } } else { -lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; -lean_dec(x_295); +lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; +lean_dec(x_296); +lean_dec(x_280); lean_dec(x_279); lean_dec(x_278); lean_dec(x_277); lean_dec(x_276); -lean_dec(x_275); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_360 = lean_ctor_get(x_296, 1); -lean_inc(x_360); -if (lean_is_exclusive(x_296)) { - lean_ctor_release(x_296, 0); - lean_ctor_release(x_296, 1); - x_361 = x_296; +x_361 = lean_ctor_get(x_297, 1); +lean_inc(x_361); +if (lean_is_exclusive(x_297)) { + lean_ctor_release(x_297, 0); + lean_ctor_release(x_297, 1); + x_362 = x_297; } else { - lean_dec_ref(x_296); - x_361 = lean_box(0); + lean_dec_ref(x_297); + x_362 = lean_box(0); } -x_362 = lean_box(0); -if (lean_is_scalar(x_361)) { - x_363 = lean_alloc_ctor(0, 2, 0); +x_363 = lean_box(0); +if (lean_is_scalar(x_362)) { + x_364 = lean_alloc_ctor(0, 2, 0); } else { - x_363 = x_361; - lean_ctor_set_tag(x_363, 0); + x_364 = x_362; + lean_ctor_set_tag(x_364, 0); } -lean_ctor_set(x_363, 0, x_362); -lean_ctor_set(x_363, 1, x_360); -return x_363; +lean_ctor_set(x_364, 0, x_363); +lean_ctor_set(x_364, 1, x_361); +return x_364; } } } else { -lean_object* x_364; lean_object* x_365; +lean_object* x_365; lean_object* x_366; lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_364 = lean_box(0); -x_365 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_365, 0, x_364); -lean_ctor_set(x_365, 1, x_6); -return x_365; +x_365 = lean_box(0); +x_366 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_366, 0, x_365); +lean_ctor_set(x_366, 1, x_6); +return x_366; } } else { -lean_object* x_366; lean_object* x_367; +lean_object* x_367; lean_object* x_368; +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_366 = lean_box(0); -x_367 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_367, 0, x_366); -lean_ctor_set(x_367, 1, x_6); -return x_367; +x_367 = lean_box(0); +x_368 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_368, 0, x_367); +lean_ctor_set(x_368, 1, x_6); +return x_368; } } else { -lean_object* x_368; lean_object* x_369; +lean_object* x_369; lean_object* x_370; lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_368 = lean_box(0); -x_369 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_369, 0, x_368); -lean_ctor_set(x_369, 1, x_6); -return x_369; +x_369 = lean_box(0); +x_370 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_370, 0, x_369); +lean_ctor_set(x_370, 1, x_6); +return x_370; } } else { -lean_object* x_370; lean_object* x_371; +lean_object* x_371; lean_object* x_372; lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_370 = lean_box(0); -x_371 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_371, 0, x_370); -lean_ctor_set(x_371, 1, x_6); -return x_371; +x_371 = lean_box(0); +x_372 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_372, 0, x_371); +lean_ctor_set(x_372, 1, x_6); +return x_372; } } } diff --git a/stage0/stdlib/Lean/Meta/Tactic/Simp/SimpAll.c b/stage0/stdlib/Lean/Meta/Tactic/Simp/SimpAll.c index 5b2a95c13e4b..47255924b98f 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Simp/SimpAll.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Simp/SimpAll.c @@ -147,6 +147,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_SimpAll_0__Lean_Meta_ LEAN_EXPORT lean_object* l_Lean_Meta_SimpAll_main(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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_Std_Range_forIn_loop___at___private_Lean_Meta_Tactic_Simp_SimpAll_0__Lean_Meta_SimpAll_loop___spec__8___closed__16; +lean_object* l_Lean_Expr_consumeMData(lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_SimpAll_main___spec__2___closed__2; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); lean_object* l_Lean_MessageData_ofName(lean_object*); @@ -4379,36 +4380,38 @@ 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; uint8_t x_9; size_t x_10; size_t x_11; +lean_object* x_6; 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_6 = lean_array_uget(x_1, x_2); x_7 = lean_ctor_get(x_6, 3); lean_inc(x_7); -x_8 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_SimpAll_main___spec__2___closed__2; -x_9 = l_Lean_Expr_isConstOf(x_7, x_8); -x_10 = 1; -x_11 = lean_usize_add(x_2, x_10); -if (x_9 == 0) +x_8 = l_Lean_Expr_consumeMData(x_7); +x_9 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_SimpAll_main___spec__2___closed__2; +x_10 = l_Lean_Expr_isConstOf(x_8, x_9); +lean_dec(x_8); +x_11 = 1; +x_12 = lean_usize_add(x_2, x_11); +if (x_10 == 0) { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_12 = lean_ctor_get(x_6, 1); -lean_inc(x_12); -x_13 = lean_ctor_get(x_6, 4); +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_13 = lean_ctor_get(x_6, 1); lean_inc(x_13); +x_14 = lean_ctor_get(x_6, 4); +lean_inc(x_14); lean_dec(x_6); -x_14 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_14, 0, x_12); -lean_ctor_set(x_14, 1, x_7); -lean_ctor_set(x_14, 2, x_13); -x_15 = lean_array_push(x_4, x_14); -x_2 = x_11; -x_4 = x_15; +x_15 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_15, 0, x_13); +lean_ctor_set(x_15, 1, x_7); +lean_ctor_set(x_15, 2, x_14); +x_16 = lean_array_push(x_4, x_15); +x_2 = x_12; +x_4 = x_16; goto _start; } else { lean_dec(x_7); lean_dec(x_6); -x_2 = x_11; +x_2 = x_12; goto _start; } } diff --git a/stage0/stdlib/Lean/Meta/Tactic/Simp/SimpTheorems.c b/stage0/stdlib/Lean/Meta/Tactic/Simp/SimpTheorems.c index 093cbd3dc814..b95a5d22707b 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Simp/SimpTheorems.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Simp/SimpTheorems.c @@ -332,7 +332,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_ppSimpTheorem___rarg___lambda__2(lean_objec 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_checkTypeIsProp___closed__1; lean_object* l_Lean_Meta_mkEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PersistentHashMap_empty___at_Lean_Meta_SimpTheorems_lemmaNames___default___spec__1___closed__5; static lean_object* l___auto____x40_Lean_Meta_Tactic_Simp_SimpTheorems___hyg_5451____closed__23; lean_object* l_Lean_Syntax_node3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -515,6 +514,7 @@ uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_registerSimpAttr___closed__4; static lean_object* l_Lean_Parser_Command_registerSimpAttr___closed__20; 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*); +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(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_mkSimpTheoremCore___closed__3; LEAN_EXPORT lean_object* l_Lean_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_mkSimpTheorems___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -5774,7 +5774,7 @@ static lean_object* _init_l_Lean_Meta_DiscrTree_insertCore___at_Lean_Meta_addSim 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_Meta_DiscrTree_insertCore___at_Lean_Meta_addSimpTheoremEntry___spec__1___closed__1; x_2 = l_Lean_Meta_DiscrTree_insertCore___at_Lean_Meta_addSimpTheoremEntry___spec__1___closed__2; -x_3 = lean_unsigned_to_nat(366u); +x_3 = lean_unsigned_to_nat(396u); x_4 = lean_unsigned_to_nat(23u); x_5 = l_Lean_Meta_DiscrTree_insertCore___at_Lean_Meta_addSimpTheoremEntry___spec__1___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -13477,7 +13477,7 @@ lean_inc(x_16); x_17 = lean_ctor_get(x_15, 1); lean_inc(x_17); lean_dec(x_15); -x_18 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_16, x_7, x_8, x_9, x_10, x_17); +x_18 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_16, x_7, x_8, x_9, x_10, x_17); x_19 = lean_ctor_get(x_18, 0); lean_inc(x_19); x_20 = lean_ctor_get(x_18, 1); @@ -18954,7 +18954,7 @@ static lean_object* _init_l_Lean_Meta_SimpTheorem_getValue___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_Meta_SimpTheorem_getValue___closed__1; x_2 = l_Lean_Meta_SimpTheorem_getValue___closed__2; -x_3 = lean_unsigned_to_nat(1481u); +x_3 = lean_unsigned_to_nat(1485u); x_4 = lean_unsigned_to_nat(18u); x_5 = l_Lean_Meta_SimpTheorem_getValue___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); diff --git a/stage0/stdlib/Lean/Meta/Tactic/Split.c b/stage0/stdlib/Lean/Meta/Tactic/Split.c index 2cb284002fb0..2d8b3f6cdbfc 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Split.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Split.c @@ -193,7 +193,6 @@ lean_object* l_Lean_Meta_mkHEq(lean_object*, lean_object*, lean_object*, lean_ob lean_object* lean_array_to_list(lean_object*, lean_object*); lean_object* l_Lean_Meta_mkEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Split_0__Lean_Meta_Split_generalizeMatchDiscrs_withNewAltEqs_go___closed__2; -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_splitIfLocalDecl_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_hasLooseBVars(lean_object*); LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Meta_splitLocalDecl_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -288,6 +287,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Split_0__Lean_Meta_Split_s LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Split_0__Lean_Meta_Split_generalizeMatchDiscrs_mkNewTarget___lambda__4(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_Array_forInUnsafe_loop___at___private_Lean_Meta_Tactic_Split_0__Lean_Meta_Split_generalizeMatchDiscrs_mkNewTarget___spec__2___closed__6; +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_splitTarget_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Split_simpMatch___closed__3; static lean_object* l___private_Lean_Meta_Tactic_Split_0__Lean_Meta_Split_generalizeMatchDiscrs___lambda__6___closed__4; @@ -974,7 +974,7 @@ lean_inc(x_8); x_9 = lean_ctor_get(x_7, 1); lean_inc(x_9); lean_dec(x_7); -x_10 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_8, x_2, x_3, x_4, x_5, x_9); +x_10 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_8, x_2, x_3, x_4, x_5, x_9); x_11 = lean_ctor_get(x_10, 0); lean_inc(x_11); x_12 = lean_ctor_get(x_10, 1); @@ -1764,7 +1764,7 @@ lean_inc(x_10); x_11 = lean_ctor_get(x_9, 1); lean_inc(x_11); lean_dec(x_9); -x_12 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_10, x_4, x_5, x_6, x_7, x_11); +x_12 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_10, x_4, x_5, x_6, x_7, x_11); x_13 = lean_ctor_get(x_12, 0); lean_inc(x_13); x_14 = lean_ctor_get(x_12, 1); @@ -5089,7 +5089,7 @@ lean_closure_set(x_14, 3, x_3); lean_closure_set(x_14, 4, x_4); lean_closure_set(x_14, 5, x_5); lean_closure_set(x_14, 6, x_6); -x_15 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_8, x_9, x_10, x_11, x_12, x_13); +x_15 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_8, x_9, x_10, 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); @@ -11075,7 +11075,7 @@ lean_inc(x_9); x_10 = lean_ctor_get(x_8, 1); lean_inc(x_10); lean_dec(x_8); -x_11 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_9, x_3, x_4, x_5, x_6, x_10); +x_11 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_9, x_3, x_4, x_5, x_6, x_10); x_12 = lean_ctor_get(x_11, 0); lean_inc(x_12); x_13 = lean_ctor_get(x_11, 1); @@ -11343,7 +11343,7 @@ lean_inc(x_14); x_15 = lean_ctor_get(x_13, 1); lean_inc(x_15); lean_dec(x_13); -x_16 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_14, x_3, x_4, x_5, x_6, x_15); +x_16 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_14, x_3, x_4, x_5, x_6, x_15); x_17 = !lean_is_exclusive(x_16); if (x_17 == 0) { diff --git a/stage0/stdlib/Lean/Meta/Tactic/SplitIf.c b/stage0/stdlib/Lean/Meta/Tactic/SplitIf.c index 342c410847d2..9c731e7bcf5f 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/SplitIf.c +++ b/stage0/stdlib/Lean/Meta/Tactic/SplitIf.c @@ -105,7 +105,6 @@ lean_object* lean_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SplitIf_findIfToSplit_x3f___lambda__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_splitIfTarget_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_SplitIf_splitIfAt_x3f___closed__2; -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_SplitIf_discharge_x3f___lambda__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_splitIfLocalDecl_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SplitIf_splitIfAt_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -157,6 +156,7 @@ static lean_object* l_Lean_Meta_SplitIf_discharge_x3f___lambda__2___closed__7; static lean_object* l_Lean_Meta_SplitIf_discharge_x3f___closed__11; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); lean_object* l_Lean_mkApp3(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_SplitIf_discharge_x3f___closed__7; LEAN_EXPORT lean_object* l_Lean_LazyInitExtension_get___at_Lean_Meta_SplitIf_getSimpContext___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); @@ -3646,7 +3646,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_SplitIf_splitIfAt_x3f(lean_object* x_1, lea _start: { lean_object* x_9; uint8_t x_10; -x_9 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_2, x_4, x_5, x_6, x_7, x_8); +x_9 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_2, x_4, x_5, x_6, x_7, x_8); x_10 = !lean_is_exclusive(x_9); if (x_10 == 0) { diff --git a/stage0/stdlib/Lean/Meta/Tactic/Subst.c b/stage0/stdlib/Lean/Meta/Tactic/Subst.c index 5cb8c4861c18..0abdcb5514e5 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Subst.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Subst.c @@ -123,7 +123,6 @@ lean_object* l_Lean_Expr_replaceFVar(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_findSomeMAux___at_Lean_Meta_subst_findEq___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_to_list(lean_object*, lean_object*); lean_object* l_Lean_Meta_mkEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_substCore___lambda__19___closed__3; lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Subst___hyg_4294____closed__9; @@ -197,6 +196,7 @@ static lean_object* l_Lean_Meta_substCore___lambda__2___closed__6; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l_Lean_Meta_substCore___lambda__21___closed__24; LEAN_EXPORT lean_object* l_Lean_Meta_heqToEq___lambda__2(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_substCore___lambda__21___closed__8; static lean_object* l_Lean_Meta_substCore___lambda__14___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_substCore___lambda__9(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_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -2159,7 +2159,7 @@ goto block_111; block_111: { 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_33 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_32, x_16, x_17, x_18, x_19, x_31); +x_33 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_32, x_16, x_17, x_18, x_19, x_31); x_34 = lean_ctor_get(x_33, 0); lean_inc(x_34); x_35 = lean_ctor_get(x_33, 1); @@ -3890,7 +3890,7 @@ goto block_87; block_87: { lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_34 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_33, x_7, x_8, x_9, x_10, x_29); +x_34 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_33, x_7, x_8, x_9, x_10, x_29); x_35 = lean_ctor_get(x_34, 0); lean_inc(x_35); x_36 = lean_ctor_get(x_34, 1); @@ -3911,7 +3911,7 @@ goto block_86; block_86: { lean_object* x_38; -x_38 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_37, x_7, x_8, x_9, x_10, x_36); +x_38 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_37, x_7, x_8, x_9, x_10, x_36); if (lean_obj_tag(x_35) == 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; uint8_t x_45; @@ -6986,13 +6986,13 @@ if (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; uint8_t x_60; x_52 = lean_ctor_get(x_49, 0); x_53 = lean_ctor_get(x_49, 1); -x_54 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_52, x_7, x_8, x_9, x_10, x_50); +x_54 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_52, x_7, x_8, x_9, x_10, x_50); 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_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_53, x_7, x_8, x_9, x_10, x_56); +x_57 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_53, x_7, x_8, x_9, x_10, x_56); x_58 = lean_ctor_get(x_57, 0); lean_inc(x_58); x_59 = lean_ctor_get(x_57, 1); @@ -7204,13 +7204,13 @@ x_96 = lean_ctor_get(x_49, 1); lean_inc(x_96); lean_inc(x_95); lean_dec(x_49); -x_97 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_95, x_7, x_8, x_9, x_10, x_50); +x_97 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_95, x_7, x_8, x_9, x_10, x_50); 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); -x_100 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_96, x_7, x_8, x_9, x_10, x_99); +x_100 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_96, x_7, x_8, x_9, x_10, x_99); x_101 = lean_ctor_get(x_100, 0); lean_inc(x_101); x_102 = lean_ctor_get(x_100, 1); @@ -7443,13 +7443,13 @@ if (lean_is_exclusive(x_140)) { lean_dec_ref(x_140); x_144 = lean_box(0); } -x_145 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_142, x_7, x_8, x_9, x_10, x_141); +x_145 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_142, x_7, x_8, x_9, x_10, x_141); 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 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_143, x_7, x_8, x_9, x_10, x_147); +x_148 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_143, x_7, x_8, x_9, x_10, x_147); x_149 = lean_ctor_get(x_148, 0); lean_inc(x_149); x_150 = lean_ctor_get(x_148, 1); @@ -8096,13 +8096,13 @@ if (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; uint8_t x_60; x_52 = lean_ctor_get(x_49, 0); x_53 = lean_ctor_get(x_49, 1); -x_54 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_52, x_7, x_8, x_9, x_10, x_50); +x_54 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_52, x_7, x_8, x_9, x_10, x_50); 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_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_53, x_7, x_8, x_9, x_10, x_56); +x_57 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_53, x_7, x_8, x_9, x_10, x_56); x_58 = lean_ctor_get(x_57, 0); lean_inc(x_58); x_59 = lean_ctor_get(x_57, 1); @@ -8314,13 +8314,13 @@ x_96 = lean_ctor_get(x_49, 1); lean_inc(x_96); lean_inc(x_95); lean_dec(x_49); -x_97 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_95, x_7, x_8, x_9, x_10, x_50); +x_97 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_95, x_7, x_8, x_9, x_10, x_50); 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); -x_100 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_96, x_7, x_8, x_9, x_10, x_99); +x_100 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_96, x_7, x_8, x_9, x_10, x_99); x_101 = lean_ctor_get(x_100, 0); lean_inc(x_101); x_102 = lean_ctor_get(x_100, 1); @@ -8553,13 +8553,13 @@ if (lean_is_exclusive(x_140)) { lean_dec_ref(x_140); x_144 = lean_box(0); } -x_145 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_142, x_7, x_8, x_9, x_10, x_141); +x_145 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_142, x_7, x_8, x_9, x_10, x_141); 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 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_143, x_7, x_8, x_9, x_10, x_147); +x_148 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_143, x_7, x_8, x_9, x_10, x_147); x_149 = lean_ctor_get(x_148, 0); lean_inc(x_149); x_150 = lean_ctor_get(x_148, 1); diff --git a/stage0/stdlib/Lean/Meta/Tactic/Unfold.c b/stage0/stdlib/Lean/Meta/Tactic/Unfold.c index 19ec633ddf75..068349694582 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Unfold.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Unfold.c @@ -40,7 +40,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_unfold_pre___lambda__1(lean_object*, lean_o lean_object* l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_unfold_pre___lambda__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_reduceMatcher_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_FVarId_getType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_unfold_pre___closed__1; static lean_object* l_Lean_Meta_unfoldLocalDecl___lambda__1___closed__1; @@ -59,6 +58,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_unfoldLocalDecl___lambda__2(lean_object*, l static lean_object* l_Lean_Meta_unfold___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_unfoldLocalDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_applySimpResultToLocalDecl(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_indentExpr(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_unfold___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_getUnfoldEqnFor_x3f(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1104,7 +1104,7 @@ lean_inc(x_9); x_10 = lean_ctor_get(x_8, 1); lean_inc(x_10); lean_dec(x_8); -x_11 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_9, x_3, x_4, x_5, x_6, x_10); +x_11 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_9, x_3, x_4, x_5, x_6, x_10); x_12 = lean_ctor_get(x_11, 0); lean_inc(x_12); x_13 = lean_ctor_get(x_11, 1); @@ -1410,7 +1410,7 @@ x_11 = lean_ctor_get(x_9, 1); lean_inc(x_11); lean_dec(x_9); lean_inc(x_10); -x_12 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_10, x_4, x_5, x_6, x_7, x_11); +x_12 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_10, x_4, x_5, x_6, x_7, x_11); x_13 = lean_ctor_get(x_12, 0); lean_inc(x_13); x_14 = lean_ctor_get(x_12, 1); diff --git a/stage0/stdlib/Lean/Meta/Tactic/UnifyEq.c b/stage0/stdlib/Lean/Meta/Tactic/UnifyEq.c index afb31b4b891e..6895f5366c40 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/UnifyEq.c +++ b/stage0/stdlib/Lean/Meta/Tactic/UnifyEq.c @@ -38,7 +38,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_unifyEq_x3f_injection(lean_object*, lean_ob lean_object* l_Lean_FVarId_getDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); lean_object* l_Lean_Meta_mkEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_unifyEq_x3f_substEq___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); lean_object* l_Lean_MVarId_clear(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -55,6 +54,7 @@ uint8_t l_Lean_Expr_isConstructorApp(lean_object*, lean_object*); lean_object* l_Lean_LocalDecl_userName(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_unifyEq_x3f___lambda__1(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_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_indentExpr(lean_object*); uint8_t l_Lean_Expr_isHEq(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_UnifyEqResult_numNewEqs___default; @@ -1322,14 +1322,14 @@ lean_dec(x_25); x_27 = l_Lean_Expr_appArg_x21(x_14); lean_dec(x_14); lean_inc(x_26); -x_28 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_26, x_6, x_7, x_8, x_9, x_13); +x_28 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_26, x_6, x_7, x_8, x_9, 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); lean_dec(x_28); lean_inc(x_27); -x_31 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_27, x_6, x_7, x_8, x_9, x_30); +x_31 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_27, x_6, x_7, x_8, x_9, x_30); if (lean_obj_tag(x_29) == 1) { lean_object* x_32; diff --git a/stage0/stdlib/Lean/Meta/Tactic/Util.c b/stage0/stdlib/Lean/Meta/Tactic/Util.c index 851b1141e54d..c4abc7b6d641 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Util.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Util.c @@ -126,7 +126,6 @@ LEAN_EXPORT lean_object* l_Lean_HashSetImp_contains___at_Lean_MVarId_getNondepPr LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_MVarId_getNondepPropHyps___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_st_mk_ref(lean_object*, lean_object*); lean_object* lean_array_to_list(lean_object*, lean_object*); -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_appendTagSuffix(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); lean_object* l_Lean_Meta_synthInstance(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -187,6 +186,7 @@ LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_Meta_getPropH uint8_t lean_nat_dec_lt(lean_object*, lean_object*); lean_object* l_Lean_MVarId_isAssigned___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_MVarId_inferInstance___lambda__1___closed__3; +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_MVarId_admit___closed__1; static lean_object* l_Lean_Meta_throwTacticEx___rarg___closed__1; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); @@ -1199,7 +1199,7 @@ lean_inc(x_8); x_9 = lean_ctor_get(x_7, 1); lean_inc(x_9); lean_dec(x_7); -x_10 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_8, x_2, x_3, x_4, x_5, x_9); +x_10 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_8, x_2, x_3, x_4, x_5, x_9); x_11 = lean_ctor_get(x_10, 0); lean_inc(x_11); x_12 = lean_ctor_get(x_10, 1); @@ -3510,7 +3510,7 @@ 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; x_47 = l_Lean_LocalDecl_type(x_45); lean_inc(x_47); -x_48 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_47, x_6, x_7, x_8, x_9, x_10); +x_48 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_47, x_6, x_7, x_8, x_9, x_10); x_49 = lean_ctor_get(x_48, 0); lean_inc(x_49); x_50 = lean_ctor_get(x_48, 1); @@ -3568,7 +3568,7 @@ lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean x_66 = lean_ctor_get(x_63, 0); lean_inc(x_66); lean_dec(x_63); -x_67 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_66, x_6, x_7, x_8, x_9, x_62); +x_67 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_66, x_6, x_7, x_8, x_9, x_62); x_68 = lean_ctor_get(x_67, 0); lean_inc(x_68); x_69 = lean_ctor_get(x_67, 1); @@ -4112,7 +4112,7 @@ if (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; x_52 = l_Lean_LocalDecl_type(x_50); lean_inc(x_52); -x_53 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_52, x_6, x_7, x_8, x_9, x_10); +x_53 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_52, x_6, x_7, x_8, x_9, x_10); x_54 = lean_ctor_get(x_53, 0); lean_inc(x_54); x_55 = lean_ctor_get(x_53, 1); @@ -4170,7 +4170,7 @@ lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean x_71 = lean_ctor_get(x_68, 0); lean_inc(x_71); lean_dec(x_68); -x_72 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_71, x_6, x_7, x_8, x_9, x_67); +x_72 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_71, x_6, x_7, x_8, x_9, x_67); x_73 = lean_ctor_get(x_72, 0); lean_inc(x_73); x_74 = lean_ctor_get(x_72, 1); @@ -5447,7 +5447,7 @@ lean_inc(x_14); x_15 = lean_ctor_get(x_13, 1); lean_inc(x_15); lean_dec(x_13); -x_16 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_14, x_3, x_4, x_5, x_6, x_15); +x_16 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_14, x_3, x_4, x_5, x_6, x_15); x_17 = lean_ctor_get(x_16, 0); lean_inc(x_17); x_18 = lean_ctor_get(x_16, 1); diff --git a/stage0/stdlib/Lean/Meta/Transform.c b/stage0/stdlib/Lean/Meta/Transform.c index 71762505122d..c7d9919e87d1 100644 --- a/stage0/stdlib/Lean/Meta/Transform.c +++ b/stage0/stdlib/Lean/Meta/Transform.c @@ -776,7 +776,7 @@ static lean_object* _init_l_Lean_Core_transform_visit___rarg___lambda__2___close 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_Core_transform_visit___rarg___lambda__2___closed__1; x_2 = l_Lean_Core_transform_visit___rarg___lambda__2___closed__2; -x_3 = lean_unsigned_to_nat(1560u); +x_3 = lean_unsigned_to_nat(1564u); x_4 = lean_unsigned_to_nat(20u); x_5 = l_Lean_Core_transform_visit___rarg___lambda__2___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -914,7 +914,7 @@ static lean_object* _init_l_Lean_Core_transform_visit___rarg___lambda__4___close 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_Core_transform_visit___rarg___lambda__2___closed__1; x_2 = l_Lean_Core_transform_visit___rarg___lambda__4___closed__1; -x_3 = lean_unsigned_to_nat(1540u); +x_3 = lean_unsigned_to_nat(1544u); x_4 = lean_unsigned_to_nat(24u); x_5 = l_Lean_Core_transform_visit___rarg___lambda__4___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -1052,7 +1052,7 @@ static lean_object* _init_l_Lean_Core_transform_visit___rarg___lambda__6___close 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_Core_transform_visit___rarg___lambda__2___closed__1; x_2 = l_Lean_Core_transform_visit___rarg___lambda__6___closed__1; -x_3 = lean_unsigned_to_nat(1569u); +x_3 = lean_unsigned_to_nat(1573u); x_4 = lean_unsigned_to_nat(22u); x_5 = l_Lean_Core_transform_visit___rarg___lambda__6___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -1224,7 +1224,7 @@ static lean_object* _init_l_Lean_Core_transform_visit___rarg___lambda__9___close 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_Core_transform_visit___rarg___lambda__2___closed__1; x_2 = l_Lean_Core_transform_visit___rarg___lambda__9___closed__1; -x_3 = lean_unsigned_to_nat(1503u); +x_3 = lean_unsigned_to_nat(1507u); x_4 = lean_unsigned_to_nat(17u); x_5 = l_Lean_Core_transform_visit___rarg___lambda__9___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -1296,7 +1296,7 @@ static lean_object* _init_l_Lean_Core_transform_visit___rarg___lambda__10___clos 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_Core_transform_visit___rarg___lambda__2___closed__1; x_2 = l_Lean_Core_transform_visit___rarg___lambda__10___closed__1; -x_3 = lean_unsigned_to_nat(1514u); +x_3 = lean_unsigned_to_nat(1518u); x_4 = lean_unsigned_to_nat(18u); x_5 = l_Lean_Core_transform_visit___rarg___lambda__10___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); diff --git a/stage0/stdlib/Lean/Meta/UnificationHint.c b/stage0/stdlib/Lean/Meta/UnificationHint.c index e33ddf148137..b6f5ec0baa59 100644 --- a/stage0/stdlib/Lean/Meta/UnificationHint.c +++ b/stage0/stdlib/Lean/Meta/UnificationHint.c @@ -4110,7 +4110,7 @@ static lean_object* _init_l_Lean_Meta_DiscrTree_insertCore___at_Lean_Meta_Unific 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_Meta_DiscrTree_insertCore___at_Lean_Meta_UnificationHints_add___spec__1___closed__1; x_2 = l_Lean_Meta_DiscrTree_insertCore___at_Lean_Meta_UnificationHints_add___spec__1___closed__2; -x_3 = lean_unsigned_to_nat(366u); +x_3 = lean_unsigned_to_nat(396u); x_4 = lean_unsigned_to_nat(23u); x_5 = l_Lean_Meta_DiscrTree_insertCore___at_Lean_Meta_UnificationHints_add___spec__1___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); diff --git a/stage0/stdlib/Lean/Meta/WHNF.c b/stage0/stdlib/Lean/Meta/WHNF.c index 692b426d3f3c..7dc506dc91b8 100644 --- a/stage0/stdlib/Lean/Meta/WHNF.c +++ b/stage0/stdlib/Lean/Meta/WHNF.c @@ -14,8 +14,8 @@ extern "C" { #endif lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); -uint8_t l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_whnfHeadPred(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_whnfEasyCases___at_Lean_Meta_whnfImp___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Init_Util_0__outOfBounds___rarg(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_WHNF_0__Lean_Meta_isRecStuck_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -37,7 +37,6 @@ lean_object* lean_mk_empty_array_with_capacity(lean_object*); lean_object* l_Lean_mkAppN(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_reduceRecMatcher_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_WHNF_0__Lean_Meta_deltaBetaDefinition___at_Lean_Meta_whnfCore_go___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_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__2; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_WHNF_0__Lean_Meta_toCtorWhenK___spec__2(lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Meta_getStructuralRecArgPos_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); @@ -71,7 +70,6 @@ LEAN_EXPORT lean_object* lean_whnf(lean_object*, lean_object*, lean_object*, lea LEAN_EXPORT lean_object* l_Lean_Meta_reduceNat_x3f___lambda__1(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_WHNF_0__Lean_Meta_toCtorWhenStructure(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_unfoldProjInstWhenIntances_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_whnfEasyCases___at_Lean_Meta_whnfUntil___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_Meta_whnfUntil___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_canUnfoldAtMatcher___closed__5; @@ -88,6 +86,7 @@ lean_object* lean_array_push(lean_object*, lean_object*); lean_object* l_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_smartUnfoldingReduce_x3f_go___closed__1; LEAN_EXPORT lean_object* l_Lean_ofExcept___at_Lean_Meta_reduceBoolNativeUnsafe___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701_(lean_object*); LEAN_EXPORT lean_object* l_Lean_ofExcept___at_Lean_Meta_reduceBoolNativeUnsafe___spec__2(lean_object*); static lean_object* l_Lean_Meta_reduceNat_x3f___closed__16; lean_object* lean_mk_array(lean_object*, lean_object*); @@ -104,7 +103,6 @@ lean_object* l_Lean_Expr_proj___override(lean_object*, lean_object*, lean_object LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_reduceBoolNativeUnsafe___spec__3(lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getStuckMVar_x3f___lambda__10(lean_object*, 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_WHNF___hyg_10682____closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at_Lean_Meta_smartUnfoldingReduce_x3f_go___spec__2(lean_object*); lean_object* l_Lean_ConstantInfo_name(lean_object*); static lean_object* l_Lean_Meta_canUnfoldAtMatcher___closed__14; @@ -113,17 +111,16 @@ lean_object* lean_synth_pending(lean_object*, lean_object*, lean_object*, lean_o LEAN_EXPORT lean_object* l_Lean_Meta_getStuckMVar_x3f___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_canUnfoldAtMatcher___closed__28; lean_object* lean_environment_find(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__2; extern lean_object* l_Lean_matchPatternAttr; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_WHNF_0__Lean_Meta_toCtorWhenStructure___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_reduceNative_x3f___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Meta_WHNF_0__Lean_Meta_deltaBetaDefinition___at_Lean_Meta_unfoldDefinition_x3f___spec__3___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_Meta_reduceBoolNativeUnsafe___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__3; uint8_t l_Lean_ConstantInfo_hasValue(lean_object*); lean_object* l_List_find_x3f___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_WHNF_0__Lean_Meta_projectCore_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isAppOf(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__7; LEAN_EXPORT lean_object* l_Lean_Meta_getMatcherInfo_x3f___at_Lean_Meta_reduceMatcher_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_reduceMatcher_x3f___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_Meta_whnfEasyCases___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -164,7 +161,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_project_x3f(lean_object*, lean_object*, lea extern lean_object* l_Lean_projectionFnInfoExt; uint8_t l_Lean_Expr_hasMVar(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_getStuckMVar_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_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_whnfHeadPred___at_Lean_Meta_whnfUntil___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isLambda(lean_object*); size_t lean_ptr_addr(lean_object*); @@ -204,6 +200,7 @@ LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_whnfCore_go___spec__1(lean_objec LEAN_EXPORT lean_object* l_Lean_Meta_getStuckMVar_x3f___lambda__21(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_expr_eqv(lean_object*, lean_object*); lean_object* l_Lean_Meta_Match_Extension_getMatcherInfo_x3f(lean_object*, lean_object*); +uint8_t l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_12510_(uint8_t, uint8_t); static lean_object* l_Lean_Meta_markSmartUnfoldingMatchAlt___closed__1; static lean_object* l_Lean_Meta_canUnfoldAtMatcher___closed__16; LEAN_EXPORT lean_object* l_Lean_Meta_smartUnfoldingMatch_x3f___boxed(lean_object*); @@ -216,14 +213,14 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_WHNF_0__Lean_Meta_cached_x3f(uint uint8_t l_Lean_TagAttribute_hasTag(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_reduceNat_x3f___lambda__1___boxed(lean_object*); static lean_object* l_Lean_Meta_reduceNat_x3f___closed__8; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_whnfCore(lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__5; lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getStuckMVar_x3f___lambda__16(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ofExcept___at_Lean_Meta_reduceBoolNativeUnsafe___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_markSmartUnfoldingMatch(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_WHNF_0__Lean_Meta_reduceRec___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*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__10; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__12; LEAN_EXPORT lean_object* l_Lean_Meta_project_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getStuckMVar_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_canUnfoldAtMatcher___closed__39; @@ -253,11 +250,11 @@ LEAN_EXPORT lean_object* l_Lean_Meta_withNatValue(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_reduceMatcher_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_RecursorVal_getMajorIdx(lean_object*); static lean_object* l_Lean_getConstInfoCtor___at___private_Lean_Meta_WHNF_0__Lean_Meta_toCtorWhenStructure___spec__1___closed__2; -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_smartUnfoldingMatchAlt_x3f(lean_object*); lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); lean_object* l_Nat_beq___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Meta_whnfEasyCases___at_Lean_Meta_whnfCore_go___spec__3___closed__1; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__6; lean_object* l_Lean_getStructureInfo_x3f(lean_object*, lean_object*); static lean_object* l_Lean_Meta_reduceBinNatOp___closed__1; static lean_object* l_Lean_Meta_canUnfoldAtMatcher___closed__37; @@ -295,15 +292,17 @@ lean_object* l_Lean_Expr_isConstructorApp_x3f(lean_object*, lean_object*); lean_object* l_Lean_MapDeclarationExtension_find_x3f___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_reduceNative_x3f___closed__11; LEAN_EXPORT lean_object* l_Lean_Meta_mkSmartUnfoldingNameFor(lean_object*); +uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Meta_reduceNative_x3f___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_withNatValue___rarg(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*); static lean_object* l_Lean_Meta_canUnfoldAtMatcher___closed__17; lean_object* l_Nat_mul___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_WHNF_0__Lean_Meta_cached_x3f___closed__1; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__9; static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_36____closed__8; +lean_object* l_Lean_isIrreducible___at___private_Lean_Meta_GetConst_0__Lean_Meta_canUnfoldDefault___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at___private_Lean_Meta_WHNF_0__Lean_Meta_cached_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__7; LEAN_EXPORT lean_object* l_Lean_Meta_mkProjFn___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getStuckMVar_x3f___lambda__12(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_instantiateLevelParams(lean_object*, lean_object*, lean_object*); @@ -321,7 +320,6 @@ extern lean_object* l_Lean_instInhabitedProjectionFunctionInfo; lean_object* l_Lean_Meta_instInhabitedMetaM___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_whnfCore___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*); -lean_object* l_Lean_Option_register___at_Std_Format_initFn____x40_Lean_Data_Format___hyg_95____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_WHNF_0__Lean_Meta_reduceRec___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_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_WHNF_0__Lean_Meta_toCtorWhenStructure___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_EXPORT lean_object* l_Lean_Meta_reduceNatNativeUnsafe(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -336,12 +334,12 @@ lean_object* l_Lean_Meta_getConfig(lean_object*, lean_object*, lean_object*, lea uint8_t l_Lean_Expr_isConstructorApp(lean_object*, lean_object*); lean_object* l_Lean_Expr_toCtorIfLit(lean_object*); static lean_object* l_Lean_Meta_whnfEasyCases___closed__2; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__11; LEAN_EXPORT lean_object* l_Lean_Meta_whnfEasyCases___at_Lean_Meta_whnfCore_go___spec__3(uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_reduceBinNatOp___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_reduceRecMatcher_x3f___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_reduceNat_x3f___closed__18; LEAN_EXPORT lean_object* l___private_Lean_Meta_WHNF_0__Lean_Meta_whnfDelayedAssigned_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__11; static lean_object* l_Lean_Meta_smartUnfoldingSuffix___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_WHNF_0__Lean_Meta_getRecRuleFor(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mkProjFn(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -378,12 +376,13 @@ uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_reduceMatcher_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_reduceNat_x3f___closed__11; LEAN_EXPORT lean_object* l_Lean_Meta_whnfEasyCases___at_Lean_Meta_whnfUntil___spec__2___lambda__1(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_WHNF___hyg_10682____closed__6; static lean_object* l_Lean_Meta_whnfEasyCases___at_Lean_Meta_whnfCore_go___spec__3___closed__2; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__8; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l_Lean_Meta_reduceNat_x3f___closed__13; static lean_object* l_Lean_Meta_reduceNat_x3f___closed__2; static lean_object* l_Lean_Meta_smartUnfoldingReduce_x3f_go___lambda__2___closed__1; +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_36____closed__2; lean_object* l_Lean_mkRawNatLit(lean_object*); static lean_object* l_Lean_Meta_reduceBinNatOp___closed__6; @@ -393,6 +392,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_WHNF_0__Lean_Meta_deltaBetaDefini lean_object* l_Lean_getExprMVarAssignment_x3f___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_reduceRecMatcher_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__9; LEAN_EXPORT lean_object* l_Lean_Meta_reduceBinNatOp___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_WHNF_0__Lean_Meta_whnfMatcher___lambda__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_indentExpr(lean_object*); @@ -400,7 +400,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_whnfEasyCases___at_Lean_Meta_whnfCore_go___ lean_object* l_Array_shrink___rarg(lean_object*, lean_object*); lean_object* l_Lean_Core_checkMaxHeartbeats(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_WHNF_0__Lean_Meta_mkNullaryCtor___boxed(lean_object*, 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_12336_(uint8_t, uint8_t); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__13; LEAN_EXPORT lean_object* l_Lean_Meta_unfoldDefinition_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isConstOf(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_reduceBoolNativeUnsafe___spec__3___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -418,12 +418,12 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_WHNF_0__Lean LEAN_EXPORT lean_object* l_Lean_Meta_reduceBinNatOp___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_canUnfoldAtMatcher___closed__30; lean_object* lean_nat_sub(lean_object*, lean_object*); +lean_object* l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_isAuxDef___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_getAppFn(lean_object*); static lean_object* l_Lean_Meta_unfoldDefinition___closed__1; static lean_object* l_Lean_Meta_canUnfoldAtMatcher___closed__38; static lean_object* l_Lean_Meta_whnfEasyCases___closed__1; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__12; static lean_object* l_Lean_Meta_reduceNat_x3f___closed__15; LEAN_EXPORT lean_object* l_Lean_Meta_reduceBinNatOp___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_Meta_WHNF_0__Lean_Meta_whnfMatcher(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -435,6 +435,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_WHNF_0__Lean_Meta_getFirstCtor___ static lean_object* l_Lean_Meta_canUnfoldAtMatcher___closed__13; lean_object* l_Array_ofSubarray___rarg(lean_object*); lean_object* l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__1; lean_object* l_Lean_Environment_getProjectionStructureName_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_getStuckMVar_x3f___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_Array_forInUnsafe_loop___at_Lean_Meta_getStuckMVar_x3f___spec__6(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -486,6 +487,7 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_getStuckMVar_x3 lean_object* lean_array_get_size(lean_object*); lean_object* l_Lean_PersistentHashMap_insert___at___private_Lean_Meta_InferType_0__Lean_Meta_checkInferTypeCache___spec__4(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_whnfEasyCases___at_Lean_Meta_whnfCore_go___spec__3___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_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__5; static lean_object* l___private_Lean_Meta_WHNF_0__Lean_Meta_cached_x3f___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_getStuckMVar_x3f___lambda__15(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_getStuckMVar_x3f___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -517,9 +519,9 @@ LEAN_EXPORT lean_object* l_Lean_Meta_getStuckMVar_x3f___lambda__17___boxed(lean_ static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_36____closed__6; LEAN_EXPORT lean_object* l___private_Lean_Meta_WHNF_0__Lean_Meta_deltaBetaDefinition___at_Lean_Meta_unfoldDefinition_x3f___spec__3___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getConstInfoCtor___at___private_Lean_Meta_WHNF_0__Lean_Meta_toCtorWhenStructure___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__13; LEAN_EXPORT lean_object* l_Lean_Meta_getStuckMVar_x3f___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_reduceProj_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__10; lean_object* l_Lean_Expr_mvarId_x21(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_WHNF_0__Lean_Meta_cache(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Meta_ParamInfo_isExplicit(lean_object*); @@ -538,7 +540,6 @@ static lean_object* l_Lean_Meta_whnfEasyCases___closed__5; LEAN_EXPORT lean_object* l___private_Lean_Meta_WHNF_0__Lean_Meta_useWHNFCache(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_canUnfoldAtMatcher(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_smartUnfoldingMatchAlt_x3f___boxed(lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__8; uint8_t l_Lean_Expr_hasExprMVar(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_getStuckMVar_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* l_Array_get_x3f___rarg(lean_object*, lean_object*); @@ -677,7 +678,7 @@ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_36____closed__2; x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_36____closed__5; x_4 = l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_36____closed__8; -x_5 = l_Lean_Option_register___at_Std_Format_initFn____x40_Lean_Data_Format___hyg_95____spec__1(x_2, x_3, x_4, x_1); +x_5 = l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__1(x_2, x_3, x_4, x_1); return x_5; } } @@ -1611,7 +1612,7 @@ lean_inc(x_12); x_13 = lean_ctor_get(x_11, 1); lean_inc(x_13); lean_dec(x_11); -x_14 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_12, x_3, x_4, x_5, x_6, x_13); +x_14 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_12, x_3, x_4, x_5, x_6, x_13); x_15 = !lean_is_exclusive(x_14); if (x_15 == 0) { @@ -4086,7 +4087,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_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_20, x_4, x_5, x_6, x_7, x_21); +x_22 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_20, x_4, x_5, x_6, x_7, x_21); x_23 = !lean_is_exclusive(x_22); if (x_23 == 0) { @@ -4283,7 +4284,7 @@ lean_inc(x_55); x_56 = lean_ctor_get(x_54, 1); lean_inc(x_56); lean_dec(x_54); -x_57 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_55, x_4, x_5, x_6, x_7, x_56); +x_57 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_55, x_4, x_5, x_6, x_7, x_56); x_58 = lean_ctor_get(x_57, 0); lean_inc(x_58); x_59 = lean_ctor_get(x_57, 1); @@ -11996,7 +11997,7 @@ switch (lean_obj_tag(x_1)) { case 2: { lean_object* x_7; lean_object* x_8; -x_7 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_1, x_2, x_3, x_4, x_5, x_6); +x_7 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_1, x_2, x_3, x_4, x_5, x_6); x_8 = lean_ctor_get(x_7, 0); lean_inc(x_8); if (lean_obj_tag(x_8) == 2) @@ -12060,7 +12061,7 @@ case 2: { lean_object* x_21; uint8_t x_22; lean_dec(x_20); -x_21 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_1, x_2, x_3, x_4, x_5, x_6); +x_21 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_1, x_2, x_3, x_4, x_5, x_6); x_22 = !lean_is_exclusive(x_21); if (x_22 == 0) { @@ -14201,2370 +14202,1259 @@ uint8_t x_6; lean_object* x_7; x_6 = lean_ctor_get_uint8(x_1, 5); x_7 = lean_box(x_6); switch (lean_obj_tag(x_7)) { -case 2: -{ -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 = l_Lean_ConstantInfo_name(x_2); -lean_inc(x_8); -x_9 = l_Lean_isReducible___at___private_Lean_Meta_GetConst_0__Lean_Meta_canUnfoldDefault___spec__3(x_8, x_3, x_4, x_5); -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_st_ref_get(x_4, x_11); -x_13 = lean_unbox(x_10); -lean_dec(x_10); -if (x_13 == 0) -{ -uint8_t x_14; -x_14 = !lean_is_exclusive(x_12); -if (x_14 == 0) -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_15 = lean_ctor_get(x_12, 0); -x_16 = lean_ctor_get(x_12, 1); -x_17 = lean_ctor_get(x_15, 0); -lean_inc(x_17); -lean_dec(x_15); -lean_inc(x_8); -x_18 = lean_is_instance(x_17, x_8); -if (x_18 == 0) -{ -lean_object* x_19; uint8_t x_20; -lean_free_object(x_12); -x_19 = lean_st_ref_get(x_4, x_16); -x_20 = !lean_is_exclusive(x_19); -if (x_20 == 0) -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; -x_21 = lean_ctor_get(x_19, 0); -x_22 = lean_ctor_get(x_21, 0); -lean_inc(x_22); -lean_dec(x_21); -x_23 = l_Lean_Meta_canUnfoldAtMatcher___closed__1; -lean_inc(x_8); -x_24 = l_Lean_TagAttribute_hasTag(x_23, x_22, x_8); -if (x_24 == 0) -{ -lean_object* x_25; uint8_t x_26; -x_25 = l_Lean_Meta_canUnfoldAtMatcher___closed__3; -x_26 = lean_name_eq(x_8, x_25); -if (x_26 == 0) -{ -lean_object* x_27; uint8_t x_28; -x_27 = l_Lean_Meta_canUnfoldAtMatcher___closed__5; -x_28 = lean_name_eq(x_8, x_27); -if (x_28 == 0) -{ -lean_object* x_29; uint8_t x_30; -x_29 = l_Lean_Meta_canUnfoldAtMatcher___closed__7; -x_30 = lean_name_eq(x_8, x_29); -if (x_30 == 0) -{ -lean_object* x_31; uint8_t x_32; -x_31 = l_Lean_Meta_canUnfoldAtMatcher___closed__9; -x_32 = lean_name_eq(x_8, x_31); -if (x_32 == 0) -{ -lean_object* x_33; uint8_t x_34; -x_33 = l_Lean_Meta_canUnfoldAtMatcher___closed__12; -x_34 = lean_name_eq(x_8, x_33); -if (x_34 == 0) -{ -lean_object* x_35; uint8_t x_36; -x_35 = l_Lean_Meta_canUnfoldAtMatcher___closed__14; -x_36 = lean_name_eq(x_8, x_35); -if (x_36 == 0) -{ -lean_object* x_37; uint8_t x_38; -x_37 = l_Lean_Meta_canUnfoldAtMatcher___closed__16; -x_38 = lean_name_eq(x_8, x_37); -if (x_38 == 0) -{ -lean_object* x_39; uint8_t x_40; -x_39 = l_Lean_Meta_canUnfoldAtMatcher___closed__19; -x_40 = lean_name_eq(x_8, x_39); -if (x_40 == 0) -{ -lean_object* x_41; uint8_t x_42; -x_41 = l_Lean_Meta_canUnfoldAtMatcher___closed__21; -x_42 = lean_name_eq(x_8, x_41); -if (x_42 == 0) -{ -lean_object* x_43; uint8_t x_44; -x_43 = l_Lean_Meta_canUnfoldAtMatcher___closed__23; -x_44 = lean_name_eq(x_8, x_43); -if (x_44 == 0) -{ -lean_object* x_45; uint8_t x_46; -x_45 = l_Lean_Meta_canUnfoldAtMatcher___closed__24; -x_46 = lean_name_eq(x_8, x_45); -if (x_46 == 0) -{ -lean_object* x_47; uint8_t x_48; -x_47 = l_Lean_Meta_canUnfoldAtMatcher___closed__26; -x_48 = lean_name_eq(x_8, x_47); -if (x_48 == 0) -{ -lean_object* x_49; uint8_t x_50; -x_49 = l_Lean_Meta_canUnfoldAtMatcher___closed__27; -x_50 = lean_name_eq(x_8, x_49); -if (x_50 == 0) -{ -lean_object* x_51; uint8_t x_52; -x_51 = l_Lean_Meta_canUnfoldAtMatcher___closed__29; -x_52 = lean_name_eq(x_8, x_51); -if (x_52 == 0) -{ -lean_object* x_53; uint8_t x_54; -x_53 = l_Lean_Meta_canUnfoldAtMatcher___closed__30; -x_54 = lean_name_eq(x_8, x_53); -if (x_54 == 0) -{ -lean_object* x_55; uint8_t x_56; -x_55 = l_Lean_Meta_canUnfoldAtMatcher___closed__32; -x_56 = lean_name_eq(x_8, x_55); -if (x_56 == 0) -{ -lean_object* x_57; uint8_t x_58; -x_57 = l_Lean_Meta_canUnfoldAtMatcher___closed__33; -x_58 = lean_name_eq(x_8, x_57); -if (x_58 == 0) -{ -lean_object* x_59; uint8_t x_60; -x_59 = l_Lean_Meta_canUnfoldAtMatcher___closed__36; -x_60 = lean_name_eq(x_8, x_59); -if (x_60 == 0) -{ -lean_object* x_61; uint8_t x_62; lean_object* x_63; -x_61 = l_Lean_Meta_canUnfoldAtMatcher___closed__39; -x_62 = lean_name_eq(x_8, x_61); -lean_dec(x_8); -x_63 = lean_box(x_62); -lean_ctor_set(x_19, 0, x_63); -return x_19; -} -else -{ -uint8_t x_64; lean_object* x_65; -lean_dec(x_8); -x_64 = 1; -x_65 = lean_box(x_64); -lean_ctor_set(x_19, 0, x_65); -return x_19; -} -} -else -{ -uint8_t x_66; lean_object* x_67; -lean_dec(x_8); -x_66 = 1; -x_67 = lean_box(x_66); -lean_ctor_set(x_19, 0, x_67); -return x_19; -} -} -else -{ -uint8_t x_68; lean_object* x_69; -lean_dec(x_8); -x_68 = 1; -x_69 = lean_box(x_68); -lean_ctor_set(x_19, 0, x_69); -return x_19; -} -} -else -{ -uint8_t x_70; lean_object* x_71; -lean_dec(x_8); -x_70 = 1; -x_71 = lean_box(x_70); -lean_ctor_set(x_19, 0, x_71); -return x_19; -} -} -else -{ -uint8_t x_72; lean_object* x_73; -lean_dec(x_8); -x_72 = 1; -x_73 = lean_box(x_72); -lean_ctor_set(x_19, 0, x_73); -return x_19; -} -} -else -{ -uint8_t x_74; lean_object* x_75; -lean_dec(x_8); -x_74 = 1; -x_75 = lean_box(x_74); -lean_ctor_set(x_19, 0, x_75); -return x_19; -} -} -else -{ -uint8_t x_76; lean_object* x_77; -lean_dec(x_8); -x_76 = 1; -x_77 = lean_box(x_76); -lean_ctor_set(x_19, 0, x_77); -return x_19; -} -} -else -{ -uint8_t x_78; lean_object* x_79; -lean_dec(x_8); -x_78 = 1; -x_79 = lean_box(x_78); -lean_ctor_set(x_19, 0, x_79); -return x_19; -} -} -else -{ -uint8_t x_80; lean_object* x_81; -lean_dec(x_8); -x_80 = 1; -x_81 = lean_box(x_80); -lean_ctor_set(x_19, 0, x_81); -return x_19; -} -} -else -{ -uint8_t x_82; lean_object* x_83; -lean_dec(x_8); -x_82 = 1; -x_83 = lean_box(x_82); -lean_ctor_set(x_19, 0, x_83); -return x_19; -} -} -else +case 0: { -uint8_t x_84; lean_object* x_85; -lean_dec(x_8); -x_84 = 1; -x_85 = lean_box(x_84); -lean_ctor_set(x_19, 0, x_85); -return x_19; -} +uint8_t x_8; lean_object* x_9; lean_object* x_10; +x_8 = 1; +x_9 = lean_box(x_8); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_5); +return x_10; } -else +case 1: { -uint8_t x_86; lean_object* x_87; -lean_dec(x_8); -x_86 = 1; -x_87 = lean_box(x_86); -lean_ctor_set(x_19, 0, x_87); -return x_19; -} -} -else +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_11 = l_Lean_ConstantInfo_name(x_2); +x_12 = l_Lean_isIrreducible___at___private_Lean_Meta_GetConst_0__Lean_Meta_canUnfoldDefault___spec__1(x_11, x_3, x_4, x_5); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_unbox(x_13); +lean_dec(x_13); +if (x_14 == 0) { -uint8_t x_88; lean_object* x_89; -lean_dec(x_8); -x_88 = 1; -x_89 = lean_box(x_88); -lean_ctor_set(x_19, 0, x_89); -return x_19; -} -} -else +uint8_t x_15; +x_15 = !lean_is_exclusive(x_12); +if (x_15 == 0) { -uint8_t x_90; lean_object* x_91; -lean_dec(x_8); -x_90 = 1; -x_91 = lean_box(x_90); -lean_ctor_set(x_19, 0, x_91); -return x_19; -} +lean_object* x_16; uint8_t x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_12, 0); +lean_dec(x_16); +x_17 = 1; +x_18 = lean_box(x_17); +lean_ctor_set(x_12, 0, x_18); +return x_12; } else { -uint8_t x_92; lean_object* x_93; -lean_dec(x_8); -x_92 = 1; -x_93 = lean_box(x_92); -lean_ctor_set(x_19, 0, x_93); -return x_19; +lean_object* x_19; uint8_t x_20; lean_object* x_21; lean_object* x_22; +x_19 = lean_ctor_get(x_12, 1); +lean_inc(x_19); +lean_dec(x_12); +x_20 = 1; +x_21 = lean_box(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_19); +return x_22; } } else { -uint8_t x_94; lean_object* x_95; -lean_dec(x_8); -x_94 = 1; -x_95 = lean_box(x_94); -lean_ctor_set(x_19, 0, x_95); -return x_19; -} -} -else +uint8_t x_23; +x_23 = !lean_is_exclusive(x_12); +if (x_23 == 0) { -uint8_t x_96; lean_object* x_97; -lean_dec(x_8); -x_96 = 1; -x_97 = lean_box(x_96); -lean_ctor_set(x_19, 0, x_97); -return x_19; -} +lean_object* x_24; uint8_t x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_12, 0); +lean_dec(x_24); +x_25 = 0; +x_26 = lean_box(x_25); +lean_ctor_set(x_12, 0, x_26); +return x_12; } else { -uint8_t x_98; lean_object* x_99; -lean_dec(x_8); -x_98 = 1; -x_99 = lean_box(x_98); -lean_ctor_set(x_19, 0, x_99); -return x_19; -} +lean_object* x_27; uint8_t x_28; lean_object* x_29; lean_object* x_30; +x_27 = lean_ctor_get(x_12, 1); +lean_inc(x_27); +lean_dec(x_12); +x_28 = 0; +x_29 = lean_box(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_27); +return x_30; } -else -{ -uint8_t x_100; lean_object* x_101; -lean_dec(x_8); -x_100 = 1; -x_101 = lean_box(x_100); -lean_ctor_set(x_19, 0, x_101); -return x_19; } } -else -{ -lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; uint8_t x_106; -x_102 = lean_ctor_get(x_19, 0); -x_103 = lean_ctor_get(x_19, 1); -lean_inc(x_103); -lean_inc(x_102); -lean_dec(x_19); -x_104 = lean_ctor_get(x_102, 0); -lean_inc(x_104); -lean_dec(x_102); -x_105 = l_Lean_Meta_canUnfoldAtMatcher___closed__1; -lean_inc(x_8); -x_106 = l_Lean_TagAttribute_hasTag(x_105, x_104, x_8); -if (x_106 == 0) -{ -lean_object* x_107; uint8_t x_108; -x_107 = l_Lean_Meta_canUnfoldAtMatcher___closed__3; -x_108 = lean_name_eq(x_8, x_107); -if (x_108 == 0) -{ -lean_object* x_109; uint8_t x_110; -x_109 = l_Lean_Meta_canUnfoldAtMatcher___closed__5; -x_110 = lean_name_eq(x_8, x_109); -if (x_110 == 0) -{ -lean_object* x_111; uint8_t x_112; -x_111 = l_Lean_Meta_canUnfoldAtMatcher___closed__7; -x_112 = lean_name_eq(x_8, x_111); -if (x_112 == 0) -{ -lean_object* x_113; uint8_t x_114; -x_113 = l_Lean_Meta_canUnfoldAtMatcher___closed__9; -x_114 = lean_name_eq(x_8, x_113); -if (x_114 == 0) -{ -lean_object* x_115; uint8_t x_116; -x_115 = l_Lean_Meta_canUnfoldAtMatcher___closed__12; -x_116 = lean_name_eq(x_8, x_115); -if (x_116 == 0) -{ -lean_object* x_117; uint8_t x_118; -x_117 = l_Lean_Meta_canUnfoldAtMatcher___closed__14; -x_118 = lean_name_eq(x_8, x_117); -if (x_118 == 0) -{ -lean_object* x_119; uint8_t x_120; -x_119 = l_Lean_Meta_canUnfoldAtMatcher___closed__16; -x_120 = lean_name_eq(x_8, x_119); -if (x_120 == 0) -{ -lean_object* x_121; uint8_t x_122; -x_121 = l_Lean_Meta_canUnfoldAtMatcher___closed__19; -x_122 = lean_name_eq(x_8, x_121); -if (x_122 == 0) +default: { -lean_object* x_123; uint8_t x_124; -x_123 = l_Lean_Meta_canUnfoldAtMatcher___closed__21; -x_124 = lean_name_eq(x_8, x_123); -if (x_124 == 0) +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_dec(x_7); +x_31 = l_Lean_ConstantInfo_name(x_2); +lean_inc(x_31); +x_32 = l_Lean_isReducible___at___private_Lean_Meta_GetConst_0__Lean_Meta_canUnfoldDefault___spec__3(x_31, x_3, x_4, x_5); +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_st_ref_get(x_4, x_34); +x_36 = lean_unbox(x_33); +lean_dec(x_33); +if (x_36 == 0) { -lean_object* x_125; uint8_t x_126; -x_125 = l_Lean_Meta_canUnfoldAtMatcher___closed__23; -x_126 = lean_name_eq(x_8, x_125); -if (x_126 == 0) +uint8_t x_37; +x_37 = !lean_is_exclusive(x_35); +if (x_37 == 0) { -lean_object* x_127; uint8_t x_128; -x_127 = l_Lean_Meta_canUnfoldAtMatcher___closed__24; -x_128 = lean_name_eq(x_8, x_127); -if (x_128 == 0) +lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_38 = lean_ctor_get(x_35, 0); +x_39 = lean_ctor_get(x_35, 1); +x_40 = lean_ctor_get(x_38, 0); +lean_inc(x_40); +lean_dec(x_38); +lean_inc(x_31); +x_41 = lean_is_instance(x_40, x_31); +if (x_41 == 0) { -lean_object* x_129; uint8_t x_130; -x_129 = l_Lean_Meta_canUnfoldAtMatcher___closed__26; -x_130 = lean_name_eq(x_8, x_129); -if (x_130 == 0) +lean_object* x_42; uint8_t x_43; +lean_free_object(x_35); +x_42 = lean_st_ref_get(x_4, x_39); +x_43 = !lean_is_exclusive(x_42); +if (x_43 == 0) { -lean_object* x_131; uint8_t x_132; -x_131 = l_Lean_Meta_canUnfoldAtMatcher___closed__27; -x_132 = lean_name_eq(x_8, x_131); -if (x_132 == 0) +lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; +x_44 = lean_ctor_get(x_42, 0); +x_45 = lean_ctor_get(x_44, 0); +lean_inc(x_45); +lean_dec(x_44); +x_46 = l_Lean_Meta_canUnfoldAtMatcher___closed__1; +lean_inc(x_31); +x_47 = l_Lean_TagAttribute_hasTag(x_46, x_45, x_31); +if (x_47 == 0) { -lean_object* x_133; uint8_t x_134; -x_133 = l_Lean_Meta_canUnfoldAtMatcher___closed__29; -x_134 = lean_name_eq(x_8, x_133); -if (x_134 == 0) +lean_object* x_48; uint8_t x_49; +x_48 = l_Lean_Meta_canUnfoldAtMatcher___closed__3; +x_49 = lean_name_eq(x_31, x_48); +if (x_49 == 0) { -lean_object* x_135; uint8_t x_136; -x_135 = l_Lean_Meta_canUnfoldAtMatcher___closed__30; -x_136 = lean_name_eq(x_8, x_135); -if (x_136 == 0) +lean_object* x_50; uint8_t x_51; +x_50 = l_Lean_Meta_canUnfoldAtMatcher___closed__5; +x_51 = lean_name_eq(x_31, x_50); +if (x_51 == 0) { -lean_object* x_137; uint8_t x_138; -x_137 = l_Lean_Meta_canUnfoldAtMatcher___closed__32; -x_138 = lean_name_eq(x_8, x_137); -if (x_138 == 0) +lean_object* x_52; uint8_t x_53; +x_52 = l_Lean_Meta_canUnfoldAtMatcher___closed__7; +x_53 = lean_name_eq(x_31, x_52); +if (x_53 == 0) { -lean_object* x_139; uint8_t x_140; -x_139 = l_Lean_Meta_canUnfoldAtMatcher___closed__33; -x_140 = lean_name_eq(x_8, x_139); -if (x_140 == 0) +lean_object* x_54; uint8_t x_55; +x_54 = l_Lean_Meta_canUnfoldAtMatcher___closed__9; +x_55 = lean_name_eq(x_31, x_54); +if (x_55 == 0) { -lean_object* x_141; uint8_t x_142; -x_141 = l_Lean_Meta_canUnfoldAtMatcher___closed__36; -x_142 = lean_name_eq(x_8, x_141); -if (x_142 == 0) +lean_object* x_56; uint8_t x_57; +x_56 = l_Lean_Meta_canUnfoldAtMatcher___closed__12; +x_57 = lean_name_eq(x_31, x_56); +if (x_57 == 0) { -lean_object* x_143; uint8_t x_144; lean_object* x_145; lean_object* x_146; -x_143 = l_Lean_Meta_canUnfoldAtMatcher___closed__39; -x_144 = lean_name_eq(x_8, x_143); -lean_dec(x_8); -x_145 = lean_box(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_103); -return x_146; -} -else +lean_object* x_58; uint8_t x_59; +x_58 = l_Lean_Meta_canUnfoldAtMatcher___closed__14; +x_59 = lean_name_eq(x_31, x_58); +if (x_59 == 0) { -uint8_t x_147; lean_object* x_148; lean_object* x_149; -lean_dec(x_8); -x_147 = 1; -x_148 = lean_box(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_103); -return x_149; -} -} -else +lean_object* x_60; uint8_t x_61; +x_60 = l_Lean_Meta_canUnfoldAtMatcher___closed__16; +x_61 = lean_name_eq(x_31, x_60); +if (x_61 == 0) { -uint8_t x_150; lean_object* x_151; lean_object* x_152; -lean_dec(x_8); -x_150 = 1; -x_151 = lean_box(x_150); -x_152 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_152, 0, x_151); -lean_ctor_set(x_152, 1, x_103); -return x_152; -} -} -else +lean_object* x_62; uint8_t x_63; +x_62 = l_Lean_Meta_canUnfoldAtMatcher___closed__19; +x_63 = lean_name_eq(x_31, x_62); +if (x_63 == 0) { -uint8_t x_153; lean_object* x_154; lean_object* x_155; -lean_dec(x_8); -x_153 = 1; -x_154 = lean_box(x_153); -x_155 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_155, 0, x_154); -lean_ctor_set(x_155, 1, x_103); -return x_155; -} -} -else +lean_object* x_64; uint8_t x_65; +x_64 = l_Lean_Meta_canUnfoldAtMatcher___closed__21; +x_65 = lean_name_eq(x_31, x_64); +if (x_65 == 0) { -uint8_t x_156; lean_object* x_157; lean_object* x_158; -lean_dec(x_8); -x_156 = 1; -x_157 = lean_box(x_156); -x_158 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_158, 0, x_157); -lean_ctor_set(x_158, 1, x_103); -return x_158; -} -} -else +lean_object* x_66; uint8_t x_67; +x_66 = l_Lean_Meta_canUnfoldAtMatcher___closed__23; +x_67 = lean_name_eq(x_31, x_66); +if (x_67 == 0) { -uint8_t x_159; lean_object* x_160; lean_object* x_161; -lean_dec(x_8); -x_159 = 1; -x_160 = lean_box(x_159); -x_161 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_161, 0, x_160); -lean_ctor_set(x_161, 1, x_103); -return x_161; -} -} -else +lean_object* x_68; uint8_t x_69; +x_68 = l_Lean_Meta_canUnfoldAtMatcher___closed__24; +x_69 = lean_name_eq(x_31, x_68); +if (x_69 == 0) { -uint8_t x_162; lean_object* x_163; lean_object* x_164; -lean_dec(x_8); -x_162 = 1; -x_163 = lean_box(x_162); -x_164 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_164, 0, x_163); -lean_ctor_set(x_164, 1, x_103); -return x_164; -} -} -else +lean_object* x_70; uint8_t x_71; +x_70 = l_Lean_Meta_canUnfoldAtMatcher___closed__26; +x_71 = lean_name_eq(x_31, x_70); +if (x_71 == 0) { -uint8_t x_165; lean_object* x_166; lean_object* x_167; -lean_dec(x_8); -x_165 = 1; -x_166 = lean_box(x_165); -x_167 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_167, 0, x_166); -lean_ctor_set(x_167, 1, x_103); -return x_167; -} -} -else +lean_object* x_72; uint8_t x_73; +x_72 = l_Lean_Meta_canUnfoldAtMatcher___closed__27; +x_73 = lean_name_eq(x_31, x_72); +if (x_73 == 0) { -uint8_t x_168; lean_object* x_169; lean_object* x_170; -lean_dec(x_8); -x_168 = 1; -x_169 = lean_box(x_168); -x_170 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_170, 0, x_169); -lean_ctor_set(x_170, 1, x_103); -return x_170; -} -} -else +lean_object* x_74; uint8_t x_75; +x_74 = l_Lean_Meta_canUnfoldAtMatcher___closed__29; +x_75 = lean_name_eq(x_31, x_74); +if (x_75 == 0) { -uint8_t x_171; lean_object* x_172; lean_object* x_173; -lean_dec(x_8); -x_171 = 1; -x_172 = lean_box(x_171); -x_173 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_173, 0, x_172); -lean_ctor_set(x_173, 1, x_103); -return x_173; -} -} -else +lean_object* x_76; uint8_t x_77; +x_76 = l_Lean_Meta_canUnfoldAtMatcher___closed__30; +x_77 = lean_name_eq(x_31, x_76); +if (x_77 == 0) { -uint8_t x_174; lean_object* x_175; lean_object* x_176; -lean_dec(x_8); -x_174 = 1; -x_175 = lean_box(x_174); -x_176 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_176, 0, x_175); -lean_ctor_set(x_176, 1, x_103); -return x_176; -} -} -else +lean_object* x_78; uint8_t x_79; +x_78 = l_Lean_Meta_canUnfoldAtMatcher___closed__32; +x_79 = lean_name_eq(x_31, x_78); +if (x_79 == 0) { -uint8_t x_177; lean_object* x_178; lean_object* x_179; -lean_dec(x_8); -x_177 = 1; -x_178 = lean_box(x_177); -x_179 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_179, 0, x_178); -lean_ctor_set(x_179, 1, x_103); -return x_179; -} -} -else +lean_object* x_80; uint8_t x_81; +x_80 = l_Lean_Meta_canUnfoldAtMatcher___closed__33; +x_81 = lean_name_eq(x_31, x_80); +if (x_81 == 0) { -uint8_t x_180; lean_object* x_181; lean_object* x_182; -lean_dec(x_8); -x_180 = 1; -x_181 = lean_box(x_180); -x_182 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_182, 0, x_181); -lean_ctor_set(x_182, 1, x_103); -return x_182; -} -} -else +lean_object* x_82; uint8_t x_83; +x_82 = l_Lean_Meta_canUnfoldAtMatcher___closed__36; +x_83 = lean_name_eq(x_31, x_82); +if (x_83 == 0) { -uint8_t x_183; lean_object* x_184; lean_object* x_185; -lean_dec(x_8); -x_183 = 1; -x_184 = lean_box(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_103); -return x_185; -} +lean_object* x_84; uint8_t x_85; lean_object* x_86; +x_84 = l_Lean_Meta_canUnfoldAtMatcher___closed__39; +x_85 = lean_name_eq(x_31, x_84); +lean_dec(x_31); +x_86 = lean_box(x_85); +lean_ctor_set(x_42, 0, x_86); +return x_42; } else { -uint8_t x_186; lean_object* x_187; lean_object* x_188; -lean_dec(x_8); -x_186 = 1; -x_187 = lean_box(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_103); -return x_188; +uint8_t x_87; lean_object* x_88; +lean_dec(x_31); +x_87 = 1; +x_88 = lean_box(x_87); +lean_ctor_set(x_42, 0, x_88); +return x_42; } } else { -uint8_t x_189; lean_object* x_190; lean_object* x_191; -lean_dec(x_8); -x_189 = 1; -x_190 = lean_box(x_189); -x_191 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_191, 0, x_190); -lean_ctor_set(x_191, 1, x_103); -return x_191; +uint8_t x_89; lean_object* x_90; +lean_dec(x_31); +x_89 = 1; +x_90 = lean_box(x_89); +lean_ctor_set(x_42, 0, x_90); +return x_42; } } else { -uint8_t x_192; lean_object* x_193; lean_object* x_194; -lean_dec(x_8); -x_192 = 1; -x_193 = lean_box(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_103); -return x_194; +uint8_t x_91; lean_object* x_92; +lean_dec(x_31); +x_91 = 1; +x_92 = lean_box(x_91); +lean_ctor_set(x_42, 0, x_92); +return x_42; } } else { -uint8_t x_195; lean_object* x_196; lean_object* x_197; -lean_dec(x_8); -x_195 = 1; -x_196 = lean_box(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_103); -return x_197; +uint8_t x_93; lean_object* x_94; +lean_dec(x_31); +x_93 = 1; +x_94 = lean_box(x_93); +lean_ctor_set(x_42, 0, x_94); +return x_42; } } else { -uint8_t x_198; lean_object* x_199; lean_object* x_200; -lean_dec(x_8); -x_198 = 1; -x_199 = lean_box(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_103); -return x_200; +uint8_t x_95; lean_object* x_96; +lean_dec(x_31); +x_95 = 1; +x_96 = lean_box(x_95); +lean_ctor_set(x_42, 0, x_96); +return x_42; } } else { -uint8_t x_201; lean_object* x_202; lean_object* x_203; -lean_dec(x_8); -x_201 = 1; -x_202 = lean_box(x_201); -x_203 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_203, 0, x_202); -lean_ctor_set(x_203, 1, x_103); -return x_203; -} +uint8_t x_97; lean_object* x_98; +lean_dec(x_31); +x_97 = 1; +x_98 = lean_box(x_97); +lean_ctor_set(x_42, 0, x_98); +return x_42; } } else { -uint8_t x_204; lean_object* x_205; -lean_dec(x_8); -x_204 = 1; -x_205 = lean_box(x_204); -lean_ctor_set(x_12, 0, x_205); -return x_12; +uint8_t x_99; lean_object* x_100; +lean_dec(x_31); +x_99 = 1; +x_100 = lean_box(x_99); +lean_ctor_set(x_42, 0, x_100); +return x_42; } } else { -lean_object* x_206; lean_object* x_207; lean_object* x_208; uint8_t x_209; -x_206 = lean_ctor_get(x_12, 0); -x_207 = lean_ctor_get(x_12, 1); -lean_inc(x_207); -lean_inc(x_206); -lean_dec(x_12); -x_208 = lean_ctor_get(x_206, 0); -lean_inc(x_208); -lean_dec(x_206); -lean_inc(x_8); -x_209 = lean_is_instance(x_208, x_8); -if (x_209 == 0) -{ -lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; uint8_t x_216; -x_210 = lean_st_ref_get(x_4, x_207); -x_211 = lean_ctor_get(x_210, 0); -lean_inc(x_211); -x_212 = lean_ctor_get(x_210, 1); -lean_inc(x_212); -if (lean_is_exclusive(x_210)) { - lean_ctor_release(x_210, 0); - lean_ctor_release(x_210, 1); - x_213 = x_210; -} else { - lean_dec_ref(x_210); - x_213 = lean_box(0); -} -x_214 = lean_ctor_get(x_211, 0); -lean_inc(x_214); -lean_dec(x_211); -x_215 = l_Lean_Meta_canUnfoldAtMatcher___closed__1; -lean_inc(x_8); -x_216 = l_Lean_TagAttribute_hasTag(x_215, x_214, x_8); -if (x_216 == 0) -{ -lean_object* x_217; uint8_t x_218; -x_217 = l_Lean_Meta_canUnfoldAtMatcher___closed__3; -x_218 = lean_name_eq(x_8, x_217); -if (x_218 == 0) -{ -lean_object* x_219; uint8_t x_220; -x_219 = l_Lean_Meta_canUnfoldAtMatcher___closed__5; -x_220 = lean_name_eq(x_8, x_219); -if (x_220 == 0) -{ -lean_object* x_221; uint8_t x_222; -x_221 = l_Lean_Meta_canUnfoldAtMatcher___closed__7; -x_222 = lean_name_eq(x_8, x_221); -if (x_222 == 0) -{ -lean_object* x_223; uint8_t x_224; -x_223 = l_Lean_Meta_canUnfoldAtMatcher___closed__9; -x_224 = lean_name_eq(x_8, x_223); -if (x_224 == 0) -{ -lean_object* x_225; uint8_t x_226; -x_225 = l_Lean_Meta_canUnfoldAtMatcher___closed__12; -x_226 = lean_name_eq(x_8, x_225); -if (x_226 == 0) -{ -lean_object* x_227; uint8_t x_228; -x_227 = l_Lean_Meta_canUnfoldAtMatcher___closed__14; -x_228 = lean_name_eq(x_8, x_227); -if (x_228 == 0) -{ -lean_object* x_229; uint8_t x_230; -x_229 = l_Lean_Meta_canUnfoldAtMatcher___closed__16; -x_230 = lean_name_eq(x_8, x_229); -if (x_230 == 0) -{ -lean_object* x_231; uint8_t x_232; -x_231 = l_Lean_Meta_canUnfoldAtMatcher___closed__19; -x_232 = lean_name_eq(x_8, x_231); -if (x_232 == 0) -{ -lean_object* x_233; uint8_t x_234; -x_233 = l_Lean_Meta_canUnfoldAtMatcher___closed__21; -x_234 = lean_name_eq(x_8, x_233); -if (x_234 == 0) -{ -lean_object* x_235; uint8_t x_236; -x_235 = l_Lean_Meta_canUnfoldAtMatcher___closed__23; -x_236 = lean_name_eq(x_8, x_235); -if (x_236 == 0) -{ -lean_object* x_237; uint8_t x_238; -x_237 = l_Lean_Meta_canUnfoldAtMatcher___closed__24; -x_238 = lean_name_eq(x_8, x_237); -if (x_238 == 0) -{ -lean_object* x_239; uint8_t x_240; -x_239 = l_Lean_Meta_canUnfoldAtMatcher___closed__26; -x_240 = lean_name_eq(x_8, x_239); -if (x_240 == 0) -{ -lean_object* x_241; uint8_t x_242; -x_241 = l_Lean_Meta_canUnfoldAtMatcher___closed__27; -x_242 = lean_name_eq(x_8, x_241); -if (x_242 == 0) -{ -lean_object* x_243; uint8_t x_244; -x_243 = l_Lean_Meta_canUnfoldAtMatcher___closed__29; -x_244 = lean_name_eq(x_8, x_243); -if (x_244 == 0) -{ -lean_object* x_245; uint8_t x_246; -x_245 = l_Lean_Meta_canUnfoldAtMatcher___closed__30; -x_246 = lean_name_eq(x_8, x_245); -if (x_246 == 0) -{ -lean_object* x_247; uint8_t x_248; -x_247 = l_Lean_Meta_canUnfoldAtMatcher___closed__32; -x_248 = lean_name_eq(x_8, x_247); -if (x_248 == 0) -{ -lean_object* x_249; uint8_t x_250; -x_249 = l_Lean_Meta_canUnfoldAtMatcher___closed__33; -x_250 = lean_name_eq(x_8, x_249); -if (x_250 == 0) -{ -lean_object* x_251; uint8_t x_252; -x_251 = l_Lean_Meta_canUnfoldAtMatcher___closed__36; -x_252 = lean_name_eq(x_8, x_251); -if (x_252 == 0) -{ -lean_object* x_253; uint8_t x_254; lean_object* x_255; lean_object* x_256; -x_253 = l_Lean_Meta_canUnfoldAtMatcher___closed__39; -x_254 = lean_name_eq(x_8, x_253); -lean_dec(x_8); -x_255 = lean_box(x_254); -if (lean_is_scalar(x_213)) { - x_256 = lean_alloc_ctor(0, 2, 0); -} else { - x_256 = x_213; +uint8_t x_101; lean_object* x_102; +lean_dec(x_31); +x_101 = 1; +x_102 = lean_box(x_101); +lean_ctor_set(x_42, 0, x_102); +return x_42; } -lean_ctor_set(x_256, 0, x_255); -lean_ctor_set(x_256, 1, x_212); -return x_256; } else { -uint8_t x_257; lean_object* x_258; lean_object* x_259; -lean_dec(x_8); -x_257 = 1; -x_258 = lean_box(x_257); -if (lean_is_scalar(x_213)) { - x_259 = lean_alloc_ctor(0, 2, 0); -} else { - x_259 = x_213; -} -lean_ctor_set(x_259, 0, x_258); -lean_ctor_set(x_259, 1, x_212); -return x_259; +uint8_t x_103; lean_object* x_104; +lean_dec(x_31); +x_103 = 1; +x_104 = lean_box(x_103); +lean_ctor_set(x_42, 0, x_104); +return x_42; } } else { -uint8_t x_260; lean_object* x_261; lean_object* x_262; -lean_dec(x_8); -x_260 = 1; -x_261 = lean_box(x_260); -if (lean_is_scalar(x_213)) { - x_262 = lean_alloc_ctor(0, 2, 0); -} else { - x_262 = x_213; -} -lean_ctor_set(x_262, 0, x_261); -lean_ctor_set(x_262, 1, x_212); -return x_262; +uint8_t x_105; lean_object* x_106; +lean_dec(x_31); +x_105 = 1; +x_106 = lean_box(x_105); +lean_ctor_set(x_42, 0, x_106); +return x_42; } } else { -uint8_t x_263; lean_object* x_264; lean_object* x_265; -lean_dec(x_8); -x_263 = 1; -x_264 = lean_box(x_263); -if (lean_is_scalar(x_213)) { - x_265 = lean_alloc_ctor(0, 2, 0); -} else { - x_265 = x_213; -} -lean_ctor_set(x_265, 0, x_264); -lean_ctor_set(x_265, 1, x_212); -return x_265; +uint8_t x_107; lean_object* x_108; +lean_dec(x_31); +x_107 = 1; +x_108 = lean_box(x_107); +lean_ctor_set(x_42, 0, x_108); +return x_42; } } else { -uint8_t x_266; lean_object* x_267; lean_object* x_268; -lean_dec(x_8); -x_266 = 1; -x_267 = lean_box(x_266); -if (lean_is_scalar(x_213)) { - x_268 = lean_alloc_ctor(0, 2, 0); -} else { - x_268 = x_213; -} -lean_ctor_set(x_268, 0, x_267); -lean_ctor_set(x_268, 1, x_212); -return x_268; +uint8_t x_109; lean_object* x_110; +lean_dec(x_31); +x_109 = 1; +x_110 = lean_box(x_109); +lean_ctor_set(x_42, 0, x_110); +return x_42; } } else { -uint8_t x_269; lean_object* x_270; lean_object* x_271; -lean_dec(x_8); -x_269 = 1; -x_270 = lean_box(x_269); -if (lean_is_scalar(x_213)) { - x_271 = lean_alloc_ctor(0, 2, 0); -} else { - x_271 = x_213; -} -lean_ctor_set(x_271, 0, x_270); -lean_ctor_set(x_271, 1, x_212); -return x_271; +uint8_t x_111; lean_object* x_112; +lean_dec(x_31); +x_111 = 1; +x_112 = lean_box(x_111); +lean_ctor_set(x_42, 0, x_112); +return x_42; } } else { -uint8_t x_272; lean_object* x_273; lean_object* x_274; -lean_dec(x_8); -x_272 = 1; -x_273 = lean_box(x_272); -if (lean_is_scalar(x_213)) { - x_274 = lean_alloc_ctor(0, 2, 0); -} else { - x_274 = x_213; -} -lean_ctor_set(x_274, 0, x_273); -lean_ctor_set(x_274, 1, x_212); -return x_274; +uint8_t x_113; lean_object* x_114; +lean_dec(x_31); +x_113 = 1; +x_114 = lean_box(x_113); +lean_ctor_set(x_42, 0, x_114); +return x_42; } } else { -uint8_t x_275; lean_object* x_276; lean_object* x_277; -lean_dec(x_8); -x_275 = 1; -x_276 = lean_box(x_275); -if (lean_is_scalar(x_213)) { - x_277 = lean_alloc_ctor(0, 2, 0); -} else { - x_277 = x_213; -} -lean_ctor_set(x_277, 0, x_276); -lean_ctor_set(x_277, 1, x_212); -return x_277; +uint8_t x_115; lean_object* x_116; +lean_dec(x_31); +x_115 = 1; +x_116 = lean_box(x_115); +lean_ctor_set(x_42, 0, x_116); +return x_42; } } else { -uint8_t x_278; lean_object* x_279; lean_object* x_280; -lean_dec(x_8); -x_278 = 1; -x_279 = lean_box(x_278); -if (lean_is_scalar(x_213)) { - x_280 = lean_alloc_ctor(0, 2, 0); -} else { - x_280 = x_213; -} -lean_ctor_set(x_280, 0, x_279); -lean_ctor_set(x_280, 1, x_212); -return x_280; +uint8_t x_117; lean_object* x_118; +lean_dec(x_31); +x_117 = 1; +x_118 = lean_box(x_117); +lean_ctor_set(x_42, 0, x_118); +return x_42; } } else { -uint8_t x_281; lean_object* x_282; lean_object* x_283; -lean_dec(x_8); -x_281 = 1; -x_282 = lean_box(x_281); -if (lean_is_scalar(x_213)) { - x_283 = lean_alloc_ctor(0, 2, 0); -} else { - x_283 = x_213; -} -lean_ctor_set(x_283, 0, x_282); -lean_ctor_set(x_283, 1, x_212); -return x_283; +uint8_t x_119; lean_object* x_120; +lean_dec(x_31); +x_119 = 1; +x_120 = lean_box(x_119); +lean_ctor_set(x_42, 0, x_120); +return x_42; } } else { -uint8_t x_284; lean_object* x_285; lean_object* x_286; -lean_dec(x_8); -x_284 = 1; -x_285 = lean_box(x_284); -if (lean_is_scalar(x_213)) { - x_286 = lean_alloc_ctor(0, 2, 0); -} else { - x_286 = x_213; -} -lean_ctor_set(x_286, 0, x_285); -lean_ctor_set(x_286, 1, x_212); -return x_286; +uint8_t x_121; lean_object* x_122; +lean_dec(x_31); +x_121 = 1; +x_122 = lean_box(x_121); +lean_ctor_set(x_42, 0, x_122); +return x_42; } } else { -uint8_t x_287; lean_object* x_288; lean_object* x_289; -lean_dec(x_8); -x_287 = 1; -x_288 = lean_box(x_287); -if (lean_is_scalar(x_213)) { - x_289 = lean_alloc_ctor(0, 2, 0); -} else { - x_289 = x_213; -} -lean_ctor_set(x_289, 0, x_288); -lean_ctor_set(x_289, 1, x_212); -return x_289; +uint8_t x_123; lean_object* x_124; +lean_dec(x_31); +x_123 = 1; +x_124 = lean_box(x_123); +lean_ctor_set(x_42, 0, x_124); +return x_42; } } else { -uint8_t x_290; lean_object* x_291; lean_object* x_292; -lean_dec(x_8); -x_290 = 1; -x_291 = lean_box(x_290); -if (lean_is_scalar(x_213)) { - x_292 = lean_alloc_ctor(0, 2, 0); -} else { - x_292 = x_213; -} -lean_ctor_set(x_292, 0, x_291); -lean_ctor_set(x_292, 1, x_212); -return x_292; -} -} -else +lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; uint8_t x_129; +x_125 = lean_ctor_get(x_42, 0); +x_126 = lean_ctor_get(x_42, 1); +lean_inc(x_126); +lean_inc(x_125); +lean_dec(x_42); +x_127 = lean_ctor_get(x_125, 0); +lean_inc(x_127); +lean_dec(x_125); +x_128 = l_Lean_Meta_canUnfoldAtMatcher___closed__1; +lean_inc(x_31); +x_129 = l_Lean_TagAttribute_hasTag(x_128, x_127, x_31); +if (x_129 == 0) { -uint8_t x_293; lean_object* x_294; lean_object* x_295; -lean_dec(x_8); -x_293 = 1; -x_294 = lean_box(x_293); -if (lean_is_scalar(x_213)) { - x_295 = lean_alloc_ctor(0, 2, 0); -} else { - x_295 = x_213; -} -lean_ctor_set(x_295, 0, x_294); -lean_ctor_set(x_295, 1, x_212); -return x_295; -} -} -else +lean_object* x_130; uint8_t x_131; +x_130 = l_Lean_Meta_canUnfoldAtMatcher___closed__3; +x_131 = lean_name_eq(x_31, x_130); +if (x_131 == 0) { -uint8_t x_296; lean_object* x_297; lean_object* x_298; -lean_dec(x_8); -x_296 = 1; -x_297 = lean_box(x_296); -if (lean_is_scalar(x_213)) { - x_298 = lean_alloc_ctor(0, 2, 0); -} else { - x_298 = x_213; -} -lean_ctor_set(x_298, 0, x_297); -lean_ctor_set(x_298, 1, x_212); -return x_298; -} -} -else +lean_object* x_132; uint8_t x_133; +x_132 = l_Lean_Meta_canUnfoldAtMatcher___closed__5; +x_133 = lean_name_eq(x_31, x_132); +if (x_133 == 0) { -uint8_t x_299; lean_object* x_300; lean_object* x_301; -lean_dec(x_8); -x_299 = 1; -x_300 = lean_box(x_299); -if (lean_is_scalar(x_213)) { - x_301 = lean_alloc_ctor(0, 2, 0); -} else { - x_301 = x_213; -} -lean_ctor_set(x_301, 0, x_300); -lean_ctor_set(x_301, 1, x_212); -return x_301; -} -} -else +lean_object* x_134; uint8_t x_135; +x_134 = l_Lean_Meta_canUnfoldAtMatcher___closed__7; +x_135 = lean_name_eq(x_31, x_134); +if (x_135 == 0) { -uint8_t x_302; lean_object* x_303; lean_object* x_304; -lean_dec(x_8); -x_302 = 1; -x_303 = lean_box(x_302); -if (lean_is_scalar(x_213)) { - x_304 = lean_alloc_ctor(0, 2, 0); -} else { - x_304 = x_213; -} -lean_ctor_set(x_304, 0, x_303); -lean_ctor_set(x_304, 1, x_212); -return x_304; -} -} -else +lean_object* x_136; uint8_t x_137; +x_136 = l_Lean_Meta_canUnfoldAtMatcher___closed__9; +x_137 = lean_name_eq(x_31, x_136); +if (x_137 == 0) { -uint8_t x_305; lean_object* x_306; lean_object* x_307; -lean_dec(x_8); -x_305 = 1; -x_306 = lean_box(x_305); -if (lean_is_scalar(x_213)) { - x_307 = lean_alloc_ctor(0, 2, 0); -} else { - x_307 = x_213; -} -lean_ctor_set(x_307, 0, x_306); -lean_ctor_set(x_307, 1, x_212); -return x_307; -} -} -else +lean_object* x_138; uint8_t x_139; +x_138 = l_Lean_Meta_canUnfoldAtMatcher___closed__12; +x_139 = lean_name_eq(x_31, x_138); +if (x_139 == 0) { -uint8_t x_308; lean_object* x_309; lean_object* x_310; -lean_dec(x_8); -x_308 = 1; -x_309 = lean_box(x_308); -if (lean_is_scalar(x_213)) { - x_310 = lean_alloc_ctor(0, 2, 0); -} else { - x_310 = x_213; -} -lean_ctor_set(x_310, 0, x_309); -lean_ctor_set(x_310, 1, x_212); -return x_310; -} -} -else +lean_object* x_140; uint8_t x_141; +x_140 = l_Lean_Meta_canUnfoldAtMatcher___closed__14; +x_141 = lean_name_eq(x_31, x_140); +if (x_141 == 0) { -uint8_t x_311; lean_object* x_312; lean_object* x_313; -lean_dec(x_8); -x_311 = 1; -x_312 = lean_box(x_311); -if (lean_is_scalar(x_213)) { - x_313 = lean_alloc_ctor(0, 2, 0); -} else { - x_313 = x_213; -} -lean_ctor_set(x_313, 0, x_312); -lean_ctor_set(x_313, 1, x_212); -return x_313; -} -} -else +lean_object* x_142; uint8_t x_143; +x_142 = l_Lean_Meta_canUnfoldAtMatcher___closed__16; +x_143 = lean_name_eq(x_31, x_142); +if (x_143 == 0) { -uint8_t x_314; lean_object* x_315; lean_object* x_316; -lean_dec(x_8); -x_314 = 1; -x_315 = lean_box(x_314); -x_316 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_316, 0, x_315); -lean_ctor_set(x_316, 1, x_207); -return x_316; -} -} -} -else +lean_object* x_144; uint8_t x_145; +x_144 = l_Lean_Meta_canUnfoldAtMatcher___closed__19; +x_145 = lean_name_eq(x_31, x_144); +if (x_145 == 0) { -uint8_t x_317; -lean_dec(x_8); -x_317 = !lean_is_exclusive(x_12); -if (x_317 == 0) +lean_object* x_146; uint8_t x_147; +x_146 = l_Lean_Meta_canUnfoldAtMatcher___closed__21; +x_147 = lean_name_eq(x_31, x_146); +if (x_147 == 0) { -lean_object* x_318; uint8_t x_319; lean_object* x_320; -x_318 = lean_ctor_get(x_12, 0); -lean_dec(x_318); -x_319 = 1; -x_320 = lean_box(x_319); -lean_ctor_set(x_12, 0, x_320); -return x_12; -} -else +lean_object* x_148; uint8_t x_149; +x_148 = l_Lean_Meta_canUnfoldAtMatcher___closed__23; +x_149 = lean_name_eq(x_31, x_148); +if (x_149 == 0) { -lean_object* x_321; uint8_t x_322; lean_object* x_323; lean_object* x_324; -x_321 = lean_ctor_get(x_12, 1); -lean_inc(x_321); -lean_dec(x_12); -x_322 = 1; -x_323 = lean_box(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_321); -return x_324; -} -} -} -case 3: +lean_object* x_150; uint8_t x_151; +x_150 = l_Lean_Meta_canUnfoldAtMatcher___closed__24; +x_151 = lean_name_eq(x_31, x_150); +if (x_151 == 0) { -lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; uint8_t x_330; -x_325 = l_Lean_ConstantInfo_name(x_2); -lean_inc(x_325); -x_326 = l_Lean_isReducible___at___private_Lean_Meta_GetConst_0__Lean_Meta_canUnfoldDefault___spec__3(x_325, x_3, x_4, x_5); -x_327 = lean_ctor_get(x_326, 0); -lean_inc(x_327); -x_328 = lean_ctor_get(x_326, 1); -lean_inc(x_328); -lean_dec(x_326); -x_329 = lean_st_ref_get(x_4, x_328); -x_330 = lean_unbox(x_327); -lean_dec(x_327); -if (x_330 == 0) -{ -uint8_t x_331; -x_331 = !lean_is_exclusive(x_329); -if (x_331 == 0) +lean_object* x_152; uint8_t x_153; +x_152 = l_Lean_Meta_canUnfoldAtMatcher___closed__26; +x_153 = lean_name_eq(x_31, x_152); +if (x_153 == 0) { -lean_object* x_332; lean_object* x_333; lean_object* x_334; uint8_t x_335; -x_332 = lean_ctor_get(x_329, 0); -x_333 = lean_ctor_get(x_329, 1); -x_334 = lean_ctor_get(x_332, 0); -lean_inc(x_334); -lean_dec(x_332); -lean_inc(x_325); -x_335 = lean_is_instance(x_334, x_325); -if (x_335 == 0) -{ -lean_object* x_336; uint8_t x_337; -lean_free_object(x_329); -x_336 = lean_st_ref_get(x_4, x_333); -x_337 = !lean_is_exclusive(x_336); -if (x_337 == 0) -{ -lean_object* x_338; lean_object* x_339; lean_object* x_340; uint8_t x_341; -x_338 = lean_ctor_get(x_336, 0); -x_339 = lean_ctor_get(x_338, 0); -lean_inc(x_339); -lean_dec(x_338); -x_340 = l_Lean_Meta_canUnfoldAtMatcher___closed__1; -lean_inc(x_325); -x_341 = l_Lean_TagAttribute_hasTag(x_340, x_339, x_325); -if (x_341 == 0) -{ -lean_object* x_342; uint8_t x_343; -x_342 = l_Lean_Meta_canUnfoldAtMatcher___closed__3; -x_343 = lean_name_eq(x_325, x_342); -if (x_343 == 0) -{ -lean_object* x_344; uint8_t x_345; -x_344 = l_Lean_Meta_canUnfoldAtMatcher___closed__5; -x_345 = lean_name_eq(x_325, x_344); -if (x_345 == 0) -{ -lean_object* x_346; uint8_t x_347; -x_346 = l_Lean_Meta_canUnfoldAtMatcher___closed__7; -x_347 = lean_name_eq(x_325, x_346); -if (x_347 == 0) -{ -lean_object* x_348; uint8_t x_349; -x_348 = l_Lean_Meta_canUnfoldAtMatcher___closed__9; -x_349 = lean_name_eq(x_325, x_348); -if (x_349 == 0) -{ -lean_object* x_350; uint8_t x_351; -x_350 = l_Lean_Meta_canUnfoldAtMatcher___closed__12; -x_351 = lean_name_eq(x_325, x_350); -if (x_351 == 0) -{ -lean_object* x_352; uint8_t x_353; -x_352 = l_Lean_Meta_canUnfoldAtMatcher___closed__14; -x_353 = lean_name_eq(x_325, x_352); -if (x_353 == 0) -{ -lean_object* x_354; uint8_t x_355; -x_354 = l_Lean_Meta_canUnfoldAtMatcher___closed__16; -x_355 = lean_name_eq(x_325, x_354); -if (x_355 == 0) +lean_object* x_154; uint8_t x_155; +x_154 = l_Lean_Meta_canUnfoldAtMatcher___closed__27; +x_155 = lean_name_eq(x_31, x_154); +if (x_155 == 0) { -lean_object* x_356; uint8_t x_357; -x_356 = l_Lean_Meta_canUnfoldAtMatcher___closed__19; -x_357 = lean_name_eq(x_325, x_356); -if (x_357 == 0) +lean_object* x_156; uint8_t x_157; +x_156 = l_Lean_Meta_canUnfoldAtMatcher___closed__29; +x_157 = lean_name_eq(x_31, x_156); +if (x_157 == 0) { -lean_object* x_358; uint8_t x_359; -x_358 = l_Lean_Meta_canUnfoldAtMatcher___closed__21; -x_359 = lean_name_eq(x_325, x_358); -if (x_359 == 0) +lean_object* x_158; uint8_t x_159; +x_158 = l_Lean_Meta_canUnfoldAtMatcher___closed__30; +x_159 = lean_name_eq(x_31, x_158); +if (x_159 == 0) { -lean_object* x_360; uint8_t x_361; -x_360 = l_Lean_Meta_canUnfoldAtMatcher___closed__23; -x_361 = lean_name_eq(x_325, x_360); -if (x_361 == 0) +lean_object* x_160; uint8_t x_161; +x_160 = l_Lean_Meta_canUnfoldAtMatcher___closed__32; +x_161 = lean_name_eq(x_31, x_160); +if (x_161 == 0) { -lean_object* x_362; uint8_t x_363; -x_362 = l_Lean_Meta_canUnfoldAtMatcher___closed__24; -x_363 = lean_name_eq(x_325, x_362); -if (x_363 == 0) -{ -lean_object* x_364; uint8_t x_365; -x_364 = l_Lean_Meta_canUnfoldAtMatcher___closed__26; -x_365 = lean_name_eq(x_325, x_364); -if (x_365 == 0) -{ -lean_object* x_366; uint8_t x_367; -x_366 = l_Lean_Meta_canUnfoldAtMatcher___closed__27; -x_367 = lean_name_eq(x_325, x_366); -if (x_367 == 0) -{ -lean_object* x_368; uint8_t x_369; -x_368 = l_Lean_Meta_canUnfoldAtMatcher___closed__29; -x_369 = lean_name_eq(x_325, x_368); -if (x_369 == 0) -{ -lean_object* x_370; uint8_t x_371; -x_370 = l_Lean_Meta_canUnfoldAtMatcher___closed__30; -x_371 = lean_name_eq(x_325, x_370); -if (x_371 == 0) -{ -lean_object* x_372; uint8_t x_373; -x_372 = l_Lean_Meta_canUnfoldAtMatcher___closed__32; -x_373 = lean_name_eq(x_325, x_372); -if (x_373 == 0) -{ -lean_object* x_374; uint8_t x_375; -x_374 = l_Lean_Meta_canUnfoldAtMatcher___closed__33; -x_375 = lean_name_eq(x_325, x_374); -if (x_375 == 0) -{ -lean_object* x_376; uint8_t x_377; -x_376 = l_Lean_Meta_canUnfoldAtMatcher___closed__36; -x_377 = lean_name_eq(x_325, x_376); -if (x_377 == 0) -{ -lean_object* x_378; uint8_t x_379; lean_object* x_380; -x_378 = l_Lean_Meta_canUnfoldAtMatcher___closed__39; -x_379 = lean_name_eq(x_325, x_378); -lean_dec(x_325); -x_380 = lean_box(x_379); -lean_ctor_set(x_336, 0, x_380); -return x_336; -} -else +lean_object* x_162; uint8_t x_163; +x_162 = l_Lean_Meta_canUnfoldAtMatcher___closed__33; +x_163 = lean_name_eq(x_31, x_162); +if (x_163 == 0) { -uint8_t x_381; lean_object* x_382; -lean_dec(x_325); -x_381 = 1; -x_382 = lean_box(x_381); -lean_ctor_set(x_336, 0, x_382); -return x_336; -} -} -else +lean_object* x_164; uint8_t x_165; +x_164 = l_Lean_Meta_canUnfoldAtMatcher___closed__36; +x_165 = lean_name_eq(x_31, x_164); +if (x_165 == 0) { -uint8_t x_383; lean_object* x_384; -lean_dec(x_325); -x_383 = 1; -x_384 = lean_box(x_383); -lean_ctor_set(x_336, 0, x_384); -return x_336; -} +lean_object* x_166; uint8_t x_167; lean_object* x_168; lean_object* x_169; +x_166 = l_Lean_Meta_canUnfoldAtMatcher___closed__39; +x_167 = lean_name_eq(x_31, x_166); +lean_dec(x_31); +x_168 = lean_box(x_167); +x_169 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_169, 0, x_168); +lean_ctor_set(x_169, 1, x_126); +return x_169; } else { -uint8_t x_385; lean_object* x_386; -lean_dec(x_325); -x_385 = 1; -x_386 = lean_box(x_385); -lean_ctor_set(x_336, 0, x_386); -return x_336; +uint8_t x_170; lean_object* x_171; lean_object* x_172; +lean_dec(x_31); +x_170 = 1; +x_171 = lean_box(x_170); +x_172 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_172, 0, x_171); +lean_ctor_set(x_172, 1, x_126); +return x_172; } } else { -uint8_t x_387; lean_object* x_388; -lean_dec(x_325); -x_387 = 1; -x_388 = lean_box(x_387); -lean_ctor_set(x_336, 0, x_388); -return x_336; +uint8_t x_173; lean_object* x_174; lean_object* x_175; +lean_dec(x_31); +x_173 = 1; +x_174 = lean_box(x_173); +x_175 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_175, 0, x_174); +lean_ctor_set(x_175, 1, x_126); +return x_175; } } else { -uint8_t x_389; lean_object* x_390; -lean_dec(x_325); -x_389 = 1; -x_390 = lean_box(x_389); -lean_ctor_set(x_336, 0, x_390); -return x_336; +uint8_t x_176; lean_object* x_177; lean_object* x_178; +lean_dec(x_31); +x_176 = 1; +x_177 = lean_box(x_176); +x_178 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_178, 0, x_177); +lean_ctor_set(x_178, 1, x_126); +return x_178; } } else { -uint8_t x_391; lean_object* x_392; -lean_dec(x_325); -x_391 = 1; -x_392 = lean_box(x_391); -lean_ctor_set(x_336, 0, x_392); -return x_336; +uint8_t x_179; lean_object* x_180; lean_object* x_181; +lean_dec(x_31); +x_179 = 1; +x_180 = lean_box(x_179); +x_181 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_181, 0, x_180); +lean_ctor_set(x_181, 1, x_126); +return x_181; } } else { -uint8_t x_393; lean_object* x_394; -lean_dec(x_325); -x_393 = 1; -x_394 = lean_box(x_393); -lean_ctor_set(x_336, 0, x_394); -return x_336; +uint8_t x_182; lean_object* x_183; lean_object* x_184; +lean_dec(x_31); +x_182 = 1; +x_183 = lean_box(x_182); +x_184 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_184, 0, x_183); +lean_ctor_set(x_184, 1, x_126); +return x_184; } } else { -uint8_t x_395; lean_object* x_396; -lean_dec(x_325); -x_395 = 1; -x_396 = lean_box(x_395); -lean_ctor_set(x_336, 0, x_396); -return x_336; +uint8_t x_185; lean_object* x_186; lean_object* x_187; +lean_dec(x_31); +x_185 = 1; +x_186 = lean_box(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_126); +return x_187; } } else { -uint8_t x_397; lean_object* x_398; -lean_dec(x_325); -x_397 = 1; -x_398 = lean_box(x_397); -lean_ctor_set(x_336, 0, x_398); -return x_336; +uint8_t x_188; lean_object* x_189; lean_object* x_190; +lean_dec(x_31); +x_188 = 1; +x_189 = lean_box(x_188); +x_190 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_190, 0, x_189); +lean_ctor_set(x_190, 1, x_126); +return x_190; } } else { -uint8_t x_399; lean_object* x_400; -lean_dec(x_325); -x_399 = 1; -x_400 = lean_box(x_399); -lean_ctor_set(x_336, 0, x_400); -return x_336; +uint8_t x_191; lean_object* x_192; lean_object* x_193; +lean_dec(x_31); +x_191 = 1; +x_192 = lean_box(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_126); +return x_193; } } else { -uint8_t x_401; lean_object* x_402; -lean_dec(x_325); -x_401 = 1; -x_402 = lean_box(x_401); -lean_ctor_set(x_336, 0, x_402); -return x_336; +uint8_t x_194; lean_object* x_195; lean_object* x_196; +lean_dec(x_31); +x_194 = 1; +x_195 = lean_box(x_194); +x_196 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_196, 0, x_195); +lean_ctor_set(x_196, 1, x_126); +return x_196; } } else { -uint8_t x_403; lean_object* x_404; -lean_dec(x_325); -x_403 = 1; -x_404 = lean_box(x_403); -lean_ctor_set(x_336, 0, x_404); -return x_336; +uint8_t x_197; lean_object* x_198; lean_object* x_199; +lean_dec(x_31); +x_197 = 1; +x_198 = lean_box(x_197); +x_199 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_199, 0, x_198); +lean_ctor_set(x_199, 1, x_126); +return x_199; } } else { -uint8_t x_405; lean_object* x_406; -lean_dec(x_325); -x_405 = 1; -x_406 = lean_box(x_405); -lean_ctor_set(x_336, 0, x_406); -return x_336; +uint8_t x_200; lean_object* x_201; lean_object* x_202; +lean_dec(x_31); +x_200 = 1; +x_201 = lean_box(x_200); +x_202 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_202, 0, x_201); +lean_ctor_set(x_202, 1, x_126); +return x_202; } } else { -uint8_t x_407; lean_object* x_408; -lean_dec(x_325); -x_407 = 1; -x_408 = lean_box(x_407); -lean_ctor_set(x_336, 0, x_408); -return x_336; +uint8_t x_203; lean_object* x_204; lean_object* x_205; +lean_dec(x_31); +x_203 = 1; +x_204 = lean_box(x_203); +x_205 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_205, 0, x_204); +lean_ctor_set(x_205, 1, x_126); +return x_205; } } else { -uint8_t x_409; lean_object* x_410; -lean_dec(x_325); -x_409 = 1; -x_410 = lean_box(x_409); -lean_ctor_set(x_336, 0, x_410); -return x_336; +uint8_t x_206; lean_object* x_207; lean_object* x_208; +lean_dec(x_31); +x_206 = 1; +x_207 = lean_box(x_206); +x_208 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_208, 0, x_207); +lean_ctor_set(x_208, 1, x_126); +return x_208; } } else { -uint8_t x_411; lean_object* x_412; -lean_dec(x_325); -x_411 = 1; -x_412 = lean_box(x_411); -lean_ctor_set(x_336, 0, x_412); -return x_336; +uint8_t x_209; lean_object* x_210; lean_object* x_211; +lean_dec(x_31); +x_209 = 1; +x_210 = lean_box(x_209); +x_211 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_211, 0, x_210); +lean_ctor_set(x_211, 1, x_126); +return x_211; } } else { -uint8_t x_413; lean_object* x_414; -lean_dec(x_325); -x_413 = 1; -x_414 = lean_box(x_413); -lean_ctor_set(x_336, 0, x_414); -return x_336; +uint8_t x_212; lean_object* x_213; lean_object* x_214; +lean_dec(x_31); +x_212 = 1; +x_213 = lean_box(x_212); +x_214 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_214, 0, x_213); +lean_ctor_set(x_214, 1, x_126); +return x_214; } } else { -uint8_t x_415; lean_object* x_416; -lean_dec(x_325); -x_415 = 1; -x_416 = lean_box(x_415); -lean_ctor_set(x_336, 0, x_416); -return x_336; +uint8_t x_215; lean_object* x_216; lean_object* x_217; +lean_dec(x_31); +x_215 = 1; +x_216 = lean_box(x_215); +x_217 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_217, 0, x_216); +lean_ctor_set(x_217, 1, x_126); +return x_217; } } else { -uint8_t x_417; lean_object* x_418; -lean_dec(x_325); -x_417 = 1; -x_418 = lean_box(x_417); -lean_ctor_set(x_336, 0, x_418); -return x_336; +uint8_t x_218; lean_object* x_219; lean_object* x_220; +lean_dec(x_31); +x_218 = 1; +x_219 = lean_box(x_218); +x_220 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_220, 0, x_219); +lean_ctor_set(x_220, 1, x_126); +return x_220; } } else { -lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; uint8_t x_423; -x_419 = lean_ctor_get(x_336, 0); -x_420 = lean_ctor_get(x_336, 1); -lean_inc(x_420); -lean_inc(x_419); -lean_dec(x_336); -x_421 = lean_ctor_get(x_419, 0); -lean_inc(x_421); -lean_dec(x_419); -x_422 = l_Lean_Meta_canUnfoldAtMatcher___closed__1; -lean_inc(x_325); -x_423 = l_Lean_TagAttribute_hasTag(x_422, x_421, x_325); -if (x_423 == 0) -{ -lean_object* x_424; uint8_t x_425; -x_424 = l_Lean_Meta_canUnfoldAtMatcher___closed__3; -x_425 = lean_name_eq(x_325, x_424); -if (x_425 == 0) -{ -lean_object* x_426; uint8_t x_427; -x_426 = l_Lean_Meta_canUnfoldAtMatcher___closed__5; -x_427 = lean_name_eq(x_325, x_426); -if (x_427 == 0) -{ -lean_object* x_428; uint8_t x_429; -x_428 = l_Lean_Meta_canUnfoldAtMatcher___closed__7; -x_429 = lean_name_eq(x_325, x_428); -if (x_429 == 0) -{ -lean_object* x_430; uint8_t x_431; -x_430 = l_Lean_Meta_canUnfoldAtMatcher___closed__9; -x_431 = lean_name_eq(x_325, x_430); -if (x_431 == 0) -{ -lean_object* x_432; uint8_t x_433; -x_432 = l_Lean_Meta_canUnfoldAtMatcher___closed__12; -x_433 = lean_name_eq(x_325, x_432); -if (x_433 == 0) -{ -lean_object* x_434; uint8_t x_435; -x_434 = l_Lean_Meta_canUnfoldAtMatcher___closed__14; -x_435 = lean_name_eq(x_325, x_434); -if (x_435 == 0) -{ -lean_object* x_436; uint8_t x_437; -x_436 = l_Lean_Meta_canUnfoldAtMatcher___closed__16; -x_437 = lean_name_eq(x_325, x_436); -if (x_437 == 0) -{ -lean_object* x_438; uint8_t x_439; -x_438 = l_Lean_Meta_canUnfoldAtMatcher___closed__19; -x_439 = lean_name_eq(x_325, x_438); -if (x_439 == 0) -{ -lean_object* x_440; uint8_t x_441; -x_440 = l_Lean_Meta_canUnfoldAtMatcher___closed__21; -x_441 = lean_name_eq(x_325, x_440); -if (x_441 == 0) -{ -lean_object* x_442; uint8_t x_443; -x_442 = l_Lean_Meta_canUnfoldAtMatcher___closed__23; -x_443 = lean_name_eq(x_325, x_442); -if (x_443 == 0) -{ -lean_object* x_444; uint8_t x_445; -x_444 = l_Lean_Meta_canUnfoldAtMatcher___closed__24; -x_445 = lean_name_eq(x_325, x_444); -if (x_445 == 0) -{ -lean_object* x_446; uint8_t x_447; -x_446 = l_Lean_Meta_canUnfoldAtMatcher___closed__26; -x_447 = lean_name_eq(x_325, x_446); -if (x_447 == 0) -{ -lean_object* x_448; uint8_t x_449; -x_448 = l_Lean_Meta_canUnfoldAtMatcher___closed__27; -x_449 = lean_name_eq(x_325, x_448); -if (x_449 == 0) -{ -lean_object* x_450; uint8_t x_451; -x_450 = l_Lean_Meta_canUnfoldAtMatcher___closed__29; -x_451 = lean_name_eq(x_325, x_450); -if (x_451 == 0) -{ -lean_object* x_452; uint8_t x_453; -x_452 = l_Lean_Meta_canUnfoldAtMatcher___closed__30; -x_453 = lean_name_eq(x_325, x_452); -if (x_453 == 0) -{ -lean_object* x_454; uint8_t x_455; -x_454 = l_Lean_Meta_canUnfoldAtMatcher___closed__32; -x_455 = lean_name_eq(x_325, x_454); -if (x_455 == 0) -{ -lean_object* x_456; uint8_t x_457; -x_456 = l_Lean_Meta_canUnfoldAtMatcher___closed__33; -x_457 = lean_name_eq(x_325, x_456); -if (x_457 == 0) -{ -lean_object* x_458; uint8_t x_459; -x_458 = l_Lean_Meta_canUnfoldAtMatcher___closed__36; -x_459 = lean_name_eq(x_325, x_458); -if (x_459 == 0) -{ -lean_object* x_460; uint8_t x_461; lean_object* x_462; lean_object* x_463; -x_460 = l_Lean_Meta_canUnfoldAtMatcher___closed__39; -x_461 = lean_name_eq(x_325, x_460); -lean_dec(x_325); -x_462 = lean_box(x_461); -x_463 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_463, 0, x_462); -lean_ctor_set(x_463, 1, x_420); -return x_463; -} -else -{ -uint8_t x_464; lean_object* x_465; lean_object* x_466; -lean_dec(x_325); -x_464 = 1; -x_465 = lean_box(x_464); -x_466 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_466, 0, x_465); -lean_ctor_set(x_466, 1, x_420); -return x_466; +uint8_t x_221; lean_object* x_222; lean_object* x_223; +lean_dec(x_31); +x_221 = 1; +x_222 = lean_box(x_221); +x_223 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_223, 0, x_222); +lean_ctor_set(x_223, 1, x_126); +return x_223; } } else { -uint8_t x_467; lean_object* x_468; lean_object* x_469; -lean_dec(x_325); -x_467 = 1; -x_468 = lean_box(x_467); -x_469 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_469, 0, x_468); -lean_ctor_set(x_469, 1, x_420); -return x_469; -} +uint8_t x_224; lean_object* x_225; lean_object* x_226; +lean_dec(x_31); +x_224 = 1; +x_225 = lean_box(x_224); +x_226 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_226, 0, x_225); +lean_ctor_set(x_226, 1, x_126); +return x_226; } -else -{ -uint8_t x_470; lean_object* x_471; lean_object* x_472; -lean_dec(x_325); -x_470 = 1; -x_471 = lean_box(x_470); -x_472 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_472, 0, x_471); -lean_ctor_set(x_472, 1, x_420); -return x_472; } } else { -uint8_t x_473; lean_object* x_474; lean_object* x_475; -lean_dec(x_325); -x_473 = 1; -x_474 = lean_box(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_420); -return x_475; +uint8_t x_227; lean_object* x_228; +lean_dec(x_31); +x_227 = 1; +x_228 = lean_box(x_227); +lean_ctor_set(x_35, 0, x_228); +return x_35; } } else { -uint8_t x_476; lean_object* x_477; lean_object* x_478; -lean_dec(x_325); -x_476 = 1; -x_477 = lean_box(x_476); -x_478 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_478, 0, x_477); -lean_ctor_set(x_478, 1, x_420); -return x_478; -} -} -else +lean_object* x_229; lean_object* x_230; lean_object* x_231; uint8_t x_232; +x_229 = lean_ctor_get(x_35, 0); +x_230 = lean_ctor_get(x_35, 1); +lean_inc(x_230); +lean_inc(x_229); +lean_dec(x_35); +x_231 = lean_ctor_get(x_229, 0); +lean_inc(x_231); +lean_dec(x_229); +lean_inc(x_31); +x_232 = lean_is_instance(x_231, x_31); +if (x_232 == 0) { -uint8_t x_479; lean_object* x_480; lean_object* x_481; -lean_dec(x_325); -x_479 = 1; -x_480 = lean_box(x_479); -x_481 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_481, 0, x_480); -lean_ctor_set(x_481, 1, x_420); -return x_481; -} +lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; uint8_t x_239; +x_233 = lean_st_ref_get(x_4, x_230); +x_234 = lean_ctor_get(x_233, 0); +lean_inc(x_234); +x_235 = lean_ctor_get(x_233, 1); +lean_inc(x_235); +if (lean_is_exclusive(x_233)) { + lean_ctor_release(x_233, 0); + lean_ctor_release(x_233, 1); + x_236 = x_233; +} else { + lean_dec_ref(x_233); + x_236 = lean_box(0); } -else +x_237 = lean_ctor_get(x_234, 0); +lean_inc(x_237); +lean_dec(x_234); +x_238 = l_Lean_Meta_canUnfoldAtMatcher___closed__1; +lean_inc(x_31); +x_239 = l_Lean_TagAttribute_hasTag(x_238, x_237, x_31); +if (x_239 == 0) { -uint8_t x_482; lean_object* x_483; lean_object* x_484; -lean_dec(x_325); -x_482 = 1; -x_483 = lean_box(x_482); -x_484 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_484, 0, x_483); -lean_ctor_set(x_484, 1, x_420); -return x_484; -} -} -else +lean_object* x_240; uint8_t x_241; +x_240 = l_Lean_Meta_canUnfoldAtMatcher___closed__3; +x_241 = lean_name_eq(x_31, x_240); +if (x_241 == 0) { -uint8_t x_485; lean_object* x_486; lean_object* x_487; -lean_dec(x_325); -x_485 = 1; -x_486 = lean_box(x_485); -x_487 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_487, 0, x_486); -lean_ctor_set(x_487, 1, x_420); -return x_487; -} -} -else +lean_object* x_242; uint8_t x_243; +x_242 = l_Lean_Meta_canUnfoldAtMatcher___closed__5; +x_243 = lean_name_eq(x_31, x_242); +if (x_243 == 0) { -uint8_t x_488; lean_object* x_489; lean_object* x_490; -lean_dec(x_325); -x_488 = 1; -x_489 = lean_box(x_488); -x_490 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_490, 0, x_489); -lean_ctor_set(x_490, 1, x_420); -return x_490; -} -} -else +lean_object* x_244; uint8_t x_245; +x_244 = l_Lean_Meta_canUnfoldAtMatcher___closed__7; +x_245 = lean_name_eq(x_31, x_244); +if (x_245 == 0) { -uint8_t x_491; lean_object* x_492; lean_object* x_493; -lean_dec(x_325); -x_491 = 1; -x_492 = lean_box(x_491); -x_493 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_493, 0, x_492); -lean_ctor_set(x_493, 1, x_420); -return x_493; -} -} -else +lean_object* x_246; uint8_t x_247; +x_246 = l_Lean_Meta_canUnfoldAtMatcher___closed__9; +x_247 = lean_name_eq(x_31, x_246); +if (x_247 == 0) { -uint8_t x_494; lean_object* x_495; lean_object* x_496; -lean_dec(x_325); -x_494 = 1; -x_495 = lean_box(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_420); -return x_496; -} -} -else +lean_object* x_248; uint8_t x_249; +x_248 = l_Lean_Meta_canUnfoldAtMatcher___closed__12; +x_249 = lean_name_eq(x_31, x_248); +if (x_249 == 0) { -uint8_t x_497; lean_object* x_498; lean_object* x_499; -lean_dec(x_325); -x_497 = 1; -x_498 = lean_box(x_497); -x_499 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_499, 0, x_498); -lean_ctor_set(x_499, 1, x_420); -return x_499; -} -} -else +lean_object* x_250; uint8_t x_251; +x_250 = l_Lean_Meta_canUnfoldAtMatcher___closed__14; +x_251 = lean_name_eq(x_31, x_250); +if (x_251 == 0) { -uint8_t x_500; lean_object* x_501; lean_object* x_502; -lean_dec(x_325); -x_500 = 1; -x_501 = lean_box(x_500); -x_502 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_502, 0, x_501); -lean_ctor_set(x_502, 1, x_420); -return x_502; -} -} -else +lean_object* x_252; uint8_t x_253; +x_252 = l_Lean_Meta_canUnfoldAtMatcher___closed__16; +x_253 = lean_name_eq(x_31, x_252); +if (x_253 == 0) { -uint8_t x_503; lean_object* x_504; lean_object* x_505; -lean_dec(x_325); -x_503 = 1; -x_504 = lean_box(x_503); -x_505 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_505, 0, x_504); -lean_ctor_set(x_505, 1, x_420); -return x_505; -} -} -else +lean_object* x_254; uint8_t x_255; +x_254 = l_Lean_Meta_canUnfoldAtMatcher___closed__19; +x_255 = lean_name_eq(x_31, x_254); +if (x_255 == 0) { -uint8_t x_506; lean_object* x_507; lean_object* x_508; -lean_dec(x_325); -x_506 = 1; -x_507 = lean_box(x_506); -x_508 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_508, 0, x_507); -lean_ctor_set(x_508, 1, x_420); -return x_508; -} -} -else +lean_object* x_256; uint8_t x_257; +x_256 = l_Lean_Meta_canUnfoldAtMatcher___closed__21; +x_257 = lean_name_eq(x_31, x_256); +if (x_257 == 0) { -uint8_t x_509; lean_object* x_510; lean_object* x_511; -lean_dec(x_325); -x_509 = 1; -x_510 = lean_box(x_509); -x_511 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_511, 0, x_510); -lean_ctor_set(x_511, 1, x_420); -return x_511; -} -} -else +lean_object* x_258; uint8_t x_259; +x_258 = l_Lean_Meta_canUnfoldAtMatcher___closed__23; +x_259 = lean_name_eq(x_31, x_258); +if (x_259 == 0) { -uint8_t x_512; lean_object* x_513; lean_object* x_514; -lean_dec(x_325); -x_512 = 1; -x_513 = lean_box(x_512); -x_514 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_514, 0, x_513); -lean_ctor_set(x_514, 1, x_420); -return x_514; -} -} -else +lean_object* x_260; uint8_t x_261; +x_260 = l_Lean_Meta_canUnfoldAtMatcher___closed__24; +x_261 = lean_name_eq(x_31, x_260); +if (x_261 == 0) { -uint8_t x_515; lean_object* x_516; lean_object* x_517; -lean_dec(x_325); -x_515 = 1; -x_516 = lean_box(x_515); -x_517 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_517, 0, x_516); -lean_ctor_set(x_517, 1, x_420); -return x_517; -} -} -else +lean_object* x_262; uint8_t x_263; +x_262 = l_Lean_Meta_canUnfoldAtMatcher___closed__26; +x_263 = lean_name_eq(x_31, x_262); +if (x_263 == 0) { -uint8_t x_518; lean_object* x_519; lean_object* x_520; -lean_dec(x_325); -x_518 = 1; -x_519 = lean_box(x_518); -x_520 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_520, 0, x_519); -lean_ctor_set(x_520, 1, x_420); -return x_520; -} -} -} -else +lean_object* x_264; uint8_t x_265; +x_264 = l_Lean_Meta_canUnfoldAtMatcher___closed__27; +x_265 = lean_name_eq(x_31, x_264); +if (x_265 == 0) { -uint8_t x_521; lean_object* x_522; -lean_dec(x_325); -x_521 = 1; -x_522 = lean_box(x_521); -lean_ctor_set(x_329, 0, x_522); -return x_329; -} -} -else +lean_object* x_266; uint8_t x_267; +x_266 = l_Lean_Meta_canUnfoldAtMatcher___closed__29; +x_267 = lean_name_eq(x_31, x_266); +if (x_267 == 0) { -lean_object* x_523; lean_object* x_524; lean_object* x_525; uint8_t x_526; -x_523 = lean_ctor_get(x_329, 0); -x_524 = lean_ctor_get(x_329, 1); -lean_inc(x_524); -lean_inc(x_523); -lean_dec(x_329); -x_525 = lean_ctor_get(x_523, 0); -lean_inc(x_525); -lean_dec(x_523); -lean_inc(x_325); -x_526 = lean_is_instance(x_525, x_325); -if (x_526 == 0) -{ -lean_object* x_527; lean_object* x_528; lean_object* x_529; lean_object* x_530; lean_object* x_531; lean_object* x_532; uint8_t x_533; -x_527 = lean_st_ref_get(x_4, x_524); -x_528 = lean_ctor_get(x_527, 0); -lean_inc(x_528); -x_529 = lean_ctor_get(x_527, 1); -lean_inc(x_529); -if (lean_is_exclusive(x_527)) { - lean_ctor_release(x_527, 0); - lean_ctor_release(x_527, 1); - x_530 = x_527; -} else { - lean_dec_ref(x_527); - x_530 = lean_box(0); -} -x_531 = lean_ctor_get(x_528, 0); -lean_inc(x_531); -lean_dec(x_528); -x_532 = l_Lean_Meta_canUnfoldAtMatcher___closed__1; -lean_inc(x_325); -x_533 = l_Lean_TagAttribute_hasTag(x_532, x_531, x_325); -if (x_533 == 0) -{ -lean_object* x_534; uint8_t x_535; -x_534 = l_Lean_Meta_canUnfoldAtMatcher___closed__3; -x_535 = lean_name_eq(x_325, x_534); -if (x_535 == 0) -{ -lean_object* x_536; uint8_t x_537; -x_536 = l_Lean_Meta_canUnfoldAtMatcher___closed__5; -x_537 = lean_name_eq(x_325, x_536); -if (x_537 == 0) -{ -lean_object* x_538; uint8_t x_539; -x_538 = l_Lean_Meta_canUnfoldAtMatcher___closed__7; -x_539 = lean_name_eq(x_325, x_538); -if (x_539 == 0) -{ -lean_object* x_540; uint8_t x_541; -x_540 = l_Lean_Meta_canUnfoldAtMatcher___closed__9; -x_541 = lean_name_eq(x_325, x_540); -if (x_541 == 0) -{ -lean_object* x_542; uint8_t x_543; -x_542 = l_Lean_Meta_canUnfoldAtMatcher___closed__12; -x_543 = lean_name_eq(x_325, x_542); -if (x_543 == 0) -{ -lean_object* x_544; uint8_t x_545; -x_544 = l_Lean_Meta_canUnfoldAtMatcher___closed__14; -x_545 = lean_name_eq(x_325, x_544); -if (x_545 == 0) -{ -lean_object* x_546; uint8_t x_547; -x_546 = l_Lean_Meta_canUnfoldAtMatcher___closed__16; -x_547 = lean_name_eq(x_325, x_546); -if (x_547 == 0) -{ -lean_object* x_548; uint8_t x_549; -x_548 = l_Lean_Meta_canUnfoldAtMatcher___closed__19; -x_549 = lean_name_eq(x_325, x_548); -if (x_549 == 0) -{ -lean_object* x_550; uint8_t x_551; -x_550 = l_Lean_Meta_canUnfoldAtMatcher___closed__21; -x_551 = lean_name_eq(x_325, x_550); -if (x_551 == 0) -{ -lean_object* x_552; uint8_t x_553; -x_552 = l_Lean_Meta_canUnfoldAtMatcher___closed__23; -x_553 = lean_name_eq(x_325, x_552); -if (x_553 == 0) -{ -lean_object* x_554; uint8_t x_555; -x_554 = l_Lean_Meta_canUnfoldAtMatcher___closed__24; -x_555 = lean_name_eq(x_325, x_554); -if (x_555 == 0) -{ -lean_object* x_556; uint8_t x_557; -x_556 = l_Lean_Meta_canUnfoldAtMatcher___closed__26; -x_557 = lean_name_eq(x_325, x_556); -if (x_557 == 0) -{ -lean_object* x_558; uint8_t x_559; -x_558 = l_Lean_Meta_canUnfoldAtMatcher___closed__27; -x_559 = lean_name_eq(x_325, x_558); -if (x_559 == 0) -{ -lean_object* x_560; uint8_t x_561; -x_560 = l_Lean_Meta_canUnfoldAtMatcher___closed__29; -x_561 = lean_name_eq(x_325, x_560); -if (x_561 == 0) -{ -lean_object* x_562; uint8_t x_563; -x_562 = l_Lean_Meta_canUnfoldAtMatcher___closed__30; -x_563 = lean_name_eq(x_325, x_562); -if (x_563 == 0) -{ -lean_object* x_564; uint8_t x_565; -x_564 = l_Lean_Meta_canUnfoldAtMatcher___closed__32; -x_565 = lean_name_eq(x_325, x_564); -if (x_565 == 0) -{ -lean_object* x_566; uint8_t x_567; -x_566 = l_Lean_Meta_canUnfoldAtMatcher___closed__33; -x_567 = lean_name_eq(x_325, x_566); -if (x_567 == 0) -{ -lean_object* x_568; uint8_t x_569; -x_568 = l_Lean_Meta_canUnfoldAtMatcher___closed__36; -x_569 = lean_name_eq(x_325, x_568); -if (x_569 == 0) -{ -lean_object* x_570; uint8_t x_571; lean_object* x_572; lean_object* x_573; -x_570 = l_Lean_Meta_canUnfoldAtMatcher___closed__39; -x_571 = lean_name_eq(x_325, x_570); -lean_dec(x_325); -x_572 = lean_box(x_571); -if (lean_is_scalar(x_530)) { - x_573 = lean_alloc_ctor(0, 2, 0); +lean_object* x_268; uint8_t x_269; +x_268 = l_Lean_Meta_canUnfoldAtMatcher___closed__30; +x_269 = lean_name_eq(x_31, x_268); +if (x_269 == 0) +{ +lean_object* x_270; uint8_t x_271; +x_270 = l_Lean_Meta_canUnfoldAtMatcher___closed__32; +x_271 = lean_name_eq(x_31, x_270); +if (x_271 == 0) +{ +lean_object* x_272; uint8_t x_273; +x_272 = l_Lean_Meta_canUnfoldAtMatcher___closed__33; +x_273 = lean_name_eq(x_31, x_272); +if (x_273 == 0) +{ +lean_object* x_274; uint8_t x_275; +x_274 = l_Lean_Meta_canUnfoldAtMatcher___closed__36; +x_275 = lean_name_eq(x_31, x_274); +if (x_275 == 0) +{ +lean_object* x_276; uint8_t x_277; lean_object* x_278; lean_object* x_279; +x_276 = l_Lean_Meta_canUnfoldAtMatcher___closed__39; +x_277 = lean_name_eq(x_31, x_276); +lean_dec(x_31); +x_278 = lean_box(x_277); +if (lean_is_scalar(x_236)) { + x_279 = lean_alloc_ctor(0, 2, 0); } else { - x_573 = x_530; + x_279 = x_236; } -lean_ctor_set(x_573, 0, x_572); -lean_ctor_set(x_573, 1, x_529); -return x_573; +lean_ctor_set(x_279, 0, x_278); +lean_ctor_set(x_279, 1, x_235); +return x_279; } else { -uint8_t x_574; lean_object* x_575; lean_object* x_576; -lean_dec(x_325); -x_574 = 1; -x_575 = lean_box(x_574); -if (lean_is_scalar(x_530)) { - x_576 = lean_alloc_ctor(0, 2, 0); +uint8_t x_280; lean_object* x_281; lean_object* x_282; +lean_dec(x_31); +x_280 = 1; +x_281 = lean_box(x_280); +if (lean_is_scalar(x_236)) { + x_282 = lean_alloc_ctor(0, 2, 0); } else { - x_576 = x_530; + x_282 = x_236; } -lean_ctor_set(x_576, 0, x_575); -lean_ctor_set(x_576, 1, x_529); -return x_576; +lean_ctor_set(x_282, 0, x_281); +lean_ctor_set(x_282, 1, x_235); +return x_282; } } else { -uint8_t x_577; lean_object* x_578; lean_object* x_579; -lean_dec(x_325); -x_577 = 1; -x_578 = lean_box(x_577); -if (lean_is_scalar(x_530)) { - x_579 = lean_alloc_ctor(0, 2, 0); +uint8_t x_283; lean_object* x_284; lean_object* x_285; +lean_dec(x_31); +x_283 = 1; +x_284 = lean_box(x_283); +if (lean_is_scalar(x_236)) { + x_285 = lean_alloc_ctor(0, 2, 0); } else { - x_579 = x_530; + x_285 = x_236; } -lean_ctor_set(x_579, 0, x_578); -lean_ctor_set(x_579, 1, x_529); -return x_579; +lean_ctor_set(x_285, 0, x_284); +lean_ctor_set(x_285, 1, x_235); +return x_285; } } else { -uint8_t x_580; lean_object* x_581; lean_object* x_582; -lean_dec(x_325); -x_580 = 1; -x_581 = lean_box(x_580); -if (lean_is_scalar(x_530)) { - x_582 = lean_alloc_ctor(0, 2, 0); +uint8_t x_286; lean_object* x_287; lean_object* x_288; +lean_dec(x_31); +x_286 = 1; +x_287 = lean_box(x_286); +if (lean_is_scalar(x_236)) { + x_288 = lean_alloc_ctor(0, 2, 0); } else { - x_582 = x_530; + x_288 = x_236; } -lean_ctor_set(x_582, 0, x_581); -lean_ctor_set(x_582, 1, x_529); -return x_582; +lean_ctor_set(x_288, 0, x_287); +lean_ctor_set(x_288, 1, x_235); +return x_288; } } else { -uint8_t x_583; lean_object* x_584; lean_object* x_585; -lean_dec(x_325); -x_583 = 1; -x_584 = lean_box(x_583); -if (lean_is_scalar(x_530)) { - x_585 = lean_alloc_ctor(0, 2, 0); +uint8_t x_289; lean_object* x_290; lean_object* x_291; +lean_dec(x_31); +x_289 = 1; +x_290 = lean_box(x_289); +if (lean_is_scalar(x_236)) { + x_291 = lean_alloc_ctor(0, 2, 0); } else { - x_585 = x_530; + x_291 = x_236; } -lean_ctor_set(x_585, 0, x_584); -lean_ctor_set(x_585, 1, x_529); -return x_585; +lean_ctor_set(x_291, 0, x_290); +lean_ctor_set(x_291, 1, x_235); +return x_291; } } else { -uint8_t x_586; lean_object* x_587; lean_object* x_588; -lean_dec(x_325); -x_586 = 1; -x_587 = lean_box(x_586); -if (lean_is_scalar(x_530)) { - x_588 = lean_alloc_ctor(0, 2, 0); +uint8_t x_292; lean_object* x_293; lean_object* x_294; +lean_dec(x_31); +x_292 = 1; +x_293 = lean_box(x_292); +if (lean_is_scalar(x_236)) { + x_294 = lean_alloc_ctor(0, 2, 0); } else { - x_588 = x_530; + x_294 = x_236; } -lean_ctor_set(x_588, 0, x_587); -lean_ctor_set(x_588, 1, x_529); -return x_588; +lean_ctor_set(x_294, 0, x_293); +lean_ctor_set(x_294, 1, x_235); +return x_294; } } else { -uint8_t x_589; lean_object* x_590; lean_object* x_591; -lean_dec(x_325); -x_589 = 1; -x_590 = lean_box(x_589); -if (lean_is_scalar(x_530)) { - x_591 = lean_alloc_ctor(0, 2, 0); +uint8_t x_295; lean_object* x_296; lean_object* x_297; +lean_dec(x_31); +x_295 = 1; +x_296 = lean_box(x_295); +if (lean_is_scalar(x_236)) { + x_297 = lean_alloc_ctor(0, 2, 0); } else { - x_591 = x_530; + x_297 = x_236; } -lean_ctor_set(x_591, 0, x_590); -lean_ctor_set(x_591, 1, x_529); -return x_591; +lean_ctor_set(x_297, 0, x_296); +lean_ctor_set(x_297, 1, x_235); +return x_297; } } else { -uint8_t x_592; lean_object* x_593; lean_object* x_594; -lean_dec(x_325); -x_592 = 1; -x_593 = lean_box(x_592); -if (lean_is_scalar(x_530)) { - x_594 = lean_alloc_ctor(0, 2, 0); +uint8_t x_298; lean_object* x_299; lean_object* x_300; +lean_dec(x_31); +x_298 = 1; +x_299 = lean_box(x_298); +if (lean_is_scalar(x_236)) { + x_300 = lean_alloc_ctor(0, 2, 0); } else { - x_594 = x_530; + x_300 = x_236; } -lean_ctor_set(x_594, 0, x_593); -lean_ctor_set(x_594, 1, x_529); -return x_594; +lean_ctor_set(x_300, 0, x_299); +lean_ctor_set(x_300, 1, x_235); +return x_300; } } else { -uint8_t x_595; lean_object* x_596; lean_object* x_597; -lean_dec(x_325); -x_595 = 1; -x_596 = lean_box(x_595); -if (lean_is_scalar(x_530)) { - x_597 = lean_alloc_ctor(0, 2, 0); +uint8_t x_301; lean_object* x_302; lean_object* x_303; +lean_dec(x_31); +x_301 = 1; +x_302 = lean_box(x_301); +if (lean_is_scalar(x_236)) { + x_303 = lean_alloc_ctor(0, 2, 0); } else { - x_597 = x_530; + x_303 = x_236; } -lean_ctor_set(x_597, 0, x_596); -lean_ctor_set(x_597, 1, x_529); -return x_597; +lean_ctor_set(x_303, 0, x_302); +lean_ctor_set(x_303, 1, x_235); +return x_303; } } else { -uint8_t x_598; lean_object* x_599; lean_object* x_600; -lean_dec(x_325); -x_598 = 1; -x_599 = lean_box(x_598); -if (lean_is_scalar(x_530)) { - x_600 = lean_alloc_ctor(0, 2, 0); +uint8_t x_304; lean_object* x_305; lean_object* x_306; +lean_dec(x_31); +x_304 = 1; +x_305 = lean_box(x_304); +if (lean_is_scalar(x_236)) { + x_306 = lean_alloc_ctor(0, 2, 0); } else { - x_600 = x_530; + x_306 = x_236; } -lean_ctor_set(x_600, 0, x_599); -lean_ctor_set(x_600, 1, x_529); -return x_600; +lean_ctor_set(x_306, 0, x_305); +lean_ctor_set(x_306, 1, x_235); +return x_306; } } else { -uint8_t x_601; lean_object* x_602; lean_object* x_603; -lean_dec(x_325); -x_601 = 1; -x_602 = lean_box(x_601); -if (lean_is_scalar(x_530)) { - x_603 = lean_alloc_ctor(0, 2, 0); +uint8_t x_307; lean_object* x_308; lean_object* x_309; +lean_dec(x_31); +x_307 = 1; +x_308 = lean_box(x_307); +if (lean_is_scalar(x_236)) { + x_309 = lean_alloc_ctor(0, 2, 0); } else { - x_603 = x_530; + x_309 = x_236; } -lean_ctor_set(x_603, 0, x_602); -lean_ctor_set(x_603, 1, x_529); -return x_603; +lean_ctor_set(x_309, 0, x_308); +lean_ctor_set(x_309, 1, x_235); +return x_309; } } else { -uint8_t x_604; lean_object* x_605; lean_object* x_606; -lean_dec(x_325); -x_604 = 1; -x_605 = lean_box(x_604); -if (lean_is_scalar(x_530)) { - x_606 = lean_alloc_ctor(0, 2, 0); +uint8_t x_310; lean_object* x_311; lean_object* x_312; +lean_dec(x_31); +x_310 = 1; +x_311 = lean_box(x_310); +if (lean_is_scalar(x_236)) { + x_312 = lean_alloc_ctor(0, 2, 0); } else { - x_606 = x_530; + x_312 = x_236; } -lean_ctor_set(x_606, 0, x_605); -lean_ctor_set(x_606, 1, x_529); -return x_606; +lean_ctor_set(x_312, 0, x_311); +lean_ctor_set(x_312, 1, x_235); +return x_312; } } else { -uint8_t x_607; lean_object* x_608; lean_object* x_609; -lean_dec(x_325); -x_607 = 1; -x_608 = lean_box(x_607); -if (lean_is_scalar(x_530)) { - x_609 = lean_alloc_ctor(0, 2, 0); +uint8_t x_313; lean_object* x_314; lean_object* x_315; +lean_dec(x_31); +x_313 = 1; +x_314 = lean_box(x_313); +if (lean_is_scalar(x_236)) { + x_315 = lean_alloc_ctor(0, 2, 0); } else { - x_609 = x_530; + x_315 = x_236; } -lean_ctor_set(x_609, 0, x_608); -lean_ctor_set(x_609, 1, x_529); -return x_609; +lean_ctor_set(x_315, 0, x_314); +lean_ctor_set(x_315, 1, x_235); +return x_315; } } else { -uint8_t x_610; lean_object* x_611; lean_object* x_612; -lean_dec(x_325); -x_610 = 1; -x_611 = lean_box(x_610); -if (lean_is_scalar(x_530)) { - x_612 = lean_alloc_ctor(0, 2, 0); +uint8_t x_316; lean_object* x_317; lean_object* x_318; +lean_dec(x_31); +x_316 = 1; +x_317 = lean_box(x_316); +if (lean_is_scalar(x_236)) { + x_318 = lean_alloc_ctor(0, 2, 0); } else { - x_612 = x_530; + x_318 = x_236; } -lean_ctor_set(x_612, 0, x_611); -lean_ctor_set(x_612, 1, x_529); -return x_612; +lean_ctor_set(x_318, 0, x_317); +lean_ctor_set(x_318, 1, x_235); +return x_318; } } else { -uint8_t x_613; lean_object* x_614; lean_object* x_615; -lean_dec(x_325); -x_613 = 1; -x_614 = lean_box(x_613); -if (lean_is_scalar(x_530)) { - x_615 = lean_alloc_ctor(0, 2, 0); +uint8_t x_319; lean_object* x_320; lean_object* x_321; +lean_dec(x_31); +x_319 = 1; +x_320 = lean_box(x_319); +if (lean_is_scalar(x_236)) { + x_321 = lean_alloc_ctor(0, 2, 0); } else { - x_615 = x_530; + x_321 = x_236; } -lean_ctor_set(x_615, 0, x_614); -lean_ctor_set(x_615, 1, x_529); -return x_615; +lean_ctor_set(x_321, 0, x_320); +lean_ctor_set(x_321, 1, x_235); +return x_321; } } else { -uint8_t x_616; lean_object* x_617; lean_object* x_618; -lean_dec(x_325); -x_616 = 1; -x_617 = lean_box(x_616); -if (lean_is_scalar(x_530)) { - x_618 = lean_alloc_ctor(0, 2, 0); +uint8_t x_322; lean_object* x_323; lean_object* x_324; +lean_dec(x_31); +x_322 = 1; +x_323 = lean_box(x_322); +if (lean_is_scalar(x_236)) { + x_324 = lean_alloc_ctor(0, 2, 0); } else { - x_618 = x_530; + x_324 = x_236; } -lean_ctor_set(x_618, 0, x_617); -lean_ctor_set(x_618, 1, x_529); -return x_618; +lean_ctor_set(x_324, 0, x_323); +lean_ctor_set(x_324, 1, x_235); +return x_324; } } else { -uint8_t x_619; lean_object* x_620; lean_object* x_621; -lean_dec(x_325); -x_619 = 1; -x_620 = lean_box(x_619); -if (lean_is_scalar(x_530)) { - x_621 = lean_alloc_ctor(0, 2, 0); +uint8_t x_325; lean_object* x_326; lean_object* x_327; +lean_dec(x_31); +x_325 = 1; +x_326 = lean_box(x_325); +if (lean_is_scalar(x_236)) { + x_327 = lean_alloc_ctor(0, 2, 0); } else { - x_621 = x_530; + x_327 = x_236; } -lean_ctor_set(x_621, 0, x_620); -lean_ctor_set(x_621, 1, x_529); -return x_621; +lean_ctor_set(x_327, 0, x_326); +lean_ctor_set(x_327, 1, x_235); +return x_327; } } else { -uint8_t x_622; lean_object* x_623; lean_object* x_624; -lean_dec(x_325); -x_622 = 1; -x_623 = lean_box(x_622); -if (lean_is_scalar(x_530)) { - x_624 = lean_alloc_ctor(0, 2, 0); +uint8_t x_328; lean_object* x_329; lean_object* x_330; +lean_dec(x_31); +x_328 = 1; +x_329 = lean_box(x_328); +if (lean_is_scalar(x_236)) { + x_330 = lean_alloc_ctor(0, 2, 0); } else { - x_624 = x_530; + x_330 = x_236; } -lean_ctor_set(x_624, 0, x_623); -lean_ctor_set(x_624, 1, x_529); -return x_624; +lean_ctor_set(x_330, 0, x_329); +lean_ctor_set(x_330, 1, x_235); +return x_330; } } else { -uint8_t x_625; lean_object* x_626; lean_object* x_627; -lean_dec(x_325); -x_625 = 1; -x_626 = lean_box(x_625); -if (lean_is_scalar(x_530)) { - x_627 = lean_alloc_ctor(0, 2, 0); +uint8_t x_331; lean_object* x_332; lean_object* x_333; +lean_dec(x_31); +x_331 = 1; +x_332 = lean_box(x_331); +if (lean_is_scalar(x_236)) { + x_333 = lean_alloc_ctor(0, 2, 0); } else { - x_627 = x_530; + x_333 = x_236; } -lean_ctor_set(x_627, 0, x_626); -lean_ctor_set(x_627, 1, x_529); -return x_627; +lean_ctor_set(x_333, 0, x_332); +lean_ctor_set(x_333, 1, x_235); +return x_333; } } else { -uint8_t x_628; lean_object* x_629; lean_object* x_630; -lean_dec(x_325); -x_628 = 1; -x_629 = lean_box(x_628); -if (lean_is_scalar(x_530)) { - x_630 = lean_alloc_ctor(0, 2, 0); +uint8_t x_334; lean_object* x_335; lean_object* x_336; +lean_dec(x_31); +x_334 = 1; +x_335 = lean_box(x_334); +if (lean_is_scalar(x_236)) { + x_336 = lean_alloc_ctor(0, 2, 0); } else { - x_630 = x_530; + x_336 = x_236; } -lean_ctor_set(x_630, 0, x_629); -lean_ctor_set(x_630, 1, x_529); -return x_630; +lean_ctor_set(x_336, 0, x_335); +lean_ctor_set(x_336, 1, x_235); +return x_336; } } else { -uint8_t x_631; lean_object* x_632; lean_object* x_633; -lean_dec(x_325); -x_631 = 1; -x_632 = lean_box(x_631); -x_633 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_633, 0, x_632); -lean_ctor_set(x_633, 1, x_524); -return x_633; +uint8_t x_337; lean_object* x_338; lean_object* x_339; +lean_dec(x_31); +x_337 = 1; +x_338 = lean_box(x_337); +x_339 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_339, 0, x_338); +lean_ctor_set(x_339, 1, x_230); +return x_339; } } } else { -uint8_t x_634; -lean_dec(x_325); -x_634 = !lean_is_exclusive(x_329); -if (x_634 == 0) -{ -lean_object* x_635; uint8_t x_636; lean_object* x_637; -x_635 = lean_ctor_get(x_329, 0); -lean_dec(x_635); -x_636 = 1; -x_637 = lean_box(x_636); -lean_ctor_set(x_329, 0, x_637); -return x_329; +uint8_t x_340; +lean_dec(x_31); +x_340 = !lean_is_exclusive(x_35); +if (x_340 == 0) +{ +lean_object* x_341; uint8_t x_342; lean_object* x_343; +x_341 = lean_ctor_get(x_35, 0); +lean_dec(x_341); +x_342 = 1; +x_343 = lean_box(x_342); +lean_ctor_set(x_35, 0, x_343); +return x_35; } else { -lean_object* x_638; uint8_t x_639; lean_object* x_640; lean_object* x_641; -x_638 = lean_ctor_get(x_329, 1); -lean_inc(x_638); -lean_dec(x_329); -x_639 = 1; -x_640 = lean_box(x_639); -x_641 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_641, 0, x_640); -lean_ctor_set(x_641, 1, x_638); -return x_641; -} +lean_object* x_344; uint8_t x_345; lean_object* x_346; lean_object* x_347; +x_344 = lean_ctor_get(x_35, 1); +lean_inc(x_344); +lean_dec(x_35); +x_345 = 1; +x_346 = lean_box(x_345); +x_347 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_347, 0, x_346); +lean_ctor_set(x_347, 1, x_344); +return x_347; } } -default: -{ -uint8_t x_642; lean_object* x_643; lean_object* x_644; -lean_dec(x_7); -x_642 = 1; -x_643 = lean_box(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_5); -return x_644; } } } @@ -16869,7 +15759,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_12336_(x_11, x_10); +x_12 = l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_12510_(x_11, x_10); if (x_12 == 0) { lean_object* x_13; uint8_t x_14; lean_object* x_15; @@ -18565,7 +17455,7 @@ if (x_36 == 0) lean_object* x_37; lean_object* x_38; uint8_t x_39; lean_free_object(x_12); x_37 = l_Lean_Expr_mvar___override(x_26); -x_38 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_37, x_3, x_4, x_5, x_6, x_23); +x_38 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_37, x_3, x_4, x_5, x_6, x_23); x_39 = !lean_is_exclusive(x_38); if (x_39 == 0) { @@ -18686,7 +17576,7 @@ if (x_69 == 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_Expr_mvar___override(x_59); -x_71 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_70, x_3, x_4, x_5, x_6, x_57); +x_71 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_70, x_3, x_4, x_5, x_6, x_57); x_72 = lean_ctor_get(x_71, 0); lean_inc(x_72); x_73 = lean_ctor_get(x_71, 1); @@ -18796,7 +17686,7 @@ if (x_98 == 0) 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_86); x_99 = l_Lean_Expr_mvar___override(x_88); -x_100 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_99, x_3, x_4, x_5, x_6, x_85); +x_100 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_99, x_3, x_4, x_5, x_6, x_85); x_101 = lean_ctor_get(x_100, 0); lean_inc(x_101); x_102 = lean_ctor_get(x_100, 1); @@ -23219,7 +22109,7 @@ x_14 = lean_ctor_get(x_12, 0); lean_inc(x_14); lean_dec(x_12); x_38 = l_Lean_Meta_unfoldDefinition_x3f___closed__1; -x_39 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_9, x_38); +x_39 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_9, x_38); lean_dec(x_9); if (x_39 == 0) { @@ -23396,7 +22286,7 @@ x_48 = lean_ctor_get(x_46, 0); lean_inc(x_48); lean_dec(x_46); x_68 = l_Lean_Meta_unfoldDefinition_x3f___closed__1; -x_69 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_9, x_68); +x_69 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_9, x_68); lean_dec(x_9); if (x_69 == 0) { @@ -23634,7 +22524,7 @@ lean_object* x_95; lean_object* x_96; uint8_t x_97; x_95 = lean_ctor_get(x_4, 2); lean_inc(x_95); x_96 = l_Lean_Meta_unfoldDefinition_x3f___closed__1; -x_97 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_95, x_96); +x_97 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_95, x_96); lean_dec(x_95); if (x_97 == 0) { @@ -23997,7 +22887,7 @@ lean_object* x_182; lean_object* x_183; uint8_t x_184; x_182 = lean_ctor_get(x_4, 2); lean_inc(x_182); x_183 = l_Lean_Meta_unfoldDefinition_x3f___closed__1; -x_184 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_182, x_183); +x_184 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_182, x_183); lean_dec(x_182); if (x_184 == 0) { @@ -24324,7 +23214,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_12336_(x_12, x_11); +x_13 = l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_12510_(x_12, x_11); if (x_13 == 0) { lean_object* x_14; @@ -24355,7 +23245,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_12336_(x_19, x_18); +x_20 = l___private_Init_Meta_0__Lean_Meta_beqTransparencyMode____x40_Init_Meta___hyg_12510_(x_19, x_18); if (x_20 == 0) { lean_object* x_21; lean_object* x_22; @@ -35743,7 +34633,7 @@ return x_38; } } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__1() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -35753,17 +34643,17 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__2() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__1; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__1; x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_36____closed__7; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__3() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__3() { _start: { lean_object* x_1; @@ -35771,17 +34661,17 @@ x_1 = lean_mk_string_from_bytes("initFn", 6); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__4() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__2; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__3; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__2; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__3; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__5() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__5() { _start: { lean_object* x_1; @@ -35789,37 +34679,37 @@ x_1 = lean_mk_string_from_bytes("_@", 2); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__6() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__4; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__5; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__4; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__7() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__6; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__6; x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_36____closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__8() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__7; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__7; x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_36____closed__7; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__9() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__9() { _start: { lean_object* x_1; @@ -35827,17 +34717,17 @@ x_1 = lean_mk_string_from_bytes("WHNF", 4); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__10() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__8; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__9; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__8; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__9; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__11() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__11() { _start: { lean_object* x_1; @@ -35845,33 +34735,33 @@ x_1 = lean_mk_string_from_bytes("_hyg", 4); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__12() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__10; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__11; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__10; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__11; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__13() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__12; -x_2 = lean_unsigned_to_nat(10682u); +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__12; +x_2 = lean_unsigned_to_nat(10701u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; x_2 = l_Lean_Meta_whnfEasyCases___at_Lean_Meta_whnfCore_go___spec__3___closed__2; x_3 = 0; -x_4 = l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__13; +x_4 = l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__13; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); if (lean_obj_tag(x_5) == 0) { @@ -36209,33 +35099,33 @@ l___private_Lean_Meta_WHNF_0__Lean_Meta_cache___closed__1 = _init_l___private_Le lean_mark_persistent(l___private_Lean_Meta_WHNF_0__Lean_Meta_cache___closed__1); l___private_Lean_Meta_WHNF_0__Lean_Meta_cache___closed__2 = _init_l___private_Lean_Meta_WHNF_0__Lean_Meta_cache___closed__2(); lean_mark_persistent(l___private_Lean_Meta_WHNF_0__Lean_Meta_cache___closed__2); -l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__1 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__1(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__1); -l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__2 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__2(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__2); -l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__3 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__3(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__3); -l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__4 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__4(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__4); -l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__5 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__5(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__5); -l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__6 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__6(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__6); -l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__7 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__7(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__7); -l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__8 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__8(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__8); -l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__9 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__9(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__9); -l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__10 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__10(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__10); -l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__11 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__11(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__11); -l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__12 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__12(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__12); -l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__13 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__13(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682____closed__13); -res = l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10682_(lean_io_mk_world()); +l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__1 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__1(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__1); +l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__2 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__2(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__2); +l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__3 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__3(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__3); +l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__4 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__4(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__4); +l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__5 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__5(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__5); +l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__6 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__6(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__6); +l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__7 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__7(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__7); +l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__8 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__8(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__8); +l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__9 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__9(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__9); +l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__10 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__10(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__10); +l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__11 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__11(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__11); +l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__12 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__12(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__12); +l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__13 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__13(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701____closed__13); +res = l_Lean_Meta_initFn____x40_Lean_Meta_WHNF___hyg_10701_(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/MetavarContext.c b/stage0/stdlib/Lean/MetavarContext.c index cec924fceb1c..9f06cbe6b405 100644 --- a/stage0/stdlib/Lean/MetavarContext.c +++ b/stage0/stdlib/Lean/MetavarContext.c @@ -12793,7 +12793,7 @@ static lean_object* _init_l_Lean_instantiateExprMVars___rarg___lambda__9___close 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_instantiateExprMVars___rarg___lambda__9___closed__1; x_2 = l_Lean_instantiateExprMVars___rarg___lambda__9___closed__2; -x_3 = lean_unsigned_to_nat(1492u); +x_3 = lean_unsigned_to_nat(1496u); x_4 = lean_unsigned_to_nat(14u); x_5 = l_Lean_instantiateExprMVars___rarg___lambda__9___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -12874,7 +12874,7 @@ static lean_object* _init_l_Lean_instantiateExprMVars___rarg___lambda__10___clos 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_instantiateExprMVars___rarg___lambda__9___closed__1; x_2 = l_Lean_instantiateExprMVars___rarg___lambda__10___closed__1; -x_3 = lean_unsigned_to_nat(1481u); +x_3 = lean_unsigned_to_nat(1485u); x_4 = lean_unsigned_to_nat(18u); x_5 = l_Lean_instantiateExprMVars___rarg___lambda__10___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -12956,7 +12956,7 @@ static lean_object* _init_l_Lean_instantiateExprMVars___rarg___lambda__11___clos 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_instantiateExprMVars___rarg___lambda__9___closed__1; x_2 = l_Lean_instantiateExprMVars___rarg___lambda__11___closed__1; -x_3 = lean_unsigned_to_nat(1560u); +x_3 = lean_unsigned_to_nat(1564u); x_4 = lean_unsigned_to_nat(20u); x_5 = l_Lean_instantiateExprMVars___rarg___lambda__11___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -13094,7 +13094,7 @@ static lean_object* _init_l_Lean_instantiateExprMVars___rarg___lambda__13___clos 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_instantiateExprMVars___rarg___lambda__9___closed__1; x_2 = l_Lean_instantiateExprMVars___rarg___lambda__13___closed__1; -x_3 = lean_unsigned_to_nat(1540u); +x_3 = lean_unsigned_to_nat(1544u); x_4 = lean_unsigned_to_nat(24u); x_5 = l_Lean_instantiateExprMVars___rarg___lambda__13___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -13232,7 +13232,7 @@ static lean_object* _init_l_Lean_instantiateExprMVars___rarg___lambda__15___clos 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_instantiateExprMVars___rarg___lambda__9___closed__1; x_2 = l_Lean_instantiateExprMVars___rarg___lambda__15___closed__1; -x_3 = lean_unsigned_to_nat(1569u); +x_3 = lean_unsigned_to_nat(1573u); x_4 = lean_unsigned_to_nat(22u); x_5 = l_Lean_instantiateExprMVars___rarg___lambda__15___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -13400,7 +13400,7 @@ static lean_object* _init_l_Lean_instantiateExprMVars___rarg___lambda__18___clos 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_instantiateExprMVars___rarg___lambda__9___closed__1; x_2 = l_Lean_instantiateExprMVars___rarg___lambda__18___closed__1; -x_3 = lean_unsigned_to_nat(1503u); +x_3 = lean_unsigned_to_nat(1507u); x_4 = lean_unsigned_to_nat(17u); x_5 = l_Lean_instantiateExprMVars___rarg___lambda__18___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -13484,7 +13484,7 @@ static lean_object* _init_l_Lean_instantiateExprMVars___rarg___lambda__19___clos 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_instantiateExprMVars___rarg___lambda__9___closed__1; x_2 = l_Lean_instantiateExprMVars___rarg___lambda__19___closed__1; -x_3 = lean_unsigned_to_nat(1514u); +x_3 = lean_unsigned_to_nat(1518u); x_4 = lean_unsigned_to_nat(18u); x_5 = l_Lean_instantiateExprMVars___rarg___lambda__19___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); diff --git a/stage0/stdlib/Lean/Parser.c b/stage0/stdlib/Lean/Parser.c index 912f79eee6e5..5cfd8a2fa473 100644 --- a/stage0/stdlib/Lean/Parser.c +++ b/stage0/stdlib/Lean/Parser.c @@ -19,6 +19,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__215; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__201; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__199; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__294; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__84; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__122; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__236; @@ -102,6 +103,7 @@ static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__4 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_Parser_Term_scientific_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__214; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__291; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__5; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__112; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__19; @@ -113,6 +115,7 @@ lean_object* l_Lean_Parser_optional(lean_object*); 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_object* l_Lean_Parser_nameLit_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_interpretParserDescr___elambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Parser_hygieneInfo_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__181; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__55; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_interpretParserDescr___elambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -130,6 +133,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__269; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__148; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__60; +extern lean_object* l_Lean_Parser_hygieneInfo; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_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__24; lean_object* l_Lean_PrettyPrinter_Parenthesizer_checkColEq_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -320,6 +324,7 @@ lean_object* l_Lean_Parser_charLit_parenthesizer(lean_object*, lean_object*, lea lean_object* l_Lean_Parser_getBinaryAlias___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__98; LEAN_EXPORT lean_object* lean_mk_antiquot_formatter(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Parser_hygieneInfo_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Parenthesizer_ident_parenthesizer(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_interpretParserDescr___elambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_andthen(lean_object*, lean_object*); @@ -335,6 +340,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__165; lean_object* l_Lean_Parser_atomic(lean_object*); lean_object* l_Lean_Parser_ident_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__299; extern lean_object* l_Lean_Parser_charLit; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__286; extern lean_object* l_Lean_PrettyPrinter_Formatter_formatterAliasesRef; @@ -352,6 +358,7 @@ extern lean_object* l_Lean_PrettyPrinter_combinatorFormatterAttribute; lean_object* l_Lean_Parser_getUnaryAlias___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__175; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__124; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__297; 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__117; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_numLit_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -372,6 +379,7 @@ static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__8 static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__272; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__125; lean_object* l_Lean_PrettyPrinter_Formatter_checkLineEq_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__295; lean_object* l_Lean_Parser_withoutForbidden_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__268; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__220; @@ -435,6 +443,7 @@ lean_object* l_Lean_Parser_checkWsBefore(lean_object*); static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__277; 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_initFn____x40_Lean_Parser___hyg_8____closed__119; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__296; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__193; lean_object* l_Lean_Parser_lookahead(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_interpretParserDescr___elambda__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -442,6 +451,7 @@ static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__4 static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__134; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__93; lean_object* l_Lean_PrettyPrinter_Parenthesizer_notFollowedBy_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__292; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Formatter_charLit_formatter___closed__1; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__205; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__237; @@ -494,6 +504,7 @@ lean_object* l_Lean_Parser_withoutPosition(lean_object*); static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__44; lean_object* l_Lean_Parser_getConstAlias___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__223; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__298; lean_object* l_Lean_Parser_withoutForbidden_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__76; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__38; @@ -505,6 +516,7 @@ lean_object* l_Lean_Parser_scientificLit_parenthesizer(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*); lean_object* l_Lean_PrettyPrinter_Parenthesizer_checkLineEq_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*); 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__293; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__140; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_interpretParserDescr___elambda__9(lean_object*, uint8_t, 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__114; @@ -1493,7 +1505,7 @@ static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____clo _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("colGt", 5); +x_1 = lean_mk_string_from_bytes("hygieneInfo", 11); return x_1; } } @@ -1510,52 +1522,137 @@ return x_3; static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__106() { _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___hyg_8____closed__3; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__4; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__104; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +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_hygieneInfo; +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__108() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____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___hyg_8____closed__109() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Parser_hygieneInfo_formatter), 5, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__110() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__109; +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() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Parser_hygieneInfo_parenthesizer), 5, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__112() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____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___hyg_8____closed__113() { +_start: +{ +lean_object* x_1; +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__114() { +_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__113; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__115() { +_start: +{ lean_object* x_1; 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__107() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__116() { _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___hyg_8____closed__3; x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__4; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__106; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__115; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -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__117() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__106; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__115; x_2 = l_Lean_Parser_checkColGt(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__118() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__108; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__117; 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__110() { +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__107; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__116; 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__111() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__120() { _start: { lean_object* x_1; @@ -1563,17 +1660,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__112() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__121() { _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__120; 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__113() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__122() { _start: { lean_object* x_1; @@ -1581,17 +1678,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__114() { +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__113; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__122; 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__124() { _start: { lean_object* x_1; @@ -1599,17 +1696,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__116() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__125() { _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__115; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__124; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -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__126() { _start: { lean_object* x_1; @@ -1617,47 +1714,47 @@ 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__118() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__127() { _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___hyg_8____closed__3; x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__4; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__117; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__126; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -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__128() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__117; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__126; x_2 = l_Lean_Parser_checkColGe(x_1); return x_2; } } -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__129() { _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__128; 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__121() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__130() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__118; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__127; 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__122() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__131() { _start: { lean_object* x_1; @@ -1665,17 +1762,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__123() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__132() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__122; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__131; 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__124() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__133() { _start: { lean_object* x_1; @@ -1683,17 +1780,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__125() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__134() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__124; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__133; 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__126() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__135() { _start: { lean_object* x_1; @@ -1701,17 +1798,17 @@ x_1 = lean_mk_string_from_bytes("colEq", 5); return x_1; } } -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__136() { _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__126; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__135; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -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__137() { _start: { lean_object* x_1; @@ -1719,47 +1816,47 @@ x_1 = lean_mk_string_from_bytes("checkColEq", 10); return x_1; } } -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__138() { _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___hyg_8____closed__3; x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__4; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__128; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__137; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -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__139() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__128; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__137; x_2 = l_Lean_Parser_checkColEq(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__140() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__130; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__139; 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__132() { +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__129; +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__133() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__142() { _start: { lean_object* x_1; @@ -1767,17 +1864,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_checkColEq_forma 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__143() { _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__142; 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__135() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__144() { _start: { lean_object* x_1; @@ -1785,17 +1882,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_checkColEq_p 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__145() { _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__144; 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__137() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__146() { _start: { lean_object* x_1; @@ -1803,17 +1900,17 @@ x_1 = lean_mk_string_from_bytes("lineEq", 6); 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__147() { _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__146; x_3 = l_Lean_Name_str___override(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__148() { _start: { lean_object* x_1; @@ -1821,47 +1918,47 @@ x_1 = lean_mk_string_from_bytes("checkLineEq", 11); 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__149() { _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___hyg_8____closed__3; x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__4; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__139; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__148; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -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__150() { _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__148; x_2 = l_Lean_Parser_checkLineEq(x_1); return x_2; } } -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__151() { _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__150; 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__143() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__152() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__140; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__149; 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__144() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__153() { _start: { lean_object* x_1; @@ -1869,17 +1966,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_checkLineEq_form return x_1; } } -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__154() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__144; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__153; 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__146() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__155() { _start: { lean_object* x_1; @@ -1887,17 +1984,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_checkLineEq_ return x_1; } } -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__156() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__146; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__155; 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__148() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__157() { _start: { lean_object* x_1; @@ -1905,28 +2002,28 @@ 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__149() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__158() { _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__148; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__157; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -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__159() { _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___hyg_8____closed__3; x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__4; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__148; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__157; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -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__160() { _start: { lean_object* x_1; @@ -1934,27 +2031,27 @@ 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__152() { +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__151; +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__153() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__162() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__150; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__159; 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__154() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__163() { _start: { lean_object* x_1; @@ -1962,17 +2059,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__155() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__164() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__154; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__163; 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__156() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__165() { _start: { lean_object* x_1; @@ -1980,17 +2077,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__157() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__166() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__156; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__165; 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__158() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__167() { _start: { lean_object* x_1; @@ -1998,28 +2095,28 @@ 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__159() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__168() { _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__158; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__167; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -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__169() { _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___hyg_8____closed__3; x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__4; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__158; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__167; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -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__170() { _start: { lean_object* x_1; @@ -2027,27 +2124,27 @@ 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__162() { +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__161; +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__163() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__172() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__160; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__169; 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__164() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__173() { _start: { lean_object* x_1; lean_object* x_2; uint8_t x_3; lean_object* x_4; @@ -2061,7 +2158,7 @@ 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___hyg_8____closed__165() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__174() { _start: { lean_object* x_1; @@ -2069,17 +2166,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_atomic_formatter), 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__175() { _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__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__167() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__176() { _start: { lean_object* x_1; @@ -2087,17 +2184,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__168() { +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__167; +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__169() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__178() { _start: { lean_object* x_1; @@ -2105,28 +2202,28 @@ 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__170() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__179() { _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__169; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__178; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -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__180() { _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___hyg_8____closed__3; x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__4; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__169; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__178; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -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__181() { _start: { lean_object* x_1; @@ -2134,27 +2231,27 @@ 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__173() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__182() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__172; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__181; 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__174() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__183() { _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__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__175() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__184() { _start: { lean_object* x_1; @@ -2162,17 +2259,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__176() { +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__175; +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__177() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__186() { _start: { lean_object* x_1; @@ -2180,17 +2277,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__178() { +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__177; +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__179() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__188() { _start: { lean_object* x_1; @@ -2198,28 +2295,28 @@ 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__180() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__189() { _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__179; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__188; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -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__190() { _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___hyg_8____closed__3; x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__4; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__179; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__188; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -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__191() { _start: { lean_object* x_1; @@ -2227,27 +2324,27 @@ 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__183() { +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__182; +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__184() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__193() { _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__190; 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__194() { _start: { lean_object* x_1; @@ -2255,17 +2352,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__186() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__195() { _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__194; 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__196() { _start: { lean_object* x_1; @@ -2273,17 +2370,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__188() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__197() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__187; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__196; 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__189() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__198() { _start: { lean_object* x_1; @@ -2291,28 +2388,28 @@ 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__190() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__199() { _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__189; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__198; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -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__200() { _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___hyg_8____closed__3; x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__4; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__189; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__198; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -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__201() { _start: { lean_object* x_1; @@ -2320,27 +2417,27 @@ 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__193() { +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__192; +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__194() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__203() { _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__200; 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__204() { _start: { lean_object* x_1; @@ -2348,17 +2445,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__196() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__205() { _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__204; 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__206() { _start: { lean_object* x_1; @@ -2366,17 +2463,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__198() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__207() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__197; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__206; 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__199() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__208() { _start: { lean_object* x_1; @@ -2384,28 +2481,28 @@ 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__200() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__209() { _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__199; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__208; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -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__210() { _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___hyg_8____closed__3; x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__4; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__199; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__208; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -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__211() { _start: { lean_object* x_1; @@ -2413,27 +2510,27 @@ 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__203() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__212() { _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__211; 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__213() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__201; +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__205() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__214() { _start: { lean_object* x_1; @@ -2441,17 +2538,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__206() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__215() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__205; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__214; 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__207() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__216() { _start: { lean_object* x_1; @@ -2459,17 +2556,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__208() { +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__207; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__216; 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__209() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__218() { _start: { lean_object* x_1; @@ -2477,28 +2574,28 @@ 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__210() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__219() { _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__209; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__218; x_3 = l_Lean_Name_str___override(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__220() { _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___hyg_8____closed__3; x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__4; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__209; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__218; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -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__221() { _start: { lean_object* x_1; @@ -2506,27 +2603,27 @@ 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__213() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__222() { _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__221; 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__214() { +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__211; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__220; 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__215() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__224() { _start: { lean_object* x_1; lean_object* x_2; uint8_t x_3; lean_object* x_4; @@ -2540,7 +2637,7 @@ 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___hyg_8____closed__216() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__225() { _start: { lean_object* x_1; @@ -2548,17 +2645,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__217() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__226() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__216; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__225; 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__218() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__227() { _start: { lean_object* x_1; @@ -2566,17 +2663,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__219() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__228() { _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__227; 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__220() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__229() { _start: { lean_object* x_1; @@ -2584,28 +2681,28 @@ 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__221() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__230() { _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__220; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__229; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -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__231() { _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___hyg_8____closed__3; x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__4; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__220; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__229; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -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__232() { _start: { lean_object* x_1; @@ -2613,27 +2710,27 @@ 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__224() { +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__223; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__232; 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__225() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__234() { _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__231; 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__226() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__235() { _start: { lean_object* x_1; @@ -2641,17 +2738,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_withPosition_formatter), 6, 0); 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__236() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__226; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__235; 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__228() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__237() { _start: { lean_object* x_1; @@ -2659,17 +2756,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__229() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__238() { _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__237; 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__239() { _start: { lean_object* x_1; @@ -2677,28 +2774,28 @@ x_1 = lean_mk_string_from_bytes("withoutPosition", 15); 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__240() { _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__230; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__239; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__232() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__241() { _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___hyg_8____closed__3; x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__4; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__230; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__239; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__233() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__242() { _start: { lean_object* x_1; @@ -2706,27 +2803,27 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_withoutPosition), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__234() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__243() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__233; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__242; 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__235() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__244() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__232; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__241; 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__236() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__245() { _start: { lean_object* x_1; @@ -2734,17 +2831,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_withoutPosition_formatter), 6, 0) return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__237() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__246() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__236; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__245; 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__238() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__247() { _start: { lean_object* x_1; @@ -2752,17 +2849,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_withoutPosition_parenthesizer), 6 return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__239() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__248() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__238; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__247; 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__240() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__249() { _start: { lean_object* x_1; @@ -2770,28 +2867,28 @@ x_1 = lean_mk_string_from_bytes("withoutForbidden", 16); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__241() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__250() { _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__240; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__249; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__242() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__251() { _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___hyg_8____closed__3; x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__4; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__240; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__249; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__243() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__252() { _start: { lean_object* x_1; @@ -2799,27 +2896,27 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_withoutForbidden), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__244() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__253() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__243; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__252; 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__245() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__254() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__242; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__251; 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__246() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__255() { _start: { lean_object* x_1; @@ -2827,17 +2924,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_withoutForbidden_formatter), 6, 0 return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__247() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__256() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__246; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__255; 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__248() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__257() { _start: { lean_object* x_1; @@ -2845,17 +2942,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_withoutForbidden_parenthesizer), return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__249() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__258() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__248; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__257; 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__250() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__259() { _start: { lean_object* x_1; @@ -2863,28 +2960,28 @@ 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__251() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__260() { _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__250; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__259; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__252() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__261() { _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___hyg_8____closed__3; x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__4; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__250; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__259; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__253() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__262() { _start: { lean_object* x_1; @@ -2892,17 +2989,17 @@ 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__254() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__263() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__253; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__262; 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__255() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__264() { _start: { lean_object* x_1; @@ -2910,27 +3007,27 @@ 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__256() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__265() { _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__255; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__264; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__257() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__266() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__256; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__265; 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__258() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__267() { _start: { lean_object* x_1; @@ -2938,17 +3035,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__259() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__268() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__258; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__267; 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__260() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__269() { _start: { lean_object* x_1; @@ -2956,17 +3053,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__261() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__270() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__260; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__269; 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__262() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__271() { _start: { lean_object* x_1; @@ -2974,28 +3071,28 @@ 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__263() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__272() { _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__262; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__271; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__264() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__273() { _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___hyg_8____closed__3; x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__4; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__262; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__271; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__265() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__274() { _start: { lean_object* x_1; @@ -3003,27 +3100,27 @@ 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__266() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__275() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__265; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__274; 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__267() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__276() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__264; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__273; 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__268() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__277() { _start: { lean_object* x_1; @@ -3031,17 +3128,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__269() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__278() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__268; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__277; 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__270() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__279() { _start: { lean_object* x_1; @@ -3049,17 +3146,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__271() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__280() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__270; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__279; 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__272() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__281() { _start: { lean_object* x_1; @@ -3067,28 +3164,28 @@ 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__273() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__282() { _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__272; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__281; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__274() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__283() { _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___hyg_8____closed__3; x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__4; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__272; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__281; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__275() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__284() { _start: { lean_object* x_1; @@ -3096,27 +3193,27 @@ 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__276() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__285() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__275; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__284; 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__277() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__286() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__274; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__283; 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__278() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__287() { _start: { lean_object* x_1; @@ -3124,17 +3221,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__279() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__288() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__278; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__287; 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__280() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__289() { _start: { lean_object* x_1; @@ -3142,17 +3239,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__281() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__290() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__280; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__289; 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__282() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__291() { _start: { lean_object* x_1; @@ -3160,28 +3257,28 @@ 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__283() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__292() { _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__282; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__291; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__284() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__293() { _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___hyg_8____closed__3; x_2 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__4; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__282; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__291; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__285() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__294() { _start: { lean_object* x_1; @@ -3189,17 +3286,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8 return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__286() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__295() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__285; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__294; 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__287() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__296() { _start: { lean_object* x_1; @@ -3207,17 +3304,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__288() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__297() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__287; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__296; 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__289() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__298() { _start: { lean_object* x_1; @@ -3225,11 +3322,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__290() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__299() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__289; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__298; x_2 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; @@ -3487,17 +3584,17 @@ 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__105; -x_115 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__107; -x_116 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__109; -x_117 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__110; -x_118 = l_Lean_Parser_registerAlias(x_114, x_115, x_116, x_117, x_6, x_113); +x_115 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__106; +x_116 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__107; +x_117 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__108; +x_118 = l_Lean_Parser_registerAlias(x_114, x_115, x_116, x_117, x_45, x_113); if (lean_obj_tag(x_118) == 0) { lean_object* x_119; lean_object* x_120; lean_object* x_121; x_119 = lean_ctor_get(x_118, 1); lean_inc(x_119); lean_dec(x_118); -x_120 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__112; +x_120 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__110; x_121 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_114, x_120, x_119); if (lean_obj_tag(x_121) == 0) { @@ -3505,7 +3602,7 @@ lean_object* x_122; lean_object* x_123; lean_object* x_124; x_122 = lean_ctor_get(x_121, 1); lean_inc(x_122); lean_dec(x_121); -x_123 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__114; +x_123 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__112; x_124 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_114, x_123, x_122); if (lean_obj_tag(x_124) == 0) { @@ -3513,10 +3610,10 @@ lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; 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___hyg_8____closed__116; -x_127 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__118; -x_128 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__120; -x_129 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__121; +x_126 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__114; +x_127 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__116; +x_128 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__118; +x_129 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__119; x_130 = l_Lean_Parser_registerAlias(x_126, x_127, x_128, x_129, x_6, x_125); if (lean_obj_tag(x_130) == 0) { @@ -3524,7 +3621,7 @@ lean_object* x_131; lean_object* x_132; lean_object* x_133; x_131 = lean_ctor_get(x_130, 1); lean_inc(x_131); lean_dec(x_130); -x_132 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__123; +x_132 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__121; x_133 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_126, x_132, x_131); if (lean_obj_tag(x_133) == 0) { @@ -3532,7 +3629,7 @@ lean_object* x_134; lean_object* x_135; lean_object* x_136; x_134 = lean_ctor_get(x_133, 1); lean_inc(x_134); lean_dec(x_133); -x_135 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__125; +x_135 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__123; x_136 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_126, x_135, x_134); if (lean_obj_tag(x_136) == 0) { @@ -3540,10 +3637,10 @@ lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; x_137 = lean_ctor_get(x_136, 1); lean_inc(x_137); lean_dec(x_136); -x_138 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__127; -x_139 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__129; -x_140 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__131; -x_141 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__132; +x_138 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__125; +x_139 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__127; +x_140 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__129; +x_141 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__130; x_142 = l_Lean_Parser_registerAlias(x_138, x_139, x_140, x_141, x_6, x_137); if (lean_obj_tag(x_142) == 0) { @@ -3551,7 +3648,7 @@ 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__134; +x_144 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__132; x_145 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_138, x_144, x_143); if (lean_obj_tag(x_145) == 0) { @@ -3559,7 +3656,7 @@ lean_object* x_146; lean_object* x_147; lean_object* x_148; 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__136; +x_147 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__134; x_148 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_138, x_147, x_146); if (lean_obj_tag(x_148) == 0) { @@ -3567,10 +3664,10 @@ lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; x_149 = lean_ctor_get(x_148, 1); lean_inc(x_149); lean_dec(x_148); -x_150 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__138; -x_151 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__140; -x_152 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__142; -x_153 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__143; +x_150 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__136; +x_151 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__138; +x_152 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__140; +x_153 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__141; x_154 = l_Lean_Parser_registerAlias(x_150, x_151, x_152, x_153, x_6, x_149); if (lean_obj_tag(x_154) == 0) { @@ -3578,7 +3675,7 @@ 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__145; +x_156 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__143; x_157 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_150, x_156, x_155); if (lean_obj_tag(x_157) == 0) { @@ -3586,7 +3683,7 @@ 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__147; +x_159 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__145; x_160 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_150, x_159, x_158); if (lean_obj_tag(x_160) == 0) { @@ -3594,10 +3691,10 @@ lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; 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__149; -x_163 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__150; -x_164 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__152; -x_165 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__153; +x_162 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__147; +x_163 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__149; +x_164 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__151; +x_165 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__152; x_166 = l_Lean_Parser_registerAlias(x_162, x_163, x_164, x_165, x_6, x_161); if (lean_obj_tag(x_166) == 0) { @@ -3605,7 +3702,7 @@ lean_object* x_167; lean_object* x_168; lean_object* x_169; x_167 = lean_ctor_get(x_166, 1); lean_inc(x_167); lean_dec(x_166); -x_168 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__155; +x_168 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__154; x_169 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_162, x_168, x_167); if (lean_obj_tag(x_169) == 0) { @@ -3613,74 +3710,74 @@ lean_object* x_170; lean_object* x_171; lean_object* x_172; x_170 = lean_ctor_get(x_169, 1); lean_inc(x_170); lean_dec(x_169); -x_171 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__157; +x_171 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__156; x_172 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_162, x_171, x_170); if (lean_obj_tag(x_172) == 0) { -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_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; x_173 = lean_ctor_get(x_172, 1); lean_inc(x_173); lean_dec(x_172); -x_174 = lean_box(0); +x_174 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__158; x_175 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__159; -x_176 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__160; +x_176 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__161; x_177 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__162; -x_178 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__163; -x_179 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__164; -x_180 = l_Lean_Parser_registerAlias(x_175, x_176, x_177, x_178, x_179, x_173); -if (lean_obj_tag(x_180) == 0) -{ -lean_object* x_181; lean_object* x_182; lean_object* x_183; -x_181 = lean_ctor_get(x_180, 1); -lean_inc(x_181); -lean_dec(x_180); -x_182 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__166; -x_183 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_175, x_182, x_181); -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__168; -x_186 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_175, x_185, x_184); -if (lean_obj_tag(x_186) == 0) -{ -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_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__170; +x_178 = l_Lean_Parser_registerAlias(x_174, x_175, x_176, x_177, x_6, x_173); +if (lean_obj_tag(x_178) == 0) +{ +lean_object* x_179; lean_object* x_180; lean_object* x_181; +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__164; +x_181 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_174, x_180, x_179); +if (lean_obj_tag(x_181) == 0) +{ +lean_object* x_182; lean_object* x_183; lean_object* x_184; +x_182 = lean_ctor_get(x_181, 1); +lean_inc(x_182); +lean_dec(x_181); +x_183 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__166; +x_184 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_174, x_183, x_182); +if (lean_obj_tag(x_184) == 0) +{ +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_185 = lean_ctor_get(x_184, 1); +lean_inc(x_185); +lean_dec(x_184); +x_186 = lean_box(0); +x_187 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__168; +x_188 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__169; x_189 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__171; -x_190 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__173; -x_191 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__174; -x_192 = l_Lean_Parser_registerAlias(x_188, x_189, x_190, x_191, x_45, x_187); +x_190 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__172; +x_191 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__173; +x_192 = l_Lean_Parser_registerAlias(x_187, x_188, x_189, x_190, x_191, x_185); if (lean_obj_tag(x_192) == 0) { lean_object* x_193; lean_object* x_194; lean_object* x_195; x_193 = lean_ctor_get(x_192, 1); lean_inc(x_193); lean_dec(x_192); -x_194 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__176; -x_195 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_188, x_194, x_193); +x_194 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__175; +x_195 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_187, x_194, x_193); if (lean_obj_tag(x_195) == 0) { lean_object* x_196; lean_object* x_197; lean_object* x_198; x_196 = lean_ctor_get(x_195, 1); lean_inc(x_196); lean_dec(x_195); -x_197 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__178; -x_198 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_188, x_197, x_196); +x_197 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__177; +x_198 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_187, x_197, x_196); if (lean_obj_tag(x_198) == 0) { 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_199 = lean_ctor_get(x_198, 1); lean_inc(x_199); lean_dec(x_198); -x_200 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__180; -x_201 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__181; -x_202 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__183; -x_203 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__184; +x_200 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__179; +x_201 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__180; +x_202 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__182; +x_203 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__183; x_204 = l_Lean_Parser_registerAlias(x_200, x_201, x_202, x_203, x_45, x_199); if (lean_obj_tag(x_204) == 0) { @@ -3688,7 +3785,7 @@ lean_object* x_205; lean_object* x_206; lean_object* x_207; x_205 = lean_ctor_get(x_204, 1); lean_inc(x_205); lean_dec(x_204); -x_206 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__186; +x_206 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__185; x_207 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_200, x_206, x_205); if (lean_obj_tag(x_207) == 0) { @@ -3696,7 +3793,7 @@ lean_object* x_208; lean_object* x_209; lean_object* x_210; x_208 = lean_ctor_get(x_207, 1); lean_inc(x_208); lean_dec(x_207); -x_209 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__188; +x_209 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__187; x_210 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_200, x_209, x_208); if (lean_obj_tag(x_210) == 0) { @@ -3704,10 +3801,10 @@ lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; 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__190; -x_213 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__191; -x_214 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__193; -x_215 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__194; +x_212 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__189; +x_213 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__190; +x_214 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__192; +x_215 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__193; x_216 = l_Lean_Parser_registerAlias(x_212, x_213, x_214, x_215, x_45, x_211); if (lean_obj_tag(x_216) == 0) { @@ -3715,7 +3812,7 @@ lean_object* x_217; lean_object* x_218; lean_object* x_219; 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__196; +x_218 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__195; x_219 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_212, x_218, x_217); if (lean_obj_tag(x_219) == 0) { @@ -3723,7 +3820,7 @@ 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__198; +x_221 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__197; x_222 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_212, x_221, x_220); if (lean_obj_tag(x_222) == 0) { @@ -3731,10 +3828,10 @@ lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; 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__200; -x_225 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__201; -x_226 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__203; -x_227 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__204; +x_224 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__199; +x_225 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__200; +x_226 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__202; +x_227 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__203; x_228 = l_Lean_Parser_registerAlias(x_224, x_225, x_226, x_227, x_45, x_223); if (lean_obj_tag(x_228) == 0) { @@ -3742,7 +3839,7 @@ lean_object* x_229; lean_object* x_230; lean_object* x_231; x_229 = lean_ctor_get(x_228, 1); lean_inc(x_229); lean_dec(x_228); -x_230 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__206; +x_230 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__205; x_231 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_224, x_230, x_229); if (lean_obj_tag(x_231) == 0) { @@ -3750,81 +3847,81 @@ lean_object* x_232; lean_object* x_233; lean_object* x_234; x_232 = lean_ctor_get(x_231, 1); lean_inc(x_232); lean_dec(x_231); -x_233 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__208; +x_233 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__207; x_234 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_224, x_233, x_232); if (lean_obj_tag(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_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; x_235 = lean_ctor_get(x_234, 1); lean_inc(x_235); lean_dec(x_234); -x_236 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__210; -x_237 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__211; -x_238 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__213; -x_239 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__214; -x_240 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__215; -x_241 = l_Lean_Parser_registerAlias(x_236, x_237, x_238, x_239, x_240, x_235); -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__217; -x_244 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_236, x_243, x_242); -if (lean_obj_tag(x_244) == 0) -{ -lean_object* x_245; lean_object* x_246; lean_object* x_247; -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__219; -x_247 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_236, x_246, x_245); -if (lean_obj_tag(x_247) == 0) -{ -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_248 = lean_ctor_get(x_247, 1); -lean_inc(x_248); -lean_dec(x_247); -x_249 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__221; +x_236 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__209; +x_237 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__210; +x_238 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__212; +x_239 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__213; +x_240 = l_Lean_Parser_registerAlias(x_236, x_237, x_238, x_239, x_45, x_235); +if (lean_obj_tag(x_240) == 0) +{ +lean_object* x_241; lean_object* x_242; lean_object* x_243; +x_241 = lean_ctor_get(x_240, 1); +lean_inc(x_241); +lean_dec(x_240); +x_242 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__215; +x_243 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_236, x_242, x_241); +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__217; +x_246 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_236, x_245, x_244); +if (lean_obj_tag(x_246) == 0) +{ +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_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__219; +x_249 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__220; x_250 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__222; -x_251 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__224; -x_252 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__225; -x_253 = l_Lean_Parser_registerAlias(x_249, x_250, x_251, x_252, x_179, x_248); +x_251 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__223; +x_252 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__224; +x_253 = l_Lean_Parser_registerAlias(x_248, x_249, x_250, x_251, x_252, x_247); if (lean_obj_tag(x_253) == 0) { lean_object* x_254; lean_object* x_255; lean_object* x_256; x_254 = lean_ctor_get(x_253, 1); lean_inc(x_254); lean_dec(x_253); -x_255 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__227; -x_256 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_249, x_255, x_254); +x_255 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__226; +x_256 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_248, x_255, x_254); if (lean_obj_tag(x_256) == 0) { lean_object* x_257; lean_object* x_258; lean_object* x_259; x_257 = lean_ctor_get(x_256, 1); lean_inc(x_257); lean_dec(x_256); -x_258 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__229; -x_259 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_249, x_258, x_257); +x_258 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__228; +x_259 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_248, x_258, x_257); if (lean_obj_tag(x_259) == 0) { lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; x_260 = lean_ctor_get(x_259, 1); lean_inc(x_260); lean_dec(x_259); -x_261 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__231; -x_262 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__232; -x_263 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__234; -x_264 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__235; -x_265 = l_Lean_Parser_registerAlias(x_261, x_262, x_263, x_264, x_179, x_260); +x_261 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__230; +x_262 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__231; +x_263 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__233; +x_264 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__234; +x_265 = l_Lean_Parser_registerAlias(x_261, x_262, x_263, x_264, x_191, x_260); if (lean_obj_tag(x_265) == 0) { lean_object* x_266; lean_object* x_267; lean_object* x_268; x_266 = lean_ctor_get(x_265, 1); lean_inc(x_266); lean_dec(x_265); -x_267 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__237; +x_267 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__236; x_268 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_261, x_267, x_266); if (lean_obj_tag(x_268) == 0) { @@ -3832,7 +3929,7 @@ lean_object* x_269; lean_object* x_270; lean_object* x_271; x_269 = lean_ctor_get(x_268, 1); lean_inc(x_269); lean_dec(x_268); -x_270 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__239; +x_270 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__238; x_271 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_261, x_270, x_269); if (lean_obj_tag(x_271) == 0) { @@ -3840,18 +3937,18 @@ lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; x_272 = lean_ctor_get(x_271, 1); lean_inc(x_272); lean_dec(x_271); -x_273 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__241; -x_274 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__242; -x_275 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__244; -x_276 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__245; -x_277 = l_Lean_Parser_registerAlias(x_273, x_274, x_275, x_276, x_179, x_272); +x_273 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__240; +x_274 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__241; +x_275 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__243; +x_276 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__244; +x_277 = l_Lean_Parser_registerAlias(x_273, x_274, x_275, x_276, x_191, x_272); if (lean_obj_tag(x_277) == 0) { lean_object* x_278; lean_object* x_279; lean_object* x_280; x_278 = lean_ctor_get(x_277, 1); lean_inc(x_278); lean_dec(x_277); -x_279 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__247; +x_279 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__246; x_280 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_273, x_279, x_278); if (lean_obj_tag(x_280) == 0) { @@ -3859,7 +3956,7 @@ lean_object* x_281; lean_object* x_282; lean_object* x_283; x_281 = lean_ctor_get(x_280, 1); lean_inc(x_281); lean_dec(x_280); -x_282 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__249; +x_282 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__248; x_283 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_273, x_282, x_281); if (lean_obj_tag(x_283) == 0) { @@ -3867,18 +3964,18 @@ lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; x_284 = lean_ctor_get(x_283, 1); lean_inc(x_284); lean_dec(x_283); -x_285 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__251; -x_286 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__252; -x_287 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__254; -x_288 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__257; -x_289 = l_Lean_Parser_registerAlias(x_285, x_286, x_287, x_288, x_45, x_284); +x_285 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__250; +x_286 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__251; +x_287 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__253; +x_288 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__254; +x_289 = l_Lean_Parser_registerAlias(x_285, x_286, x_287, x_288, x_191, x_284); if (lean_obj_tag(x_289) == 0) { lean_object* x_290; lean_object* x_291; lean_object* x_292; x_290 = lean_ctor_get(x_289, 1); lean_inc(x_290); lean_dec(x_289); -x_291 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__259; +x_291 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__256; x_292 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_285, x_291, x_290); if (lean_obj_tag(x_292) == 0) { @@ -3886,7 +3983,7 @@ lean_object* x_293; lean_object* x_294; lean_object* x_295; x_293 = lean_ctor_get(x_292, 1); lean_inc(x_293); lean_dec(x_292); -x_294 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__261; +x_294 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__258; x_295 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_285, x_294, x_293); if (lean_obj_tag(x_295) == 0) { @@ -3894,10 +3991,10 @@ lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; x_296 = lean_ctor_get(x_295, 1); lean_inc(x_296); lean_dec(x_295); -x_297 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__263; -x_298 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__264; -x_299 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__266; -x_300 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__267; +x_297 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__260; +x_298 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__261; +x_299 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__263; +x_300 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__266; x_301 = l_Lean_Parser_registerAlias(x_297, x_298, x_299, x_300, x_45, x_296); if (lean_obj_tag(x_301) == 0) { @@ -3905,7 +4002,7 @@ lean_object* x_302; lean_object* x_303; lean_object* x_304; x_302 = lean_ctor_get(x_301, 1); lean_inc(x_302); lean_dec(x_301); -x_303 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__269; +x_303 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__268; x_304 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_297, x_303, x_302); if (lean_obj_tag(x_304) == 0) { @@ -3913,7 +4010,7 @@ lean_object* x_305; lean_object* x_306; lean_object* x_307; x_305 = lean_ctor_get(x_304, 1); lean_inc(x_305); lean_dec(x_304); -x_306 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__271; +x_306 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__270; x_307 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_297, x_306, x_305); if (lean_obj_tag(x_307) == 0) { @@ -3921,18 +4018,18 @@ lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; x_308 = lean_ctor_get(x_307, 1); lean_inc(x_308); lean_dec(x_307); -x_309 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__273; -x_310 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__274; -x_311 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__276; -x_312 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__277; -x_313 = l_Lean_Parser_registerAlias(x_309, x_310, x_311, x_312, x_179, x_308); +x_309 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__272; +x_310 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__273; +x_311 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__275; +x_312 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__276; +x_313 = l_Lean_Parser_registerAlias(x_309, x_310, x_311, x_312, x_45, x_308); if (lean_obj_tag(x_313) == 0) { lean_object* x_314; lean_object* x_315; lean_object* x_316; x_314 = lean_ctor_get(x_313, 1); lean_inc(x_314); lean_dec(x_313); -x_315 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__279; +x_315 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__278; x_316 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_309, x_315, x_314); if (lean_obj_tag(x_316) == 0) { @@ -3940,121 +4037,79 @@ lean_object* x_317; lean_object* x_318; lean_object* x_319; x_317 = lean_ctor_get(x_316, 1); lean_inc(x_317); lean_dec(x_316); -x_318 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__281; +x_318 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__280; x_319 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_309, x_318, x_317); if (lean_obj_tag(x_319) == 0) { -lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; +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_320 = lean_ctor_get(x_319, 1); lean_inc(x_320); lean_dec(x_319); -x_321 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__283; -x_322 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__284; -x_323 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__286; -x_324 = l_Lean_Parser_registerAlias(x_321, x_322, x_323, x_174, x_45, x_320); -if (lean_obj_tag(x_324) == 0) -{ -lean_object* x_325; lean_object* x_326; lean_object* x_327; -x_325 = lean_ctor_get(x_324, 1); -lean_inc(x_325); -lean_dec(x_324); -x_326 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__288; -x_327 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_321, x_326, x_325); -if (lean_obj_tag(x_327) == 0) -{ -lean_object* x_328; lean_object* x_329; lean_object* x_330; -x_328 = lean_ctor_get(x_327, 1); -lean_inc(x_328); -lean_dec(x_327); -x_329 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__290; -x_330 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_321, x_329, x_328); -return x_330; -} -else -{ -uint8_t x_331; -x_331 = !lean_is_exclusive(x_327); -if (x_331 == 0) -{ -return x_327; -} -else -{ -lean_object* x_332; lean_object* x_333; lean_object* x_334; -x_332 = lean_ctor_get(x_327, 0); -x_333 = lean_ctor_get(x_327, 1); -lean_inc(x_333); +x_321 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__282; +x_322 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__283; +x_323 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__285; +x_324 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__286; +x_325 = l_Lean_Parser_registerAlias(x_321, x_322, x_323, x_324, x_191, x_320); +if (lean_obj_tag(x_325) == 0) +{ +lean_object* x_326; lean_object* x_327; lean_object* x_328; +x_326 = lean_ctor_get(x_325, 1); +lean_inc(x_326); +lean_dec(x_325); +x_327 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__288; +x_328 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_321, x_327, x_326); +if (lean_obj_tag(x_328) == 0) +{ +lean_object* x_329; lean_object* x_330; lean_object* x_331; +x_329 = lean_ctor_get(x_328, 1); +lean_inc(x_329); +lean_dec(x_328); +x_330 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__290; +x_331 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_321, x_330, x_329); +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; +x_332 = lean_ctor_get(x_331, 1); lean_inc(x_332); -lean_dec(x_327); -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 -{ -uint8_t x_335; -x_335 = !lean_is_exclusive(x_324); -if (x_335 == 0) -{ -return x_324; -} -else -{ -lean_object* x_336; lean_object* x_337; lean_object* x_338; -x_336 = lean_ctor_get(x_324, 0); -x_337 = lean_ctor_get(x_324, 1); +lean_dec(x_331); +x_333 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__292; +x_334 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__293; +x_335 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__295; +x_336 = l_Lean_Parser_registerAlias(x_333, x_334, x_335, x_186, x_45, x_332); +if (lean_obj_tag(x_336) == 0) +{ +lean_object* x_337; lean_object* x_338; lean_object* x_339; +x_337 = lean_ctor_get(x_336, 1); lean_inc(x_337); -lean_inc(x_336); -lean_dec(x_324); -x_338 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_338, 0, x_336); -lean_ctor_set(x_338, 1, x_337); -return x_338; -} -} -} -else -{ -uint8_t x_339; -x_339 = !lean_is_exclusive(x_319); -if (x_339 == 0) -{ -return x_319; -} -else +lean_dec(x_336); +x_338 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__297; +x_339 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_333, x_338, x_337); +if (lean_obj_tag(x_339) == 0) { lean_object* x_340; lean_object* x_341; lean_object* x_342; -x_340 = lean_ctor_get(x_319, 0); -x_341 = lean_ctor_get(x_319, 1); -lean_inc(x_341); +x_340 = lean_ctor_get(x_339, 1); lean_inc(x_340); -lean_dec(x_319); -x_342 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_342, 0, x_340); -lean_ctor_set(x_342, 1, x_341); +lean_dec(x_339); +x_341 = l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__299; +x_342 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_333, x_341, x_340); return x_342; } -} -} else { uint8_t x_343; -x_343 = !lean_is_exclusive(x_316); +x_343 = !lean_is_exclusive(x_339); if (x_343 == 0) { -return x_316; +return x_339; } else { lean_object* x_344; lean_object* x_345; lean_object* x_346; -x_344 = lean_ctor_get(x_316, 0); -x_345 = lean_ctor_get(x_316, 1); +x_344 = lean_ctor_get(x_339, 0); +x_345 = lean_ctor_get(x_339, 1); lean_inc(x_345); lean_inc(x_344); -lean_dec(x_316); +lean_dec(x_339); x_346 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_346, 0, x_344); lean_ctor_set(x_346, 1, x_345); @@ -4065,19 +4120,19 @@ return x_346; else { uint8_t x_347; -x_347 = !lean_is_exclusive(x_313); +x_347 = !lean_is_exclusive(x_336); if (x_347 == 0) { -return x_313; +return x_336; } else { lean_object* x_348; lean_object* x_349; lean_object* x_350; -x_348 = lean_ctor_get(x_313, 0); -x_349 = lean_ctor_get(x_313, 1); +x_348 = lean_ctor_get(x_336, 0); +x_349 = lean_ctor_get(x_336, 1); lean_inc(x_349); lean_inc(x_348); -lean_dec(x_313); +lean_dec(x_336); x_350 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_350, 0, x_348); lean_ctor_set(x_350, 1, x_349); @@ -4088,19 +4143,19 @@ return x_350; else { uint8_t x_351; -x_351 = !lean_is_exclusive(x_307); +x_351 = !lean_is_exclusive(x_331); if (x_351 == 0) { -return x_307; +return x_331; } else { lean_object* x_352; lean_object* x_353; lean_object* x_354; -x_352 = lean_ctor_get(x_307, 0); -x_353 = lean_ctor_get(x_307, 1); +x_352 = lean_ctor_get(x_331, 0); +x_353 = lean_ctor_get(x_331, 1); lean_inc(x_353); lean_inc(x_352); -lean_dec(x_307); +lean_dec(x_331); x_354 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_354, 0, x_352); lean_ctor_set(x_354, 1, x_353); @@ -4111,19 +4166,19 @@ return x_354; else { uint8_t x_355; -x_355 = !lean_is_exclusive(x_304); +x_355 = !lean_is_exclusive(x_328); if (x_355 == 0) { -return x_304; +return x_328; } else { lean_object* x_356; lean_object* x_357; lean_object* x_358; -x_356 = lean_ctor_get(x_304, 0); -x_357 = lean_ctor_get(x_304, 1); +x_356 = lean_ctor_get(x_328, 0); +x_357 = lean_ctor_get(x_328, 1); lean_inc(x_357); lean_inc(x_356); -lean_dec(x_304); +lean_dec(x_328); x_358 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_358, 0, x_356); lean_ctor_set(x_358, 1, x_357); @@ -4134,19 +4189,19 @@ return x_358; else { uint8_t x_359; -x_359 = !lean_is_exclusive(x_301); +x_359 = !lean_is_exclusive(x_325); if (x_359 == 0) { -return x_301; +return x_325; } else { lean_object* x_360; lean_object* x_361; lean_object* x_362; -x_360 = lean_ctor_get(x_301, 0); -x_361 = lean_ctor_get(x_301, 1); +x_360 = lean_ctor_get(x_325, 0); +x_361 = lean_ctor_get(x_325, 1); lean_inc(x_361); lean_inc(x_360); -lean_dec(x_301); +lean_dec(x_325); x_362 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_362, 0, x_360); lean_ctor_set(x_362, 1, x_361); @@ -4157,19 +4212,19 @@ return x_362; else { uint8_t x_363; -x_363 = !lean_is_exclusive(x_295); +x_363 = !lean_is_exclusive(x_319); if (x_363 == 0) { -return x_295; +return x_319; } else { lean_object* x_364; lean_object* x_365; lean_object* x_366; -x_364 = lean_ctor_get(x_295, 0); -x_365 = lean_ctor_get(x_295, 1); +x_364 = lean_ctor_get(x_319, 0); +x_365 = lean_ctor_get(x_319, 1); lean_inc(x_365); lean_inc(x_364); -lean_dec(x_295); +lean_dec(x_319); x_366 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_366, 0, x_364); lean_ctor_set(x_366, 1, x_365); @@ -4180,19 +4235,19 @@ return x_366; else { uint8_t x_367; -x_367 = !lean_is_exclusive(x_292); +x_367 = !lean_is_exclusive(x_316); if (x_367 == 0) { -return x_292; +return x_316; } else { lean_object* x_368; lean_object* x_369; lean_object* x_370; -x_368 = lean_ctor_get(x_292, 0); -x_369 = lean_ctor_get(x_292, 1); +x_368 = lean_ctor_get(x_316, 0); +x_369 = lean_ctor_get(x_316, 1); lean_inc(x_369); lean_inc(x_368); -lean_dec(x_292); +lean_dec(x_316); x_370 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_370, 0, x_368); lean_ctor_set(x_370, 1, x_369); @@ -4203,19 +4258,19 @@ return x_370; else { uint8_t x_371; -x_371 = !lean_is_exclusive(x_289); +x_371 = !lean_is_exclusive(x_313); if (x_371 == 0) { -return x_289; +return x_313; } else { lean_object* x_372; lean_object* x_373; lean_object* x_374; -x_372 = lean_ctor_get(x_289, 0); -x_373 = lean_ctor_get(x_289, 1); +x_372 = lean_ctor_get(x_313, 0); +x_373 = lean_ctor_get(x_313, 1); lean_inc(x_373); lean_inc(x_372); -lean_dec(x_289); +lean_dec(x_313); x_374 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_374, 0, x_372); lean_ctor_set(x_374, 1, x_373); @@ -4226,19 +4281,19 @@ return x_374; else { uint8_t x_375; -x_375 = !lean_is_exclusive(x_283); +x_375 = !lean_is_exclusive(x_307); if (x_375 == 0) { -return x_283; +return x_307; } else { lean_object* x_376; lean_object* x_377; lean_object* x_378; -x_376 = lean_ctor_get(x_283, 0); -x_377 = lean_ctor_get(x_283, 1); +x_376 = lean_ctor_get(x_307, 0); +x_377 = lean_ctor_get(x_307, 1); lean_inc(x_377); lean_inc(x_376); -lean_dec(x_283); +lean_dec(x_307); x_378 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_378, 0, x_376); lean_ctor_set(x_378, 1, x_377); @@ -4249,19 +4304,19 @@ return x_378; else { uint8_t x_379; -x_379 = !lean_is_exclusive(x_280); +x_379 = !lean_is_exclusive(x_304); if (x_379 == 0) { -return x_280; +return x_304; } else { lean_object* x_380; lean_object* x_381; lean_object* x_382; -x_380 = lean_ctor_get(x_280, 0); -x_381 = lean_ctor_get(x_280, 1); +x_380 = lean_ctor_get(x_304, 0); +x_381 = lean_ctor_get(x_304, 1); lean_inc(x_381); lean_inc(x_380); -lean_dec(x_280); +lean_dec(x_304); x_382 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_382, 0, x_380); lean_ctor_set(x_382, 1, x_381); @@ -4272,19 +4327,19 @@ return x_382; else { uint8_t x_383; -x_383 = !lean_is_exclusive(x_277); +x_383 = !lean_is_exclusive(x_301); if (x_383 == 0) { -return x_277; +return x_301; } else { lean_object* x_384; lean_object* x_385; lean_object* x_386; -x_384 = lean_ctor_get(x_277, 0); -x_385 = lean_ctor_get(x_277, 1); +x_384 = lean_ctor_get(x_301, 0); +x_385 = lean_ctor_get(x_301, 1); lean_inc(x_385); lean_inc(x_384); -lean_dec(x_277); +lean_dec(x_301); x_386 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_386, 0, x_384); lean_ctor_set(x_386, 1, x_385); @@ -4295,19 +4350,19 @@ return x_386; else { uint8_t x_387; -x_387 = !lean_is_exclusive(x_271); +x_387 = !lean_is_exclusive(x_295); if (x_387 == 0) { -return x_271; +return x_295; } else { lean_object* x_388; lean_object* x_389; lean_object* x_390; -x_388 = lean_ctor_get(x_271, 0); -x_389 = lean_ctor_get(x_271, 1); +x_388 = lean_ctor_get(x_295, 0); +x_389 = lean_ctor_get(x_295, 1); lean_inc(x_389); lean_inc(x_388); -lean_dec(x_271); +lean_dec(x_295); x_390 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_390, 0, x_388); lean_ctor_set(x_390, 1, x_389); @@ -4318,19 +4373,19 @@ return x_390; else { uint8_t x_391; -x_391 = !lean_is_exclusive(x_268); +x_391 = !lean_is_exclusive(x_292); if (x_391 == 0) { -return x_268; +return x_292; } else { lean_object* x_392; lean_object* x_393; lean_object* x_394; -x_392 = lean_ctor_get(x_268, 0); -x_393 = lean_ctor_get(x_268, 1); +x_392 = lean_ctor_get(x_292, 0); +x_393 = lean_ctor_get(x_292, 1); lean_inc(x_393); lean_inc(x_392); -lean_dec(x_268); +lean_dec(x_292); x_394 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_394, 0, x_392); lean_ctor_set(x_394, 1, x_393); @@ -4341,19 +4396,19 @@ return x_394; else { uint8_t x_395; -x_395 = !lean_is_exclusive(x_265); +x_395 = !lean_is_exclusive(x_289); if (x_395 == 0) { -return x_265; +return x_289; } else { lean_object* x_396; lean_object* x_397; lean_object* x_398; -x_396 = lean_ctor_get(x_265, 0); -x_397 = lean_ctor_get(x_265, 1); +x_396 = lean_ctor_get(x_289, 0); +x_397 = lean_ctor_get(x_289, 1); lean_inc(x_397); lean_inc(x_396); -lean_dec(x_265); +lean_dec(x_289); x_398 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_398, 0, x_396); lean_ctor_set(x_398, 1, x_397); @@ -4364,19 +4419,19 @@ return x_398; else { uint8_t x_399; -x_399 = !lean_is_exclusive(x_259); +x_399 = !lean_is_exclusive(x_283); if (x_399 == 0) { -return x_259; +return x_283; } else { lean_object* x_400; lean_object* x_401; lean_object* x_402; -x_400 = lean_ctor_get(x_259, 0); -x_401 = lean_ctor_get(x_259, 1); +x_400 = lean_ctor_get(x_283, 0); +x_401 = lean_ctor_get(x_283, 1); lean_inc(x_401); lean_inc(x_400); -lean_dec(x_259); +lean_dec(x_283); x_402 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_402, 0, x_400); lean_ctor_set(x_402, 1, x_401); @@ -4387,19 +4442,19 @@ return x_402; else { uint8_t x_403; -x_403 = !lean_is_exclusive(x_256); +x_403 = !lean_is_exclusive(x_280); if (x_403 == 0) { -return x_256; +return x_280; } else { lean_object* x_404; lean_object* x_405; lean_object* x_406; -x_404 = lean_ctor_get(x_256, 0); -x_405 = lean_ctor_get(x_256, 1); +x_404 = lean_ctor_get(x_280, 0); +x_405 = lean_ctor_get(x_280, 1); lean_inc(x_405); lean_inc(x_404); -lean_dec(x_256); +lean_dec(x_280); x_406 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_406, 0, x_404); lean_ctor_set(x_406, 1, x_405); @@ -4410,19 +4465,19 @@ return x_406; else { uint8_t x_407; -x_407 = !lean_is_exclusive(x_253); +x_407 = !lean_is_exclusive(x_277); if (x_407 == 0) { -return x_253; +return x_277; } else { lean_object* x_408; lean_object* x_409; lean_object* x_410; -x_408 = lean_ctor_get(x_253, 0); -x_409 = lean_ctor_get(x_253, 1); +x_408 = lean_ctor_get(x_277, 0); +x_409 = lean_ctor_get(x_277, 1); lean_inc(x_409); lean_inc(x_408); -lean_dec(x_253); +lean_dec(x_277); x_410 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_410, 0, x_408); lean_ctor_set(x_410, 1, x_409); @@ -4433,19 +4488,19 @@ return x_410; else { uint8_t x_411; -x_411 = !lean_is_exclusive(x_247); +x_411 = !lean_is_exclusive(x_271); if (x_411 == 0) { -return x_247; +return x_271; } else { lean_object* x_412; lean_object* x_413; lean_object* x_414; -x_412 = lean_ctor_get(x_247, 0); -x_413 = lean_ctor_get(x_247, 1); +x_412 = lean_ctor_get(x_271, 0); +x_413 = lean_ctor_get(x_271, 1); lean_inc(x_413); lean_inc(x_412); -lean_dec(x_247); +lean_dec(x_271); x_414 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_414, 0, x_412); lean_ctor_set(x_414, 1, x_413); @@ -4456,19 +4511,19 @@ return x_414; else { uint8_t x_415; -x_415 = !lean_is_exclusive(x_244); +x_415 = !lean_is_exclusive(x_268); if (x_415 == 0) { -return x_244; +return x_268; } else { lean_object* x_416; lean_object* x_417; lean_object* x_418; -x_416 = lean_ctor_get(x_244, 0); -x_417 = lean_ctor_get(x_244, 1); +x_416 = lean_ctor_get(x_268, 0); +x_417 = lean_ctor_get(x_268, 1); lean_inc(x_417); lean_inc(x_416); -lean_dec(x_244); +lean_dec(x_268); x_418 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_418, 0, x_416); lean_ctor_set(x_418, 1, x_417); @@ -4479,19 +4534,19 @@ return x_418; else { uint8_t x_419; -x_419 = !lean_is_exclusive(x_241); +x_419 = !lean_is_exclusive(x_265); if (x_419 == 0) { -return x_241; +return x_265; } else { lean_object* x_420; lean_object* x_421; lean_object* x_422; -x_420 = lean_ctor_get(x_241, 0); -x_421 = lean_ctor_get(x_241, 1); +x_420 = lean_ctor_get(x_265, 0); +x_421 = lean_ctor_get(x_265, 1); lean_inc(x_421); lean_inc(x_420); -lean_dec(x_241); +lean_dec(x_265); x_422 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_422, 0, x_420); lean_ctor_set(x_422, 1, x_421); @@ -4502,19 +4557,19 @@ return x_422; else { uint8_t x_423; -x_423 = !lean_is_exclusive(x_234); +x_423 = !lean_is_exclusive(x_259); if (x_423 == 0) { -return x_234; +return x_259; } else { lean_object* x_424; lean_object* x_425; lean_object* x_426; -x_424 = lean_ctor_get(x_234, 0); -x_425 = lean_ctor_get(x_234, 1); +x_424 = lean_ctor_get(x_259, 0); +x_425 = lean_ctor_get(x_259, 1); lean_inc(x_425); lean_inc(x_424); -lean_dec(x_234); +lean_dec(x_259); x_426 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_426, 0, x_424); lean_ctor_set(x_426, 1, x_425); @@ -4525,19 +4580,19 @@ return x_426; else { uint8_t x_427; -x_427 = !lean_is_exclusive(x_231); +x_427 = !lean_is_exclusive(x_256); if (x_427 == 0) { -return x_231; +return x_256; } else { lean_object* x_428; lean_object* x_429; lean_object* x_430; -x_428 = lean_ctor_get(x_231, 0); -x_429 = lean_ctor_get(x_231, 1); +x_428 = lean_ctor_get(x_256, 0); +x_429 = lean_ctor_get(x_256, 1); lean_inc(x_429); lean_inc(x_428); -lean_dec(x_231); +lean_dec(x_256); x_430 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_430, 0, x_428); lean_ctor_set(x_430, 1, x_429); @@ -4548,19 +4603,19 @@ return x_430; else { uint8_t x_431; -x_431 = !lean_is_exclusive(x_228); +x_431 = !lean_is_exclusive(x_253); if (x_431 == 0) { -return x_228; +return x_253; } else { lean_object* x_432; lean_object* x_433; lean_object* x_434; -x_432 = lean_ctor_get(x_228, 0); -x_433 = lean_ctor_get(x_228, 1); +x_432 = lean_ctor_get(x_253, 0); +x_433 = lean_ctor_get(x_253, 1); lean_inc(x_433); lean_inc(x_432); -lean_dec(x_228); +lean_dec(x_253); x_434 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_434, 0, x_432); lean_ctor_set(x_434, 1, x_433); @@ -4571,19 +4626,19 @@ return x_434; else { uint8_t x_435; -x_435 = !lean_is_exclusive(x_222); +x_435 = !lean_is_exclusive(x_246); if (x_435 == 0) { -return x_222; +return x_246; } else { lean_object* x_436; lean_object* x_437; lean_object* x_438; -x_436 = lean_ctor_get(x_222, 0); -x_437 = lean_ctor_get(x_222, 1); +x_436 = lean_ctor_get(x_246, 0); +x_437 = lean_ctor_get(x_246, 1); lean_inc(x_437); lean_inc(x_436); -lean_dec(x_222); +lean_dec(x_246); x_438 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_438, 0, x_436); lean_ctor_set(x_438, 1, x_437); @@ -4594,19 +4649,19 @@ return x_438; else { uint8_t x_439; -x_439 = !lean_is_exclusive(x_219); +x_439 = !lean_is_exclusive(x_243); if (x_439 == 0) { -return x_219; +return x_243; } else { lean_object* x_440; lean_object* x_441; lean_object* x_442; -x_440 = lean_ctor_get(x_219, 0); -x_441 = lean_ctor_get(x_219, 1); +x_440 = lean_ctor_get(x_243, 0); +x_441 = lean_ctor_get(x_243, 1); lean_inc(x_441); lean_inc(x_440); -lean_dec(x_219); +lean_dec(x_243); x_442 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_442, 0, x_440); lean_ctor_set(x_442, 1, x_441); @@ -4617,19 +4672,19 @@ return x_442; else { uint8_t x_443; -x_443 = !lean_is_exclusive(x_216); +x_443 = !lean_is_exclusive(x_240); if (x_443 == 0) { -return x_216; +return x_240; } else { lean_object* x_444; lean_object* x_445; lean_object* x_446; -x_444 = lean_ctor_get(x_216, 0); -x_445 = lean_ctor_get(x_216, 1); +x_444 = lean_ctor_get(x_240, 0); +x_445 = lean_ctor_get(x_240, 1); lean_inc(x_445); lean_inc(x_444); -lean_dec(x_216); +lean_dec(x_240); x_446 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_446, 0, x_444); lean_ctor_set(x_446, 1, x_445); @@ -4640,19 +4695,19 @@ return x_446; else { uint8_t x_447; -x_447 = !lean_is_exclusive(x_210); +x_447 = !lean_is_exclusive(x_234); if (x_447 == 0) { -return x_210; +return x_234; } else { lean_object* x_448; lean_object* x_449; lean_object* x_450; -x_448 = lean_ctor_get(x_210, 0); -x_449 = lean_ctor_get(x_210, 1); +x_448 = lean_ctor_get(x_234, 0); +x_449 = lean_ctor_get(x_234, 1); lean_inc(x_449); lean_inc(x_448); -lean_dec(x_210); +lean_dec(x_234); x_450 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_450, 0, x_448); lean_ctor_set(x_450, 1, x_449); @@ -4663,19 +4718,19 @@ return x_450; else { uint8_t x_451; -x_451 = !lean_is_exclusive(x_207); +x_451 = !lean_is_exclusive(x_231); if (x_451 == 0) { -return x_207; +return x_231; } else { lean_object* x_452; lean_object* x_453; lean_object* x_454; -x_452 = lean_ctor_get(x_207, 0); -x_453 = lean_ctor_get(x_207, 1); +x_452 = lean_ctor_get(x_231, 0); +x_453 = lean_ctor_get(x_231, 1); lean_inc(x_453); lean_inc(x_452); -lean_dec(x_207); +lean_dec(x_231); x_454 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_454, 0, x_452); lean_ctor_set(x_454, 1, x_453); @@ -4686,19 +4741,19 @@ return x_454; else { uint8_t x_455; -x_455 = !lean_is_exclusive(x_204); +x_455 = !lean_is_exclusive(x_228); if (x_455 == 0) { -return x_204; +return x_228; } else { lean_object* x_456; lean_object* x_457; lean_object* x_458; -x_456 = lean_ctor_get(x_204, 0); -x_457 = lean_ctor_get(x_204, 1); +x_456 = lean_ctor_get(x_228, 0); +x_457 = lean_ctor_get(x_228, 1); lean_inc(x_457); lean_inc(x_456); -lean_dec(x_204); +lean_dec(x_228); x_458 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_458, 0, x_456); lean_ctor_set(x_458, 1, x_457); @@ -4709,19 +4764,19 @@ return x_458; else { uint8_t x_459; -x_459 = !lean_is_exclusive(x_198); +x_459 = !lean_is_exclusive(x_222); if (x_459 == 0) { -return x_198; +return x_222; } else { lean_object* x_460; lean_object* x_461; lean_object* x_462; -x_460 = lean_ctor_get(x_198, 0); -x_461 = lean_ctor_get(x_198, 1); +x_460 = lean_ctor_get(x_222, 0); +x_461 = lean_ctor_get(x_222, 1); lean_inc(x_461); lean_inc(x_460); -lean_dec(x_198); +lean_dec(x_222); x_462 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_462, 0, x_460); lean_ctor_set(x_462, 1, x_461); @@ -4732,19 +4787,19 @@ return x_462; else { uint8_t x_463; -x_463 = !lean_is_exclusive(x_195); +x_463 = !lean_is_exclusive(x_219); if (x_463 == 0) { -return x_195; +return x_219; } else { lean_object* x_464; lean_object* x_465; lean_object* x_466; -x_464 = lean_ctor_get(x_195, 0); -x_465 = lean_ctor_get(x_195, 1); +x_464 = lean_ctor_get(x_219, 0); +x_465 = lean_ctor_get(x_219, 1); lean_inc(x_465); lean_inc(x_464); -lean_dec(x_195); +lean_dec(x_219); x_466 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_466, 0, x_464); lean_ctor_set(x_466, 1, x_465); @@ -4755,19 +4810,19 @@ return x_466; else { uint8_t x_467; -x_467 = !lean_is_exclusive(x_192); +x_467 = !lean_is_exclusive(x_216); if (x_467 == 0) { -return x_192; +return x_216; } else { lean_object* x_468; lean_object* x_469; lean_object* x_470; -x_468 = lean_ctor_get(x_192, 0); -x_469 = lean_ctor_get(x_192, 1); +x_468 = lean_ctor_get(x_216, 0); +x_469 = lean_ctor_get(x_216, 1); lean_inc(x_469); lean_inc(x_468); -lean_dec(x_192); +lean_dec(x_216); x_470 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_470, 0, x_468); lean_ctor_set(x_470, 1, x_469); @@ -4778,19 +4833,19 @@ return x_470; else { uint8_t x_471; -x_471 = !lean_is_exclusive(x_186); +x_471 = !lean_is_exclusive(x_210); if (x_471 == 0) { -return x_186; +return x_210; } else { lean_object* x_472; lean_object* x_473; lean_object* x_474; -x_472 = lean_ctor_get(x_186, 0); -x_473 = lean_ctor_get(x_186, 1); +x_472 = lean_ctor_get(x_210, 0); +x_473 = lean_ctor_get(x_210, 1); lean_inc(x_473); lean_inc(x_472); -lean_dec(x_186); +lean_dec(x_210); x_474 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_474, 0, x_472); lean_ctor_set(x_474, 1, x_473); @@ -4801,19 +4856,19 @@ return x_474; else { uint8_t x_475; -x_475 = !lean_is_exclusive(x_183); +x_475 = !lean_is_exclusive(x_207); if (x_475 == 0) { -return x_183; +return x_207; } else { lean_object* x_476; lean_object* x_477; lean_object* x_478; -x_476 = lean_ctor_get(x_183, 0); -x_477 = lean_ctor_get(x_183, 1); +x_476 = lean_ctor_get(x_207, 0); +x_477 = lean_ctor_get(x_207, 1); lean_inc(x_477); lean_inc(x_476); -lean_dec(x_183); +lean_dec(x_207); x_478 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_478, 0, x_476); lean_ctor_set(x_478, 1, x_477); @@ -4824,19 +4879,19 @@ return x_478; else { uint8_t x_479; -x_479 = !lean_is_exclusive(x_180); +x_479 = !lean_is_exclusive(x_204); if (x_479 == 0) { -return x_180; +return x_204; } else { lean_object* x_480; lean_object* x_481; lean_object* x_482; -x_480 = lean_ctor_get(x_180, 0); -x_481 = lean_ctor_get(x_180, 1); +x_480 = lean_ctor_get(x_204, 0); +x_481 = lean_ctor_get(x_204, 1); lean_inc(x_481); lean_inc(x_480); -lean_dec(x_180); +lean_dec(x_204); x_482 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_482, 0, x_480); lean_ctor_set(x_482, 1, x_481); @@ -4847,19 +4902,19 @@ return x_482; else { uint8_t x_483; -x_483 = !lean_is_exclusive(x_172); +x_483 = !lean_is_exclusive(x_198); if (x_483 == 0) { -return x_172; +return x_198; } else { lean_object* x_484; lean_object* x_485; lean_object* x_486; -x_484 = lean_ctor_get(x_172, 0); -x_485 = lean_ctor_get(x_172, 1); +x_484 = lean_ctor_get(x_198, 0); +x_485 = lean_ctor_get(x_198, 1); lean_inc(x_485); lean_inc(x_484); -lean_dec(x_172); +lean_dec(x_198); x_486 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_486, 0, x_484); lean_ctor_set(x_486, 1, x_485); @@ -4870,19 +4925,19 @@ return x_486; else { uint8_t x_487; -x_487 = !lean_is_exclusive(x_169); +x_487 = !lean_is_exclusive(x_195); if (x_487 == 0) { -return x_169; +return x_195; } else { lean_object* x_488; lean_object* x_489; lean_object* x_490; -x_488 = lean_ctor_get(x_169, 0); -x_489 = lean_ctor_get(x_169, 1); +x_488 = lean_ctor_get(x_195, 0); +x_489 = lean_ctor_get(x_195, 1); lean_inc(x_489); lean_inc(x_488); -lean_dec(x_169); +lean_dec(x_195); x_490 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_490, 0, x_488); lean_ctor_set(x_490, 1, x_489); @@ -4893,19 +4948,19 @@ return x_490; else { uint8_t x_491; -x_491 = !lean_is_exclusive(x_166); +x_491 = !lean_is_exclusive(x_192); if (x_491 == 0) { -return x_166; +return x_192; } else { lean_object* x_492; lean_object* x_493; lean_object* x_494; -x_492 = lean_ctor_get(x_166, 0); -x_493 = lean_ctor_get(x_166, 1); +x_492 = lean_ctor_get(x_192, 0); +x_493 = lean_ctor_get(x_192, 1); lean_inc(x_493); lean_inc(x_492); -lean_dec(x_166); +lean_dec(x_192); x_494 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_494, 0, x_492); lean_ctor_set(x_494, 1, x_493); @@ -4916,19 +4971,19 @@ return x_494; else { uint8_t x_495; -x_495 = !lean_is_exclusive(x_160); +x_495 = !lean_is_exclusive(x_184); if (x_495 == 0) { -return x_160; +return x_184; } else { lean_object* x_496; lean_object* x_497; lean_object* x_498; -x_496 = lean_ctor_get(x_160, 0); -x_497 = lean_ctor_get(x_160, 1); +x_496 = lean_ctor_get(x_184, 0); +x_497 = lean_ctor_get(x_184, 1); lean_inc(x_497); lean_inc(x_496); -lean_dec(x_160); +lean_dec(x_184); x_498 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_498, 0, x_496); lean_ctor_set(x_498, 1, x_497); @@ -4939,19 +4994,19 @@ return x_498; else { uint8_t x_499; -x_499 = !lean_is_exclusive(x_157); +x_499 = !lean_is_exclusive(x_181); if (x_499 == 0) { -return x_157; +return x_181; } else { lean_object* x_500; lean_object* x_501; lean_object* x_502; -x_500 = lean_ctor_get(x_157, 0); -x_501 = lean_ctor_get(x_157, 1); +x_500 = lean_ctor_get(x_181, 0); +x_501 = lean_ctor_get(x_181, 1); lean_inc(x_501); lean_inc(x_500); -lean_dec(x_157); +lean_dec(x_181); x_502 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_502, 0, x_500); lean_ctor_set(x_502, 1, x_501); @@ -4962,19 +5017,19 @@ return x_502; else { uint8_t x_503; -x_503 = !lean_is_exclusive(x_154); +x_503 = !lean_is_exclusive(x_178); if (x_503 == 0) { -return x_154; +return x_178; } else { lean_object* x_504; lean_object* x_505; lean_object* x_506; -x_504 = lean_ctor_get(x_154, 0); -x_505 = lean_ctor_get(x_154, 1); +x_504 = lean_ctor_get(x_178, 0); +x_505 = lean_ctor_get(x_178, 1); lean_inc(x_505); lean_inc(x_504); -lean_dec(x_154); +lean_dec(x_178); x_506 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_506, 0, x_504); lean_ctor_set(x_506, 1, x_505); @@ -4985,19 +5040,19 @@ return x_506; else { uint8_t x_507; -x_507 = !lean_is_exclusive(x_148); +x_507 = !lean_is_exclusive(x_172); if (x_507 == 0) { -return x_148; +return x_172; } else { lean_object* x_508; lean_object* x_509; lean_object* x_510; -x_508 = lean_ctor_get(x_148, 0); -x_509 = lean_ctor_get(x_148, 1); +x_508 = lean_ctor_get(x_172, 0); +x_509 = lean_ctor_get(x_172, 1); lean_inc(x_509); lean_inc(x_508); -lean_dec(x_148); +lean_dec(x_172); x_510 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_510, 0, x_508); lean_ctor_set(x_510, 1, x_509); @@ -5008,19 +5063,19 @@ return x_510; else { uint8_t x_511; -x_511 = !lean_is_exclusive(x_145); +x_511 = !lean_is_exclusive(x_169); if (x_511 == 0) { -return x_145; +return x_169; } else { lean_object* x_512; lean_object* x_513; lean_object* x_514; -x_512 = lean_ctor_get(x_145, 0); -x_513 = lean_ctor_get(x_145, 1); +x_512 = lean_ctor_get(x_169, 0); +x_513 = lean_ctor_get(x_169, 1); lean_inc(x_513); lean_inc(x_512); -lean_dec(x_145); +lean_dec(x_169); x_514 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_514, 0, x_512); lean_ctor_set(x_514, 1, x_513); @@ -5031,19 +5086,19 @@ return x_514; else { uint8_t x_515; -x_515 = !lean_is_exclusive(x_142); +x_515 = !lean_is_exclusive(x_166); if (x_515 == 0) { -return x_142; +return x_166; } else { lean_object* x_516; lean_object* x_517; lean_object* x_518; -x_516 = lean_ctor_get(x_142, 0); -x_517 = lean_ctor_get(x_142, 1); +x_516 = lean_ctor_get(x_166, 0); +x_517 = lean_ctor_get(x_166, 1); lean_inc(x_517); lean_inc(x_516); -lean_dec(x_142); +lean_dec(x_166); x_518 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_518, 0, x_516); lean_ctor_set(x_518, 1, x_517); @@ -5054,19 +5109,19 @@ return x_518; else { uint8_t x_519; -x_519 = !lean_is_exclusive(x_136); +x_519 = !lean_is_exclusive(x_160); if (x_519 == 0) { -return x_136; +return x_160; } else { lean_object* x_520; lean_object* x_521; lean_object* x_522; -x_520 = lean_ctor_get(x_136, 0); -x_521 = lean_ctor_get(x_136, 1); +x_520 = lean_ctor_get(x_160, 0); +x_521 = lean_ctor_get(x_160, 1); lean_inc(x_521); lean_inc(x_520); -lean_dec(x_136); +lean_dec(x_160); x_522 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_522, 0, x_520); lean_ctor_set(x_522, 1, x_521); @@ -5077,19 +5132,19 @@ return x_522; else { uint8_t x_523; -x_523 = !lean_is_exclusive(x_133); +x_523 = !lean_is_exclusive(x_157); if (x_523 == 0) { -return x_133; +return x_157; } else { lean_object* x_524; lean_object* x_525; lean_object* x_526; -x_524 = lean_ctor_get(x_133, 0); -x_525 = lean_ctor_get(x_133, 1); +x_524 = lean_ctor_get(x_157, 0); +x_525 = lean_ctor_get(x_157, 1); lean_inc(x_525); lean_inc(x_524); -lean_dec(x_133); +lean_dec(x_157); x_526 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_526, 0, x_524); lean_ctor_set(x_526, 1, x_525); @@ -5100,19 +5155,19 @@ return x_526; else { uint8_t x_527; -x_527 = !lean_is_exclusive(x_130); +x_527 = !lean_is_exclusive(x_154); if (x_527 == 0) { -return x_130; +return x_154; } else { lean_object* x_528; lean_object* x_529; lean_object* x_530; -x_528 = lean_ctor_get(x_130, 0); -x_529 = lean_ctor_get(x_130, 1); +x_528 = lean_ctor_get(x_154, 0); +x_529 = lean_ctor_get(x_154, 1); lean_inc(x_529); lean_inc(x_528); -lean_dec(x_130); +lean_dec(x_154); x_530 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_530, 0, x_528); lean_ctor_set(x_530, 1, x_529); @@ -5123,19 +5178,19 @@ return x_530; else { uint8_t x_531; -x_531 = !lean_is_exclusive(x_124); +x_531 = !lean_is_exclusive(x_148); if (x_531 == 0) { -return x_124; +return x_148; } else { lean_object* x_532; lean_object* x_533; lean_object* x_534; -x_532 = lean_ctor_get(x_124, 0); -x_533 = lean_ctor_get(x_124, 1); +x_532 = lean_ctor_get(x_148, 0); +x_533 = lean_ctor_get(x_148, 1); lean_inc(x_533); lean_inc(x_532); -lean_dec(x_124); +lean_dec(x_148); x_534 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_534, 0, x_532); lean_ctor_set(x_534, 1, x_533); @@ -5146,19 +5201,19 @@ return x_534; else { uint8_t x_535; -x_535 = !lean_is_exclusive(x_121); +x_535 = !lean_is_exclusive(x_145); if (x_535 == 0) { -return x_121; +return x_145; } else { lean_object* x_536; lean_object* x_537; lean_object* x_538; -x_536 = lean_ctor_get(x_121, 0); -x_537 = lean_ctor_get(x_121, 1); +x_536 = lean_ctor_get(x_145, 0); +x_537 = lean_ctor_get(x_145, 1); lean_inc(x_537); lean_inc(x_536); -lean_dec(x_121); +lean_dec(x_145); x_538 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_538, 0, x_536); lean_ctor_set(x_538, 1, x_537); @@ -5169,19 +5224,19 @@ return x_538; else { uint8_t x_539; -x_539 = !lean_is_exclusive(x_118); +x_539 = !lean_is_exclusive(x_142); if (x_539 == 0) { -return x_118; +return x_142; } else { lean_object* x_540; lean_object* x_541; lean_object* x_542; -x_540 = lean_ctor_get(x_118, 0); -x_541 = lean_ctor_get(x_118, 1); +x_540 = lean_ctor_get(x_142, 0); +x_541 = lean_ctor_get(x_142, 1); lean_inc(x_541); lean_inc(x_540); -lean_dec(x_118); +lean_dec(x_142); x_542 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_542, 0, x_540); lean_ctor_set(x_542, 1, x_541); @@ -5192,19 +5247,19 @@ return x_542; else { uint8_t x_543; -x_543 = !lean_is_exclusive(x_112); +x_543 = !lean_is_exclusive(x_136); if (x_543 == 0) { -return x_112; +return x_136; } else { lean_object* x_544; lean_object* x_545; lean_object* x_546; -x_544 = lean_ctor_get(x_112, 0); -x_545 = lean_ctor_get(x_112, 1); +x_544 = lean_ctor_get(x_136, 0); +x_545 = lean_ctor_get(x_136, 1); lean_inc(x_545); lean_inc(x_544); -lean_dec(x_112); +lean_dec(x_136); x_546 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_546, 0, x_544); lean_ctor_set(x_546, 1, x_545); @@ -5215,19 +5270,19 @@ return x_546; else { uint8_t x_547; -x_547 = !lean_is_exclusive(x_109); +x_547 = !lean_is_exclusive(x_133); if (x_547 == 0) { -return x_109; +return x_133; } else { lean_object* x_548; lean_object* x_549; lean_object* x_550; -x_548 = lean_ctor_get(x_109, 0); -x_549 = lean_ctor_get(x_109, 1); +x_548 = lean_ctor_get(x_133, 0); +x_549 = lean_ctor_get(x_133, 1); lean_inc(x_549); lean_inc(x_548); -lean_dec(x_109); +lean_dec(x_133); x_550 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_550, 0, x_548); lean_ctor_set(x_550, 1, x_549); @@ -5238,19 +5293,19 @@ return x_550; else { uint8_t x_551; -x_551 = !lean_is_exclusive(x_106); +x_551 = !lean_is_exclusive(x_130); if (x_551 == 0) { -return x_106; +return x_130; } else { lean_object* x_552; lean_object* x_553; lean_object* x_554; -x_552 = lean_ctor_get(x_106, 0); -x_553 = lean_ctor_get(x_106, 1); +x_552 = lean_ctor_get(x_130, 0); +x_553 = lean_ctor_get(x_130, 1); lean_inc(x_553); lean_inc(x_552); -lean_dec(x_106); +lean_dec(x_130); x_554 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_554, 0, x_552); lean_ctor_set(x_554, 1, x_553); @@ -5261,19 +5316,19 @@ return x_554; else { uint8_t x_555; -x_555 = !lean_is_exclusive(x_100); +x_555 = !lean_is_exclusive(x_124); if (x_555 == 0) { -return x_100; +return x_124; } else { lean_object* x_556; lean_object* x_557; lean_object* x_558; -x_556 = lean_ctor_get(x_100, 0); -x_557 = lean_ctor_get(x_100, 1); +x_556 = lean_ctor_get(x_124, 0); +x_557 = lean_ctor_get(x_124, 1); lean_inc(x_557); lean_inc(x_556); -lean_dec(x_100); +lean_dec(x_124); x_558 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_558, 0, x_556); lean_ctor_set(x_558, 1, x_557); @@ -5284,19 +5339,19 @@ return x_558; else { uint8_t x_559; -x_559 = !lean_is_exclusive(x_97); +x_559 = !lean_is_exclusive(x_121); if (x_559 == 0) { -return x_97; +return x_121; } else { lean_object* x_560; lean_object* x_561; lean_object* x_562; -x_560 = lean_ctor_get(x_97, 0); -x_561 = lean_ctor_get(x_97, 1); +x_560 = lean_ctor_get(x_121, 0); +x_561 = lean_ctor_get(x_121, 1); lean_inc(x_561); lean_inc(x_560); -lean_dec(x_97); +lean_dec(x_121); x_562 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_562, 0, x_560); lean_ctor_set(x_562, 1, x_561); @@ -5307,19 +5362,19 @@ return x_562; else { uint8_t x_563; -x_563 = !lean_is_exclusive(x_94); +x_563 = !lean_is_exclusive(x_118); if (x_563 == 0) { -return x_94; +return x_118; } else { lean_object* x_564; lean_object* x_565; lean_object* x_566; -x_564 = lean_ctor_get(x_94, 0); -x_565 = lean_ctor_get(x_94, 1); +x_564 = lean_ctor_get(x_118, 0); +x_565 = lean_ctor_get(x_118, 1); lean_inc(x_565); lean_inc(x_564); -lean_dec(x_94); +lean_dec(x_118); x_566 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_566, 0, x_564); lean_ctor_set(x_566, 1, x_565); @@ -5330,19 +5385,19 @@ return x_566; else { uint8_t x_567; -x_567 = !lean_is_exclusive(x_88); +x_567 = !lean_is_exclusive(x_112); if (x_567 == 0) { -return x_88; +return x_112; } else { lean_object* x_568; lean_object* x_569; lean_object* x_570; -x_568 = lean_ctor_get(x_88, 0); -x_569 = lean_ctor_get(x_88, 1); +x_568 = lean_ctor_get(x_112, 0); +x_569 = lean_ctor_get(x_112, 1); lean_inc(x_569); lean_inc(x_568); -lean_dec(x_88); +lean_dec(x_112); x_570 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_570, 0, x_568); lean_ctor_set(x_570, 1, x_569); @@ -5353,19 +5408,19 @@ return x_570; else { uint8_t x_571; -x_571 = !lean_is_exclusive(x_85); +x_571 = !lean_is_exclusive(x_109); if (x_571 == 0) { -return x_85; +return x_109; } else { lean_object* x_572; lean_object* x_573; lean_object* x_574; -x_572 = lean_ctor_get(x_85, 0); -x_573 = lean_ctor_get(x_85, 1); +x_572 = lean_ctor_get(x_109, 0); +x_573 = lean_ctor_get(x_109, 1); lean_inc(x_573); lean_inc(x_572); -lean_dec(x_85); +lean_dec(x_109); x_574 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_574, 0, x_572); lean_ctor_set(x_574, 1, x_573); @@ -5376,19 +5431,19 @@ return x_574; else { uint8_t x_575; -x_575 = !lean_is_exclusive(x_82); +x_575 = !lean_is_exclusive(x_106); if (x_575 == 0) { -return x_82; +return x_106; } else { lean_object* x_576; lean_object* x_577; lean_object* x_578; -x_576 = lean_ctor_get(x_82, 0); -x_577 = lean_ctor_get(x_82, 1); +x_576 = lean_ctor_get(x_106, 0); +x_577 = lean_ctor_get(x_106, 1); lean_inc(x_577); lean_inc(x_576); -lean_dec(x_82); +lean_dec(x_106); x_578 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_578, 0, x_576); lean_ctor_set(x_578, 1, x_577); @@ -5399,19 +5454,19 @@ return x_578; else { uint8_t x_579; -x_579 = !lean_is_exclusive(x_76); +x_579 = !lean_is_exclusive(x_100); if (x_579 == 0) { -return x_76; +return x_100; } else { lean_object* x_580; lean_object* x_581; lean_object* x_582; -x_580 = lean_ctor_get(x_76, 0); -x_581 = lean_ctor_get(x_76, 1); +x_580 = lean_ctor_get(x_100, 0); +x_581 = lean_ctor_get(x_100, 1); lean_inc(x_581); lean_inc(x_580); -lean_dec(x_76); +lean_dec(x_100); x_582 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_582, 0, x_580); lean_ctor_set(x_582, 1, x_581); @@ -5422,19 +5477,19 @@ return x_582; else { uint8_t x_583; -x_583 = !lean_is_exclusive(x_73); +x_583 = !lean_is_exclusive(x_97); if (x_583 == 0) { -return x_73; +return x_97; } else { lean_object* x_584; lean_object* x_585; lean_object* x_586; -x_584 = lean_ctor_get(x_73, 0); -x_585 = lean_ctor_get(x_73, 1); +x_584 = lean_ctor_get(x_97, 0); +x_585 = lean_ctor_get(x_97, 1); lean_inc(x_585); lean_inc(x_584); -lean_dec(x_73); +lean_dec(x_97); x_586 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_586, 0, x_584); lean_ctor_set(x_586, 1, x_585); @@ -5445,19 +5500,19 @@ return x_586; else { uint8_t x_587; -x_587 = !lean_is_exclusive(x_70); +x_587 = !lean_is_exclusive(x_94); if (x_587 == 0) { -return x_70; +return x_94; } else { lean_object* x_588; lean_object* x_589; lean_object* x_590; -x_588 = lean_ctor_get(x_70, 0); -x_589 = lean_ctor_get(x_70, 1); +x_588 = lean_ctor_get(x_94, 0); +x_589 = lean_ctor_get(x_94, 1); lean_inc(x_589); lean_inc(x_588); -lean_dec(x_70); +lean_dec(x_94); x_590 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_590, 0, x_588); lean_ctor_set(x_590, 1, x_589); @@ -5468,19 +5523,19 @@ return x_590; else { uint8_t x_591; -x_591 = !lean_is_exclusive(x_64); +x_591 = !lean_is_exclusive(x_88); if (x_591 == 0) { -return x_64; +return x_88; } else { lean_object* x_592; lean_object* x_593; lean_object* x_594; -x_592 = lean_ctor_get(x_64, 0); -x_593 = lean_ctor_get(x_64, 1); +x_592 = lean_ctor_get(x_88, 0); +x_593 = lean_ctor_get(x_88, 1); lean_inc(x_593); lean_inc(x_592); -lean_dec(x_64); +lean_dec(x_88); x_594 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_594, 0, x_592); lean_ctor_set(x_594, 1, x_593); @@ -5491,19 +5546,19 @@ return x_594; else { uint8_t x_595; -x_595 = !lean_is_exclusive(x_61); +x_595 = !lean_is_exclusive(x_85); if (x_595 == 0) { -return x_61; +return x_85; } else { lean_object* x_596; lean_object* x_597; lean_object* x_598; -x_596 = lean_ctor_get(x_61, 0); -x_597 = lean_ctor_get(x_61, 1); +x_596 = lean_ctor_get(x_85, 0); +x_597 = lean_ctor_get(x_85, 1); lean_inc(x_597); lean_inc(x_596); -lean_dec(x_61); +lean_dec(x_85); x_598 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_598, 0, x_596); lean_ctor_set(x_598, 1, x_597); @@ -5514,19 +5569,19 @@ return x_598; else { uint8_t x_599; -x_599 = !lean_is_exclusive(x_58); +x_599 = !lean_is_exclusive(x_82); if (x_599 == 0) { -return x_58; +return x_82; } else { lean_object* x_600; lean_object* x_601; lean_object* x_602; -x_600 = lean_ctor_get(x_58, 0); -x_601 = lean_ctor_get(x_58, 1); +x_600 = lean_ctor_get(x_82, 0); +x_601 = lean_ctor_get(x_82, 1); lean_inc(x_601); lean_inc(x_600); -lean_dec(x_58); +lean_dec(x_82); x_602 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_602, 0, x_600); lean_ctor_set(x_602, 1, x_601); @@ -5537,19 +5592,19 @@ return x_602; else { uint8_t x_603; -x_603 = !lean_is_exclusive(x_52); +x_603 = !lean_is_exclusive(x_76); if (x_603 == 0) { -return x_52; +return x_76; } else { lean_object* x_604; lean_object* x_605; lean_object* x_606; -x_604 = lean_ctor_get(x_52, 0); -x_605 = lean_ctor_get(x_52, 1); +x_604 = lean_ctor_get(x_76, 0); +x_605 = lean_ctor_get(x_76, 1); lean_inc(x_605); lean_inc(x_604); -lean_dec(x_52); +lean_dec(x_76); x_606 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_606, 0, x_604); lean_ctor_set(x_606, 1, x_605); @@ -5560,19 +5615,19 @@ return x_606; else { uint8_t x_607; -x_607 = !lean_is_exclusive(x_49); +x_607 = !lean_is_exclusive(x_73); if (x_607 == 0) { -return x_49; +return x_73; } else { lean_object* x_608; lean_object* x_609; lean_object* x_610; -x_608 = lean_ctor_get(x_49, 0); -x_609 = lean_ctor_get(x_49, 1); +x_608 = lean_ctor_get(x_73, 0); +x_609 = lean_ctor_get(x_73, 1); lean_inc(x_609); lean_inc(x_608); -lean_dec(x_49); +lean_dec(x_73); x_610 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_610, 0, x_608); lean_ctor_set(x_610, 1, x_609); @@ -5583,19 +5638,19 @@ return x_610; else { uint8_t x_611; -x_611 = !lean_is_exclusive(x_46); +x_611 = !lean_is_exclusive(x_70); if (x_611 == 0) { -return x_46; +return x_70; } else { lean_object* x_612; lean_object* x_613; lean_object* x_614; -x_612 = lean_ctor_get(x_46, 0); -x_613 = lean_ctor_get(x_46, 1); +x_612 = lean_ctor_get(x_70, 0); +x_613 = lean_ctor_get(x_70, 1); lean_inc(x_613); lean_inc(x_612); -lean_dec(x_46); +lean_dec(x_70); x_614 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_614, 0, x_612); lean_ctor_set(x_614, 1, x_613); @@ -5606,19 +5661,19 @@ return x_614; else { uint8_t x_615; -x_615 = !lean_is_exclusive(x_39); +x_615 = !lean_is_exclusive(x_64); if (x_615 == 0) { -return x_39; +return x_64; } else { lean_object* x_616; lean_object* x_617; lean_object* x_618; -x_616 = lean_ctor_get(x_39, 0); -x_617 = lean_ctor_get(x_39, 1); +x_616 = lean_ctor_get(x_64, 0); +x_617 = lean_ctor_get(x_64, 1); lean_inc(x_617); lean_inc(x_616); -lean_dec(x_39); +lean_dec(x_64); x_618 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_618, 0, x_616); lean_ctor_set(x_618, 1, x_617); @@ -5629,19 +5684,19 @@ return x_618; else { uint8_t x_619; -x_619 = !lean_is_exclusive(x_36); +x_619 = !lean_is_exclusive(x_61); if (x_619 == 0) { -return x_36; +return x_61; } else { lean_object* x_620; lean_object* x_621; lean_object* x_622; -x_620 = lean_ctor_get(x_36, 0); -x_621 = lean_ctor_get(x_36, 1); +x_620 = lean_ctor_get(x_61, 0); +x_621 = lean_ctor_get(x_61, 1); lean_inc(x_621); lean_inc(x_620); -lean_dec(x_36); +lean_dec(x_61); x_622 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_622, 0, x_620); lean_ctor_set(x_622, 1, x_621); @@ -5652,19 +5707,19 @@ return x_622; else { uint8_t x_623; -x_623 = !lean_is_exclusive(x_33); +x_623 = !lean_is_exclusive(x_58); if (x_623 == 0) { -return x_33; +return x_58; } else { lean_object* x_624; lean_object* x_625; lean_object* x_626; -x_624 = lean_ctor_get(x_33, 0); -x_625 = lean_ctor_get(x_33, 1); +x_624 = lean_ctor_get(x_58, 0); +x_625 = lean_ctor_get(x_58, 1); lean_inc(x_625); lean_inc(x_624); -lean_dec(x_33); +lean_dec(x_58); x_626 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_626, 0, x_624); lean_ctor_set(x_626, 1, x_625); @@ -5675,19 +5730,19 @@ return x_626; else { uint8_t x_627; -x_627 = !lean_is_exclusive(x_27); +x_627 = !lean_is_exclusive(x_52); if (x_627 == 0) { -return x_27; +return x_52; } else { lean_object* x_628; lean_object* x_629; lean_object* x_630; -x_628 = lean_ctor_get(x_27, 0); -x_629 = lean_ctor_get(x_27, 1); +x_628 = lean_ctor_get(x_52, 0); +x_629 = lean_ctor_get(x_52, 1); lean_inc(x_629); lean_inc(x_628); -lean_dec(x_27); +lean_dec(x_52); x_630 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_630, 0, x_628); lean_ctor_set(x_630, 1, x_629); @@ -5698,19 +5753,19 @@ return x_630; else { uint8_t x_631; -x_631 = !lean_is_exclusive(x_24); +x_631 = !lean_is_exclusive(x_49); if (x_631 == 0) { -return x_24; +return x_49; } else { lean_object* x_632; lean_object* x_633; lean_object* x_634; -x_632 = lean_ctor_get(x_24, 0); -x_633 = lean_ctor_get(x_24, 1); +x_632 = lean_ctor_get(x_49, 0); +x_633 = lean_ctor_get(x_49, 1); lean_inc(x_633); lean_inc(x_632); -lean_dec(x_24); +lean_dec(x_49); x_634 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_634, 0, x_632); lean_ctor_set(x_634, 1, x_633); @@ -5721,19 +5776,19 @@ return x_634; else { uint8_t x_635; -x_635 = !lean_is_exclusive(x_21); +x_635 = !lean_is_exclusive(x_46); if (x_635 == 0) { -return x_21; +return x_46; } else { lean_object* x_636; lean_object* x_637; lean_object* x_638; -x_636 = lean_ctor_get(x_21, 0); -x_637 = lean_ctor_get(x_21, 1); +x_636 = lean_ctor_get(x_46, 0); +x_637 = lean_ctor_get(x_46, 1); lean_inc(x_637); lean_inc(x_636); -lean_dec(x_21); +lean_dec(x_46); x_638 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_638, 0, x_636); lean_ctor_set(x_638, 1, x_637); @@ -5744,19 +5799,19 @@ return x_638; else { uint8_t x_639; -x_639 = !lean_is_exclusive(x_15); +x_639 = !lean_is_exclusive(x_39); if (x_639 == 0) { -return x_15; +return x_39; } else { lean_object* x_640; lean_object* x_641; lean_object* x_642; -x_640 = lean_ctor_get(x_15, 0); -x_641 = lean_ctor_get(x_15, 1); +x_640 = lean_ctor_get(x_39, 0); +x_641 = lean_ctor_get(x_39, 1); lean_inc(x_641); lean_inc(x_640); -lean_dec(x_15); +lean_dec(x_39); x_642 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_642, 0, x_640); lean_ctor_set(x_642, 1, x_641); @@ -5767,19 +5822,19 @@ return x_642; else { uint8_t x_643; -x_643 = !lean_is_exclusive(x_11); +x_643 = !lean_is_exclusive(x_36); if (x_643 == 0) { -return x_11; +return x_36; } else { lean_object* x_644; lean_object* x_645; lean_object* x_646; -x_644 = lean_ctor_get(x_11, 0); -x_645 = lean_ctor_get(x_11, 1); +x_644 = lean_ctor_get(x_36, 0); +x_645 = lean_ctor_get(x_36, 1); lean_inc(x_645); lean_inc(x_644); -lean_dec(x_11); +lean_dec(x_36); x_646 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_646, 0, x_644); lean_ctor_set(x_646, 1, x_645); @@ -5790,19 +5845,19 @@ return x_646; else { uint8_t x_647; -x_647 = !lean_is_exclusive(x_7); +x_647 = !lean_is_exclusive(x_33); if (x_647 == 0) { -return x_7; +return x_33; } else { lean_object* x_648; lean_object* x_649; lean_object* x_650; -x_648 = lean_ctor_get(x_7, 0); -x_649 = lean_ctor_get(x_7, 1); +x_648 = lean_ctor_get(x_33, 0); +x_649 = lean_ctor_get(x_33, 1); lean_inc(x_649); lean_inc(x_648); -lean_dec(x_7); +lean_dec(x_33); x_650 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_650, 0, x_648); lean_ctor_set(x_650, 1, x_649); @@ -5810,6 +5865,144 @@ return x_650; } } } +else +{ +uint8_t x_651; +x_651 = !lean_is_exclusive(x_27); +if (x_651 == 0) +{ +return x_27; +} +else +{ +lean_object* x_652; lean_object* x_653; lean_object* x_654; +x_652 = lean_ctor_get(x_27, 0); +x_653 = lean_ctor_get(x_27, 1); +lean_inc(x_653); +lean_inc(x_652); +lean_dec(x_27); +x_654 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_654, 0, x_652); +lean_ctor_set(x_654, 1, x_653); +return x_654; +} +} +} +else +{ +uint8_t x_655; +x_655 = !lean_is_exclusive(x_24); +if (x_655 == 0) +{ +return x_24; +} +else +{ +lean_object* x_656; lean_object* x_657; lean_object* x_658; +x_656 = lean_ctor_get(x_24, 0); +x_657 = lean_ctor_get(x_24, 1); +lean_inc(x_657); +lean_inc(x_656); +lean_dec(x_24); +x_658 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_658, 0, x_656); +lean_ctor_set(x_658, 1, x_657); +return x_658; +} +} +} +else +{ +uint8_t x_659; +x_659 = !lean_is_exclusive(x_21); +if (x_659 == 0) +{ +return x_21; +} +else +{ +lean_object* x_660; lean_object* x_661; lean_object* x_662; +x_660 = lean_ctor_get(x_21, 0); +x_661 = lean_ctor_get(x_21, 1); +lean_inc(x_661); +lean_inc(x_660); +lean_dec(x_21); +x_662 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_662, 0, x_660); +lean_ctor_set(x_662, 1, x_661); +return x_662; +} +} +} +else +{ +uint8_t x_663; +x_663 = !lean_is_exclusive(x_15); +if (x_663 == 0) +{ +return x_15; +} +else +{ +lean_object* x_664; lean_object* x_665; lean_object* x_666; +x_664 = lean_ctor_get(x_15, 0); +x_665 = lean_ctor_get(x_15, 1); +lean_inc(x_665); +lean_inc(x_664); +lean_dec(x_15); +x_666 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_666, 0, x_664); +lean_ctor_set(x_666, 1, x_665); +return x_666; +} +} +} +else +{ +uint8_t x_667; +x_667 = !lean_is_exclusive(x_11); +if (x_667 == 0) +{ +return x_11; +} +else +{ +lean_object* x_668; lean_object* x_669; lean_object* x_670; +x_668 = lean_ctor_get(x_11, 0); +x_669 = lean_ctor_get(x_11, 1); +lean_inc(x_669); +lean_inc(x_668); +lean_dec(x_11); +x_670 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_670, 0, x_668); +lean_ctor_set(x_670, 1, x_669); +return x_670; +} +} +} +else +{ +uint8_t x_671; +x_671 = !lean_is_exclusive(x_7); +if (x_671 == 0) +{ +return x_7; +} +else +{ +lean_object* x_672; lean_object* x_673; lean_object* x_674; +x_672 = lean_ctor_get(x_7, 0); +x_673 = lean_ctor_get(x_7, 1); +lean_inc(x_673); +lean_inc(x_672); +lean_dec(x_7); +x_674 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_674, 0, x_672); +lean_ctor_set(x_674, 1, x_673); +return x_674; +} +} +} } 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: @@ -8970,6 +9163,24 @@ l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__289 = _init_l_Lean_Pa lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__289); l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__290 = _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__290(); lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__290); +l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__291 = _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__291(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__291); +l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__292 = _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__292(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__292); +l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__293 = _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__293(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__293); +l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__294 = _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__294(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__294); +l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__295 = _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__295(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__295); +l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__296 = _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__296(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__296); +l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__297 = _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__297(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__297); +l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__298 = _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__298(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__298); +l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__299 = _init_l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__299(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser___hyg_8____closed__299); 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 531e76c2db58..92e995293328 100644 --- a/stage0/stdlib/Lean/Parser/Attr.c +++ b/stage0/stdlib/Lean/Parser/Attr.c @@ -24,6 +24,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Attr_simple(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Attr_recursor_declRange___closed__3; static lean_object* l_Lean_Parser_Attr_extern_parenthesizer___closed__3; 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_Attr_extern_formatter___closed__9; lean_object* l_Lean_Parser_registerBuiltinParserAttribute(lean_object*, lean_object*, uint8_t, lean_object*); static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Attr___hyg_4____closed__19; LEAN_EXPORT lean_object* l_Lean_Parser_Attr_default__instance_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -104,11 +105,13 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Attr_extern_parenthesizer(le static lean_object* l_Lean_Parser_Attr_specialize___closed__2; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Attr___hyg_4____closed__1; static lean_object* l___regBuiltin_Lean_Parser_Attr_simple_declRange___closed__2; +static lean_object* l_Lean_Parser_Attr_simple_formatter___closed__8; static lean_object* l___regBuiltin_Lean_Parser_Attr_class_declRange___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Attr_macro_parenthesizer(lean_object*); lean_object* l_Lean_Parser_optional(lean_object*); lean_object* l_Lean_Parser_nonReservedSymbol(lean_object*, uint8_t); static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Attr___hyg_4____closed__10; +static lean_object* l_Lean_Parser_Attr_instance_parenthesizer___closed__6; static lean_object* l_Lean_Parser_Attr_simple_formatter___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Attr_macro_declRange___closed__7; extern lean_object* l_Lean_Parser_strLit; @@ -134,6 +137,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Attr_specialize; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Attr_extern_formatter(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Attr_simple; static lean_object* l_Lean_Parser_Attr_externEntry_parenthesizer___closed__4; +static lean_object* l_Lean_Parser_Attr_extern___closed__13; lean_object* l_Lean_Name_mkStr5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Attr_externEntry_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Attr___hyg_4____closed__4; @@ -151,6 +155,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_priorityParser_parenthesizer(lean_object* static lean_object* l___regBuiltin_Lean_Parser_Attr_instance_formatter___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Attr_class; static lean_object* l_Lean_Parser_Priority_numPrio_parenthesizer___closed__2; +static lean_object* l_Lean_Parser_Attr_simple_parenthesizer___closed__9; static lean_object* l_Lean_Parser_Attr_specialize_formatter___closed__4; 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_Attr_simple_formatter___closed__7; @@ -162,10 +167,12 @@ static lean_object* l___regBuiltin_Lean_Parser_Attr_specialize_declRange___close static lean_object* l___regBuiltin_Lean_Parser_Attr_extern_declRange___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Priority_numPrio_declRange___closed__5; static lean_object* l_Lean_Parser_Attr_extern___closed__11; +static lean_object* l_Lean_Parser_Attr_specialize___closed__11; static lean_object* l_Lean_Parser_Attr_specialize___closed__3; static lean_object* l_Lean_Parser_Attr_class___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Attr_simple_parenthesizer___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Attr_macro; +static lean_object* l_Lean_Parser_Attr_extern_parenthesizer___closed__8; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Attr_recursor(lean_object*); static lean_object* l_Lean_Parser_Attr_extern___closed__5; 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*); @@ -197,9 +204,11 @@ static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Attr___hyg_4____clos static lean_object* l___regBuiltin_Lean_Parser_Attr_simple_declRange___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Priority_numPrio___closed__1; static lean_object* l_Lean_Parser_Attr_instance___closed__4; +static lean_object* l_Lean_Parser_Attr_simple___closed__11; static lean_object* l___regBuiltin_Lean_Parser_Attr_specialize_declRange___closed__2; static lean_object* l_Lean_Parser_Attr_extern_formatter___closed__7; static lean_object* l_Lean_Parser_Priority_numPrio_formatter___closed__1; +static lean_object* l_Lean_Parser_Attr_instance___closed__10; static lean_object* l_Lean_Parser_Attr_externEntry_formatter___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Attr_export_formatter___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Attr___hyg_4_(lean_object*); @@ -262,12 +271,12 @@ static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Attr___hyg_4____clos lean_object* l_Lean_Parser_symbol_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Attr_extern___closed__8; LEAN_EXPORT lean_object* l_Lean_Parser_attrParser_formatter(lean_object*); -static lean_object* l_Lean_Parser_Attr_default__instance___closed__9; static lean_object* l___regBuiltin_Lean_Parser_Priority_numPrio_declRange___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Attr_specialize_formatter___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Attr_simple_parenthesizer___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Attr_export_formatter___closed__2; static lean_object* l_Lean_Parser_Attr_simple_parenthesizer___closed__5; +static lean_object* l_Lean_Parser_Attr_extern_formatter___closed__8; static lean_object* l_Lean_Parser_Attr_export___closed__2; lean_object* l_Lean_Parser_withAntiquot(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Attr_instance_declRange___closed__1; @@ -322,6 +331,7 @@ static lean_object* l_Lean_Parser_Attr_recursor_parenthesizer___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Attr_export_declRange___closed__5; static lean_object* l_Lean_Parser_Attr_extern_parenthesizer___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Attr_simple_declRange___closed__4; +lean_object* l_Lean_Parser_ppSpace_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Attr_macro_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Attr_macro___closed__2; lean_object* l_Lean_Parser_ident_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -333,6 +343,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Attr_simple_formatter(lean_o static lean_object* l_Lean_Parser_Attr_specialize_formatter___closed__6; static lean_object* l_Lean_Parser_Attr_simple_formatter___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_attrParser_formatter___boxed(lean_object*); +static lean_object* l_Lean_Parser_Attr_simple_parenthesizer___closed__8; static lean_object* l_Lean_Parser_Attr_extern_formatter___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Attr_default__instance_declRange___closed__1; static lean_object* l_Lean_Parser_Attr_instance___closed__6; @@ -344,6 +355,7 @@ static lean_object* l_Lean_Parser_Attr_class_formatter___closed__2; static lean_object* l_Lean_Parser_Attr_class_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Attr_specialize_formatter___closed__3; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Attr_default__instance(lean_object*); +lean_object* l_Lean_ppSpace_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Attr_macro_parenthesizer___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_attrParser_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Attr_export_formatter___closed__3; @@ -354,6 +366,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Attr_macro_parenthesizer___closed 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_Attr_extern_parenthesizer___closed__4; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Attr___hyg_4____closed__17; +static lean_object* l_Lean_Parser_Attr_instance_formatter___closed__6; extern lean_object* l_Lean_Parser_ident; lean_object* l_Lean_Parser_optional_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Attr_externEntry_formatter___closed__1; @@ -376,6 +389,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Attr_instance; static lean_object* l_Lean_Parser_Attr_recursor_formatter___closed__2; lean_object* l_Lean_Parser_withCache(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Attr_externEntry; +static lean_object* l_Lean_Parser_Attr_specialize_parenthesizer___closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_priorityParser_formatter___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Attr_recursor_declRange(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Attr_default__instance_declRange___closed__2; @@ -384,9 +398,11 @@ static lean_object* l_Lean_Parser_Attr_recursor___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Attr_recursor_declRange___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Attr_macro_declRange___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Attr_specialize_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_Parser_skip; static lean_object* l_Lean_Parser_Attr_default__instance___closed__3; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Attr_extern(lean_object*); static lean_object* l_Lean_Parser_Attr_class_parenthesizer___closed__2; +static lean_object* l_Lean_Parser_Attr_externEntry_formatter___closed__9; static lean_object* l___regBuiltin_Lean_Parser_Attr_class_declRange___closed__3; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Attr___hyg_4____closed__13; LEAN_EXPORT lean_object* l_Lean_Parser_Attr_class_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -403,11 +419,13 @@ static lean_object* l_Lean_Parser_Attr_recursor___closed__2; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Attr___hyg_4____closed__6; static lean_object* l___regBuiltin_Lean_Parser_Priority_numPrio___closed__2; static lean_object* l_Lean_Parser_Attr_export_parenthesizer___closed__2; +static lean_object* l_Lean_Parser_Attr_externEntry_parenthesizer___closed__9; static lean_object* l_Lean_Parser_Attr_recursor___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Attr_macro_declRange___closed__5; 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_Attr_class___closed__5; 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_Attr_specialize_formatter___closed__7; static lean_object* l_Lean_Parser_Attr_default__instance___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Attr_instance_declRange___closed__7; extern lean_object* l_Lean_Parser_maxPrec; @@ -418,6 +436,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Attr_class_declRange___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Attr_specialize_declRange___closed__6; static lean_object* l_Lean_Parser_Attr_instance___closed__8; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Attr_macro_declRange(lean_object*); +static lean_object* l_Lean_Parser_Attr_externEntry___closed__13; static lean_object* l_Lean_Parser_Attr_macro___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Attr_class_declRange(lean_object*); static lean_object* l_Lean_Parser_Attr_default__instance___closed__1; @@ -443,6 +462,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Attr_recursor_declRange___closed_ static lean_object* l_Lean_Parser_Attr_externEntry___closed__8; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Attr_default__instance_parenthesizer(lean_object*); extern lean_object* l_Lean_Parser_numLit; +static lean_object* l_Lean_Parser_Attr_extern_parenthesizer___closed__9; static lean_object* l_Lean_Parser_Attr_recursor___closed__9; extern lean_object* l_Lean_PrettyPrinter_formatterAttribute; lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); @@ -473,6 +493,7 @@ static lean_object* l_Lean_Parser_Attr_instance_parenthesizer___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Attr_recursor_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Attr_macro___closed__9; LEAN_EXPORT lean_object* l_Lean_Parser_Attr_export_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Attr_simple_formatter___closed__9; static lean_object* l_Lean_Parser_Attr_default__instance___closed__4; static lean_object* l_Lean_Parser_Attr_instance_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Attr_default__instance_formatter___closed__2; @@ -1249,49 +1270,59 @@ return x_3; static lean_object* _init_l_Lean_Parser_Attr_simple___closed__6() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_skip; +x_2 = l_Lean_Parser_Attr_simple___closed__5; +x_3 = l_Lean_Parser_andthen(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Attr_simple___closed__7() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Attr_simple___closed__5; +x_1 = l_Lean_Parser_Attr_simple___closed__6; x_2 = l_Lean_Parser_optional(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Attr_simple___closed__7() { +static lean_object* _init_l_Lean_Parser_Attr_simple___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_ident; -x_2 = l_Lean_Parser_Attr_simple___closed__6; +x_2 = l_Lean_Parser_Attr_simple___closed__7; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Attr_simple___closed__8() { +static lean_object* _init_l_Lean_Parser_Attr_simple___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Attr_simple___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Attr_simple___closed__7; +x_3 = l_Lean_Parser_Attr_simple___closed__8; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Attr_simple___closed__9() { +static lean_object* _init_l_Lean_Parser_Attr_simple___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Attr_simple___closed__3; -x_2 = l_Lean_Parser_Attr_simple___closed__8; +x_2 = l_Lean_Parser_Attr_simple___closed__9; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Attr_simple___closed__10() { +static lean_object* _init_l_Lean_Parser_Attr_simple___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Attr_simple___closed__2; -x_2 = l_Lean_Parser_Attr_simple___closed__9; +x_2 = l_Lean_Parser_Attr_simple___closed__10; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -1300,7 +1331,7 @@ static lean_object* _init_l_Lean_Parser_Attr_simple() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Attr_simple___closed__10; +x_1 = l_Lean_Parser_Attr_simple___closed__11; return x_1; } } @@ -1334,7 +1365,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Attr_simple_declRange___clo { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_unsigned_to_nat(35u); -x_2 = lean_unsigned_to_nat(100u); +x_2 = lean_unsigned_to_nat(113u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -1348,7 +1379,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Parser_Attr_simple_declRange___closed__1; x_2 = lean_unsigned_to_nat(23u); x_3 = l___regBuiltin_Lean_Parser_Attr_simple_declRange___closed__2; -x_4 = lean_unsigned_to_nat(100u); +x_4 = lean_unsigned_to_nat(113u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -1468,32 +1499,52 @@ return x_3; static lean_object* _init_l_Lean_Parser_Attr_simple_formatter___closed__5() { _start: { +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_ppSpace_formatter___boxed), 5, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Attr_simple_formatter___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Attr_simple_formatter___closed__5; +x_2 = l_Lean_Parser_Attr_simple_formatter___closed__4; +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; +} +} +static lean_object* _init_l_Lean_Parser_Attr_simple_formatter___closed__7() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Attr_simple_formatter___closed__4; +x_1 = l_Lean_Parser_Attr_simple_formatter___closed__6; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_optional_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Attr_simple_formatter___closed__6() { +static lean_object* _init_l_Lean_Parser_Attr_simple_formatter___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Attr_simple_formatter___closed__3; -x_2 = l_Lean_Parser_Attr_simple_formatter___closed__5; +x_2 = l_Lean_Parser_Attr_simple_formatter___closed__7; 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; } } -static lean_object* _init_l_Lean_Parser_Attr_simple_formatter___closed__7() { +static lean_object* _init_l_Lean_Parser_Attr_simple_formatter___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Attr_simple___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Attr_simple_formatter___closed__6; +x_3 = l_Lean_Parser_Attr_simple_formatter___closed__8; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -1506,7 +1557,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Attr_simple_formatter(lean_object* x_1, l { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Attr_simple_formatter___closed__1; -x_7 = l_Lean_Parser_Attr_simple_formatter___closed__7; +x_7 = l_Lean_Parser_Attr_simple_formatter___closed__9; 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; } @@ -1611,32 +1662,52 @@ return x_3; static lean_object* _init_l_Lean_Parser_Attr_simple_parenthesizer___closed__5() { _start: { +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Parser_ppSpace_parenthesizer___boxed), 4, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Attr_simple_parenthesizer___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Attr_simple_parenthesizer___closed__5; +x_2 = l_Lean_Parser_Attr_simple_parenthesizer___closed__4; +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; +} +} +static lean_object* _init_l_Lean_Parser_Attr_simple_parenthesizer___closed__7() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Attr_simple_parenthesizer___closed__4; +x_1 = l_Lean_Parser_Attr_simple_parenthesizer___closed__6; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_optional_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Attr_simple_parenthesizer___closed__6() { +static lean_object* _init_l_Lean_Parser_Attr_simple_parenthesizer___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Attr_simple_parenthesizer___closed__3; -x_2 = l_Lean_Parser_Attr_simple_parenthesizer___closed__5; +x_2 = l_Lean_Parser_Attr_simple_parenthesizer___closed__7; 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; } } -static lean_object* _init_l_Lean_Parser_Attr_simple_parenthesizer___closed__7() { +static lean_object* _init_l_Lean_Parser_Attr_simple_parenthesizer___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Attr_simple___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Attr_simple_parenthesizer___closed__6; +x_3 = l_Lean_Parser_Attr_simple_parenthesizer___closed__8; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -1649,7 +1720,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Attr_simple_parenthesizer(lean_object* x_ { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Attr_simple_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Attr_simple_parenthesizer___closed__7; +x_7 = l_Lean_Parser_Attr_simple_parenthesizer___closed__9; 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; } @@ -3340,49 +3411,59 @@ return x_2; static lean_object* _init_l_Lean_Parser_Attr_instance___closed__5() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_skip; +x_2 = l_Lean_Parser_Attr_simple___closed__4; +x_3 = l_Lean_Parser_andthen(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Attr_instance___closed__6() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Attr_simple___closed__4; +x_1 = l_Lean_Parser_Attr_instance___closed__5; x_2 = l_Lean_Parser_optional(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Attr_instance___closed__6() { +static lean_object* _init_l_Lean_Parser_Attr_instance___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Attr_instance___closed__4; -x_2 = l_Lean_Parser_Attr_instance___closed__5; +x_2 = l_Lean_Parser_Attr_instance___closed__6; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Attr_instance___closed__7() { +static lean_object* _init_l_Lean_Parser_Attr_instance___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Attr_instance___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Attr_instance___closed__6; +x_3 = l_Lean_Parser_Attr_instance___closed__7; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Attr_instance___closed__8() { +static lean_object* _init_l_Lean_Parser_Attr_instance___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Attr_instance___closed__3; -x_2 = l_Lean_Parser_Attr_instance___closed__7; +x_2 = l_Lean_Parser_Attr_instance___closed__8; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Attr_instance___closed__9() { +static lean_object* _init_l_Lean_Parser_Attr_instance___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Attr_instance___closed__2; -x_2 = l_Lean_Parser_Attr_instance___closed__8; +x_2 = l_Lean_Parser_Attr_instance___closed__9; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -3391,7 +3472,7 @@ static lean_object* _init_l_Lean_Parser_Attr_instance() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Attr_instance___closed__9; +x_1 = l_Lean_Parser_Attr_instance___closed__10; return x_1; } } @@ -3425,7 +3506,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Attr_instance_declRange___c { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_unsigned_to_nat(43u); -x_2 = lean_unsigned_to_nat(99u); +x_2 = lean_unsigned_to_nat(112u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -3439,7 +3520,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Parser_Attr_instance_declRange___closed__1; x_2 = lean_unsigned_to_nat(23u); x_3 = l___regBuiltin_Lean_Parser_Attr_instance_declRange___closed__2; -x_4 = lean_unsigned_to_nat(99u); +x_4 = lean_unsigned_to_nat(112u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -3541,32 +3622,44 @@ return x_2; static lean_object* _init_l_Lean_Parser_Attr_instance_formatter___closed__3() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Attr_simple_formatter___closed__5; +x_2 = l_Lean_Parser_Attr_simple_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; +} +} +static lean_object* _init_l_Lean_Parser_Attr_instance_formatter___closed__4() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Attr_simple_formatter___closed__2; +x_1 = l_Lean_Parser_Attr_instance_formatter___closed__3; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_optional_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Attr_instance_formatter___closed__4() { +static lean_object* _init_l_Lean_Parser_Attr_instance_formatter___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Attr_instance_formatter___closed__2; -x_2 = l_Lean_Parser_Attr_instance_formatter___closed__3; +x_2 = l_Lean_Parser_Attr_instance_formatter___closed__4; 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; } } -static lean_object* _init_l_Lean_Parser_Attr_instance_formatter___closed__5() { +static lean_object* _init_l_Lean_Parser_Attr_instance_formatter___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Attr_instance___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Attr_instance_formatter___closed__4; +x_3 = l_Lean_Parser_Attr_instance_formatter___closed__5; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -3579,7 +3672,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Attr_instance_formatter(lean_object* x_1, { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Attr_instance_formatter___closed__1; -x_7 = l_Lean_Parser_Attr_instance_formatter___closed__5; +x_7 = l_Lean_Parser_Attr_instance_formatter___closed__6; 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; } @@ -3648,32 +3741,44 @@ return x_2; static lean_object* _init_l_Lean_Parser_Attr_instance_parenthesizer___closed__3() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Attr_simple_parenthesizer___closed__5; +x_2 = l_Lean_Parser_Attr_simple_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; +} +} +static lean_object* _init_l_Lean_Parser_Attr_instance_parenthesizer___closed__4() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Attr_simple_parenthesizer___closed__2; +x_1 = l_Lean_Parser_Attr_instance_parenthesizer___closed__3; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_optional_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Attr_instance_parenthesizer___closed__4() { +static lean_object* _init_l_Lean_Parser_Attr_instance_parenthesizer___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Attr_instance_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Attr_instance_parenthesizer___closed__3; +x_2 = l_Lean_Parser_Attr_instance_parenthesizer___closed__4; 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; } } -static lean_object* _init_l_Lean_Parser_Attr_instance_parenthesizer___closed__5() { +static lean_object* _init_l_Lean_Parser_Attr_instance_parenthesizer___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Attr_instance___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Attr_instance_parenthesizer___closed__4; +x_3 = l_Lean_Parser_Attr_instance_parenthesizer___closed__5; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -3686,7 +3791,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Attr_instance_parenthesizer(lean_object* { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Attr_instance_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Attr_instance_parenthesizer___closed__5; +x_7 = l_Lean_Parser_Attr_instance_parenthesizer___closed__6; 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; } @@ -3759,58 +3864,50 @@ return x_5; static lean_object* _init_l_Lean_Parser_Attr_default__instance___closed__4() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("default_instance ", 17); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Attr_default__instance___closed__5() { -_start: -{ lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Attr_default__instance___closed__4; +x_1 = l_Lean_Parser_Attr_default__instance___closed__1; x_2 = 0; x_3 = l_Lean_Parser_nonReservedSymbol(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Attr_default__instance___closed__6() { +static lean_object* _init_l_Lean_Parser_Attr_default__instance___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Attr_default__instance___closed__5; -x_2 = l_Lean_Parser_Attr_instance___closed__5; +x_1 = l_Lean_Parser_Attr_default__instance___closed__4; +x_2 = l_Lean_Parser_Attr_instance___closed__6; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Attr_default__instance___closed__7() { +static lean_object* _init_l_Lean_Parser_Attr_default__instance___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Attr_default__instance___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Attr_default__instance___closed__6; +x_3 = l_Lean_Parser_Attr_default__instance___closed__5; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Attr_default__instance___closed__8() { +static lean_object* _init_l_Lean_Parser_Attr_default__instance___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Attr_default__instance___closed__3; -x_2 = l_Lean_Parser_Attr_default__instance___closed__7; +x_2 = l_Lean_Parser_Attr_default__instance___closed__6; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Attr_default__instance___closed__9() { +static lean_object* _init_l_Lean_Parser_Attr_default__instance___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Attr_default__instance___closed__2; -x_2 = l_Lean_Parser_Attr_default__instance___closed__8; +x_2 = l_Lean_Parser_Attr_default__instance___closed__7; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -3819,7 +3916,7 @@ static lean_object* _init_l_Lean_Parser_Attr_default__instance() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Attr_default__instance___closed__9; +x_1 = l_Lean_Parser_Attr_default__instance___closed__8; return x_1; } } @@ -3853,7 +3950,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Attr_default__instance_decl { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_unsigned_to_nat(44u); -x_2 = lean_unsigned_to_nat(126u); +x_2 = lean_unsigned_to_nat(138u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -3867,7 +3964,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Parser_Attr_default__instance_declRange___closed__1; x_2 = lean_unsigned_to_nat(23u); x_3 = l___regBuiltin_Lean_Parser_Attr_default__instance_declRange___closed__2; -x_4 = lean_unsigned_to_nat(126u); +x_4 = lean_unsigned_to_nat(138u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -3960,7 +4057,7 @@ static lean_object* _init_l_Lean_Parser_Attr_default__instance_formatter___close _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Attr_default__instance___closed__4; +x_1 = l_Lean_Parser_Attr_default__instance___closed__1; x_2 = 0; x_3 = lean_box(x_2); x_4 = lean_alloc_closure((void*)(l_Lean_Parser_nonReservedSymbol_formatter___boxed), 7, 2); @@ -3974,7 +4071,7 @@ static lean_object* _init_l_Lean_Parser_Attr_default__instance_formatter___close { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Attr_default__instance_formatter___closed__2; -x_2 = l_Lean_Parser_Attr_instance_formatter___closed__3; +x_2 = l_Lean_Parser_Attr_instance_formatter___closed__4; 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); @@ -4060,7 +4157,7 @@ static lean_object* _init_l_Lean_Parser_Attr_default__instance_parenthesizer___c _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Attr_default__instance___closed__4; +x_1 = l_Lean_Parser_Attr_default__instance___closed__1; x_2 = 0; x_3 = lean_box(x_2); x_4 = lean_alloc_closure((void*)(l_Lean_Parser_nonReservedSymbol_parenthesizer___boxed), 7, 2); @@ -4074,7 +4171,7 @@ static lean_object* _init_l_Lean_Parser_Attr_default__instance_parenthesizer___c { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Attr_default__instance_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Attr_instance_parenthesizer___closed__3; +x_2 = l_Lean_Parser_Attr_instance_parenthesizer___closed__4; 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); @@ -4193,49 +4290,59 @@ return x_3; static lean_object* _init_l_Lean_Parser_Attr_specialize___closed__6() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_skip; +x_2 = l_Lean_Parser_Attr_specialize___closed__5; +x_3 = l_Lean_Parser_andthen(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Attr_specialize___closed__7() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Attr_specialize___closed__5; +x_1 = l_Lean_Parser_Attr_specialize___closed__6; x_2 = l_Lean_Parser_many(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Attr_specialize___closed__7() { +static lean_object* _init_l_Lean_Parser_Attr_specialize___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Attr_specialize___closed__4; -x_2 = l_Lean_Parser_Attr_specialize___closed__6; +x_2 = l_Lean_Parser_Attr_specialize___closed__7; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Attr_specialize___closed__8() { +static lean_object* _init_l_Lean_Parser_Attr_specialize___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Attr_specialize___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Attr_specialize___closed__7; +x_3 = l_Lean_Parser_Attr_specialize___closed__8; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Attr_specialize___closed__9() { +static lean_object* _init_l_Lean_Parser_Attr_specialize___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Attr_specialize___closed__3; -x_2 = l_Lean_Parser_Attr_specialize___closed__8; +x_2 = l_Lean_Parser_Attr_specialize___closed__9; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Attr_specialize___closed__10() { +static lean_object* _init_l_Lean_Parser_Attr_specialize___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Attr_specialize___closed__2; -x_2 = l_Lean_Parser_Attr_specialize___closed__9; +x_2 = l_Lean_Parser_Attr_specialize___closed__10; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -4244,7 +4351,7 @@ static lean_object* _init_l_Lean_Parser_Attr_specialize() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Attr_specialize___closed__10; +x_1 = l_Lean_Parser_Attr_specialize___closed__11; return x_1; } } @@ -4278,7 +4385,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Attr_specialize_declRange__ { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_unsigned_to_nat(45u); -x_2 = lean_unsigned_to_nat(121u); +x_2 = lean_unsigned_to_nat(134u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -4292,7 +4399,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Parser_Attr_specialize_declRange___closed__1; x_2 = lean_unsigned_to_nat(23u); x_3 = l___regBuiltin_Lean_Parser_Attr_specialize_declRange___closed__2; -x_4 = lean_unsigned_to_nat(121u); +x_4 = lean_unsigned_to_nat(134u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -4409,32 +4516,44 @@ return x_3; static lean_object* _init_l_Lean_Parser_Attr_specialize_formatter___closed__4() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Attr_simple_formatter___closed__5; +x_2 = l_Lean_Parser_Attr_specialize_formatter___closed__3; +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; +} +} +static lean_object* _init_l_Lean_Parser_Attr_specialize_formatter___closed__5() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Attr_specialize_formatter___closed__3; +x_1 = l_Lean_Parser_Attr_specialize_formatter___closed__4; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_many_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Attr_specialize_formatter___closed__5() { +static lean_object* _init_l_Lean_Parser_Attr_specialize_formatter___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Attr_specialize_formatter___closed__2; -x_2 = l_Lean_Parser_Attr_specialize_formatter___closed__4; +x_2 = l_Lean_Parser_Attr_specialize_formatter___closed__5; 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; } } -static lean_object* _init_l_Lean_Parser_Attr_specialize_formatter___closed__6() { +static lean_object* _init_l_Lean_Parser_Attr_specialize_formatter___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Attr_specialize___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Attr_specialize_formatter___closed__5; +x_3 = l_Lean_Parser_Attr_specialize_formatter___closed__6; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -4447,7 +4566,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Attr_specialize_formatter(lean_object* x_ { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Attr_specialize_formatter___closed__1; -x_7 = l_Lean_Parser_Attr_specialize_formatter___closed__6; +x_7 = l_Lean_Parser_Attr_specialize_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; } @@ -4531,32 +4650,44 @@ return x_3; static lean_object* _init_l_Lean_Parser_Attr_specialize_parenthesizer___closed__4() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Attr_simple_parenthesizer___closed__5; +x_2 = l_Lean_Parser_Attr_specialize_parenthesizer___closed__3; +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; +} +} +static lean_object* _init_l_Lean_Parser_Attr_specialize_parenthesizer___closed__5() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Attr_specialize_parenthesizer___closed__3; +x_1 = l_Lean_Parser_Attr_specialize_parenthesizer___closed__4; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_many_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Attr_specialize_parenthesizer___closed__5() { +static lean_object* _init_l_Lean_Parser_Attr_specialize_parenthesizer___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Attr_specialize_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Attr_specialize_parenthesizer___closed__4; +x_2 = l_Lean_Parser_Attr_specialize_parenthesizer___closed__5; 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; } } -static lean_object* _init_l_Lean_Parser_Attr_specialize_parenthesizer___closed__6() { +static lean_object* _init_l_Lean_Parser_Attr_specialize_parenthesizer___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Attr_specialize___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Attr_specialize_parenthesizer___closed__5; +x_3 = l_Lean_Parser_Attr_specialize_parenthesizer___closed__6; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -4569,7 +4700,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Attr_specialize_parenthesizer(lean_object { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Attr_specialize_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Attr_specialize_parenthesizer___closed__6; +x_7 = l_Lean_Parser_Attr_specialize_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; } @@ -4642,13 +4773,23 @@ return x_5; static lean_object* _init_l_Lean_Parser_Attr_externEntry___closed__4() { _start: { -lean_object* x_1; lean_object* x_2; +lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_ident; +x_2 = l_Lean_Parser_skip; +x_3 = l_Lean_Parser_andthen(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Attr_externEntry___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Attr_externEntry___closed__4; x_2 = l_Lean_Parser_optional(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Attr_externEntry___closed__5() { +static lean_object* _init_l_Lean_Parser_Attr_externEntry___closed__6() { _start: { lean_object* x_1; @@ -4656,72 +4797,72 @@ x_1 = lean_mk_string_from_bytes("inline ", 7); return x_1; } } -static lean_object* _init_l_Lean_Parser_Attr_externEntry___closed__6() { +static lean_object* _init_l_Lean_Parser_Attr_externEntry___closed__7() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Attr_externEntry___closed__5; +x_1 = l_Lean_Parser_Attr_externEntry___closed__6; x_2 = 0; x_3 = l_Lean_Parser_nonReservedSymbol(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Attr_externEntry___closed__7() { +static lean_object* _init_l_Lean_Parser_Attr_externEntry___closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Attr_externEntry___closed__6; +x_1 = l_Lean_Parser_Attr_externEntry___closed__7; x_2 = l_Lean_Parser_optional(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Attr_externEntry___closed__8() { +static lean_object* _init_l_Lean_Parser_Attr_externEntry___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Attr_externEntry___closed__7; +x_1 = l_Lean_Parser_Attr_externEntry___closed__8; x_2 = l_Lean_Parser_strLit; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Attr_externEntry___closed__9() { +static lean_object* _init_l_Lean_Parser_Attr_externEntry___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Attr_externEntry___closed__4; -x_2 = l_Lean_Parser_Attr_externEntry___closed__8; +x_1 = l_Lean_Parser_Attr_externEntry___closed__5; +x_2 = l_Lean_Parser_Attr_externEntry___closed__9; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Attr_externEntry___closed__10() { +static lean_object* _init_l_Lean_Parser_Attr_externEntry___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Attr_externEntry___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Attr_externEntry___closed__9; +x_3 = l_Lean_Parser_Attr_externEntry___closed__10; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Attr_externEntry___closed__11() { +static lean_object* _init_l_Lean_Parser_Attr_externEntry___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Attr_externEntry___closed__3; -x_2 = l_Lean_Parser_Attr_externEntry___closed__10; +x_2 = l_Lean_Parser_Attr_externEntry___closed__11; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Attr_externEntry___closed__12() { +static lean_object* _init_l_Lean_Parser_Attr_externEntry___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Attr_externEntry___closed__2; -x_2 = l_Lean_Parser_Attr_externEntry___closed__11; +x_2 = l_Lean_Parser_Attr_externEntry___closed__12; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -4730,7 +4871,7 @@ static lean_object* _init_l_Lean_Parser_Attr_externEntry() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Attr_externEntry___closed__12; +x_1 = l_Lean_Parser_Attr_externEntry___closed__13; return x_1; } } @@ -4769,18 +4910,20 @@ return x_5; static lean_object* _init_l_Lean_Parser_Attr_extern___closed__4() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("extern ", 7); -return x_1; +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Attr_extern___closed__1; +x_2 = 0; +x_3 = l_Lean_Parser_nonReservedSymbol(x_1, x_2); +return x_3; } } static lean_object* _init_l_Lean_Parser_Attr_extern___closed__5() { _start: { -lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Attr_extern___closed__4; -x_2 = 0; -x_3 = l_Lean_Parser_nonReservedSymbol(x_1, x_2); +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_skip; +x_2 = l_Lean_Parser_numLit; +x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } @@ -4788,7 +4931,7 @@ static lean_object* _init_l_Lean_Parser_Attr_extern___closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_numLit; +x_1 = l_Lean_Parser_Attr_extern___closed__5; x_2 = l_Lean_Parser_optional(x_1); return x_2; } @@ -4796,59 +4939,69 @@ return x_2; static lean_object* _init_l_Lean_Parser_Attr_extern___closed__7() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_skip; +x_2 = l_Lean_Parser_Attr_externEntry; +x_3 = l_Lean_Parser_andthen(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Attr_extern___closed__8() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Attr_externEntry; +x_1 = l_Lean_Parser_Attr_extern___closed__7; x_2 = l_Lean_Parser_many(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Attr_extern___closed__8() { +static lean_object* _init_l_Lean_Parser_Attr_extern___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Attr_extern___closed__6; -x_2 = l_Lean_Parser_Attr_extern___closed__7; +x_2 = l_Lean_Parser_Attr_extern___closed__8; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Attr_extern___closed__9() { +static lean_object* _init_l_Lean_Parser_Attr_extern___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Attr_extern___closed__5; -x_2 = l_Lean_Parser_Attr_extern___closed__8; +x_1 = l_Lean_Parser_Attr_extern___closed__4; +x_2 = l_Lean_Parser_Attr_extern___closed__9; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Attr_extern___closed__10() { +static lean_object* _init_l_Lean_Parser_Attr_extern___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Attr_extern___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Attr_extern___closed__9; +x_3 = l_Lean_Parser_Attr_extern___closed__10; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Attr_extern___closed__11() { +static lean_object* _init_l_Lean_Parser_Attr_extern___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Attr_extern___closed__3; -x_2 = l_Lean_Parser_Attr_extern___closed__10; +x_2 = l_Lean_Parser_Attr_extern___closed__11; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Attr_extern___closed__12() { +static lean_object* _init_l_Lean_Parser_Attr_extern___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Attr_extern___closed__2; -x_2 = l_Lean_Parser_Attr_extern___closed__11; +x_2 = l_Lean_Parser_Attr_extern___closed__12; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -4857,7 +5010,7 @@ static lean_object* _init_l_Lean_Parser_Attr_extern() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Attr_extern___closed__12; +x_1 = l_Lean_Parser_Attr_extern___closed__13; return x_1; } } @@ -4878,7 +5031,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Attr_extern_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(48u); +x_1 = lean_unsigned_to_nat(49u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -4890,8 +5043,8 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Attr_extern_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(48u); -x_2 = lean_unsigned_to_nat(122u); +x_1 = lean_unsigned_to_nat(50u); +x_2 = lean_unsigned_to_nat(93u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -4905,7 +5058,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Parser_Attr_extern_declRange___closed__1; x_2 = lean_unsigned_to_nat(23u); x_3 = l___regBuiltin_Lean_Parser_Attr_extern_declRange___closed__2; -x_4 = lean_unsigned_to_nat(122u); +x_4 = lean_unsigned_to_nat(93u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -4918,7 +5071,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Attr_extern_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(48u); +x_1 = lean_unsigned_to_nat(49u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -4930,7 +5083,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Attr_extern_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(48u); +x_1 = lean_unsigned_to_nat(49u); x_2 = lean_unsigned_to_nat(33u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -4997,18 +5150,30 @@ return x_7; static lean_object* _init_l_Lean_Parser_Attr_externEntry_formatter___closed__2() { _start: { -lean_object* x_1; lean_object* x_2; +lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Attr_simple_formatter___closed__3; +x_2 = l_Lean_Parser_Attr_simple_formatter___closed__5; +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; +} +} +static lean_object* _init_l_Lean_Parser_Attr_externEntry_formatter___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Attr_externEntry_formatter___closed__2; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_optional_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Attr_externEntry_formatter___closed__3() { +static lean_object* _init_l_Lean_Parser_Attr_externEntry_formatter___closed__4() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Attr_externEntry___closed__5; +x_1 = l_Lean_Parser_Attr_externEntry___closed__6; x_2 = 0; x_3 = lean_box(x_2); x_4 = lean_alloc_closure((void*)(l_Lean_Parser_nonReservedSymbol_formatter___boxed), 7, 2); @@ -5017,17 +5182,17 @@ lean_closure_set(x_4, 1, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Attr_externEntry_formatter___closed__4() { +static lean_object* _init_l_Lean_Parser_Attr_externEntry_formatter___closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Attr_externEntry_formatter___closed__3; +x_1 = l_Lean_Parser_Attr_externEntry_formatter___closed__4; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_optional_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Attr_externEntry_formatter___closed__5() { +static lean_object* _init_l_Lean_Parser_Attr_externEntry_formatter___closed__6() { _start: { lean_object* x_1; @@ -5035,37 +5200,37 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_strLit_formatter), 5, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_Attr_externEntry_formatter___closed__6() { +static lean_object* _init_l_Lean_Parser_Attr_externEntry_formatter___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Attr_externEntry_formatter___closed__4; -x_2 = l_Lean_Parser_Attr_externEntry_formatter___closed__5; +x_1 = l_Lean_Parser_Attr_externEntry_formatter___closed__5; +x_2 = l_Lean_Parser_Attr_externEntry_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; } } -static lean_object* _init_l_Lean_Parser_Attr_externEntry_formatter___closed__7() { +static lean_object* _init_l_Lean_Parser_Attr_externEntry_formatter___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Attr_externEntry_formatter___closed__2; -x_2 = l_Lean_Parser_Attr_externEntry_formatter___closed__6; +x_1 = l_Lean_Parser_Attr_externEntry_formatter___closed__3; +x_2 = l_Lean_Parser_Attr_externEntry_formatter___closed__7; 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; } } -static lean_object* _init_l_Lean_Parser_Attr_externEntry_formatter___closed__8() { +static lean_object* _init_l_Lean_Parser_Attr_externEntry_formatter___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Attr_externEntry___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Attr_externEntry_formatter___closed__7; +x_3 = l_Lean_Parser_Attr_externEntry_formatter___closed__8; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -5078,7 +5243,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Attr_externEntry_formatter(lean_object* x { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Attr_externEntry_formatter___closed__1; -x_7 = l_Lean_Parser_Attr_externEntry_formatter___closed__8; +x_7 = l_Lean_Parser_Attr_externEntry_formatter___closed__9; 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; } @@ -5138,7 +5303,7 @@ static lean_object* _init_l_Lean_Parser_Attr_extern_formatter___closed__2() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Attr_extern___closed__4; +x_1 = l_Lean_Parser_Attr_extern___closed__1; x_2 = 0; x_3 = lean_box(x_2); x_4 = lean_alloc_closure((void*)(l_Lean_Parser_nonReservedSymbol_formatter___boxed), 7, 2); @@ -5150,54 +5315,78 @@ return x_4; static lean_object* _init_l_Lean_Parser_Attr_extern_formatter___closed__3() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Attr_simple_formatter___closed__5; +x_2 = l_Lean_Parser_Priority_numPrio_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; +} +} +static lean_object* _init_l_Lean_Parser_Attr_extern_formatter___closed__4() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Priority_numPrio_formatter___closed__2; +x_1 = l_Lean_Parser_Attr_extern_formatter___closed__3; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_optional_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Attr_extern_formatter___closed__4() { +static lean_object* _init_l_Lean_Parser_Attr_extern_formatter___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Attr_simple_formatter___closed__5; +x_2 = l___regBuiltin_Lean_Parser_Attr_externEntry_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; +} +} +static lean_object* _init_l_Lean_Parser_Attr_extern_formatter___closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___regBuiltin_Lean_Parser_Attr_externEntry_formatter___closed__2; +x_1 = l_Lean_Parser_Attr_extern_formatter___closed__5; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_many_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Attr_extern_formatter___closed__5() { +static lean_object* _init_l_Lean_Parser_Attr_extern_formatter___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Attr_extern_formatter___closed__3; -x_2 = l_Lean_Parser_Attr_extern_formatter___closed__4; +x_1 = l_Lean_Parser_Attr_extern_formatter___closed__4; +x_2 = l_Lean_Parser_Attr_extern_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; } } -static lean_object* _init_l_Lean_Parser_Attr_extern_formatter___closed__6() { +static lean_object* _init_l_Lean_Parser_Attr_extern_formatter___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Attr_extern_formatter___closed__2; -x_2 = l_Lean_Parser_Attr_extern_formatter___closed__5; +x_2 = l_Lean_Parser_Attr_extern_formatter___closed__7; 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; } } -static lean_object* _init_l_Lean_Parser_Attr_extern_formatter___closed__7() { +static lean_object* _init_l_Lean_Parser_Attr_extern_formatter___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Attr_extern___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Attr_extern_formatter___closed__6; +x_3 = l_Lean_Parser_Attr_extern_formatter___closed__8; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -5210,7 +5399,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Attr_extern_formatter(lean_object* x_1, l { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Attr_extern_formatter___closed__1; -x_7 = l_Lean_Parser_Attr_extern_formatter___closed__7; +x_7 = l_Lean_Parser_Attr_extern_formatter___closed__9; 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; } @@ -5269,18 +5458,30 @@ return x_7; static lean_object* _init_l_Lean_Parser_Attr_externEntry_parenthesizer___closed__2() { _start: { -lean_object* x_1; lean_object* x_2; +lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Attr_simple_parenthesizer___closed__3; +x_2 = l_Lean_Parser_Attr_simple_parenthesizer___closed__5; +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; +} +} +static lean_object* _init_l_Lean_Parser_Attr_externEntry_parenthesizer___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Attr_externEntry_parenthesizer___closed__2; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_optional_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Attr_externEntry_parenthesizer___closed__3() { +static lean_object* _init_l_Lean_Parser_Attr_externEntry_parenthesizer___closed__4() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Attr_externEntry___closed__5; +x_1 = l_Lean_Parser_Attr_externEntry___closed__6; x_2 = 0; x_3 = lean_box(x_2); x_4 = lean_alloc_closure((void*)(l_Lean_Parser_nonReservedSymbol_parenthesizer___boxed), 7, 2); @@ -5289,17 +5490,17 @@ lean_closure_set(x_4, 1, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Attr_externEntry_parenthesizer___closed__4() { +static lean_object* _init_l_Lean_Parser_Attr_externEntry_parenthesizer___closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Attr_externEntry_parenthesizer___closed__3; +x_1 = l_Lean_Parser_Attr_externEntry_parenthesizer___closed__4; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_optional_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Attr_externEntry_parenthesizer___closed__5() { +static lean_object* _init_l_Lean_Parser_Attr_externEntry_parenthesizer___closed__6() { _start: { lean_object* x_1; @@ -5307,37 +5508,37 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_strLit_parenthesizer), 5, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_Attr_externEntry_parenthesizer___closed__6() { +static lean_object* _init_l_Lean_Parser_Attr_externEntry_parenthesizer___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Attr_externEntry_parenthesizer___closed__4; -x_2 = l_Lean_Parser_Attr_externEntry_parenthesizer___closed__5; +x_1 = l_Lean_Parser_Attr_externEntry_parenthesizer___closed__5; +x_2 = l_Lean_Parser_Attr_externEntry_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; } } -static lean_object* _init_l_Lean_Parser_Attr_externEntry_parenthesizer___closed__7() { +static lean_object* _init_l_Lean_Parser_Attr_externEntry_parenthesizer___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Attr_externEntry_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Attr_externEntry_parenthesizer___closed__6; +x_1 = l_Lean_Parser_Attr_externEntry_parenthesizer___closed__3; +x_2 = l_Lean_Parser_Attr_externEntry_parenthesizer___closed__7; 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; } } -static lean_object* _init_l_Lean_Parser_Attr_externEntry_parenthesizer___closed__8() { +static lean_object* _init_l_Lean_Parser_Attr_externEntry_parenthesizer___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Attr_externEntry___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Attr_externEntry_parenthesizer___closed__7; +x_3 = l_Lean_Parser_Attr_externEntry_parenthesizer___closed__8; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -5350,7 +5551,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Attr_externEntry_parenthesizer(lean_objec { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Attr_externEntry_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Attr_externEntry_parenthesizer___closed__8; +x_7 = l_Lean_Parser_Attr_externEntry_parenthesizer___closed__9; 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; } @@ -5410,7 +5611,7 @@ static lean_object* _init_l_Lean_Parser_Attr_extern_parenthesizer___closed__2() _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Attr_extern___closed__4; +x_1 = l_Lean_Parser_Attr_extern___closed__1; x_2 = 0; x_3 = lean_box(x_2); x_4 = lean_alloc_closure((void*)(l_Lean_Parser_nonReservedSymbol_parenthesizer___boxed), 7, 2); @@ -5422,54 +5623,78 @@ return x_4; static lean_object* _init_l_Lean_Parser_Attr_extern_parenthesizer___closed__3() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Attr_simple_parenthesizer___closed__5; +x_2 = l_Lean_Parser_Priority_numPrio_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; +} +} +static lean_object* _init_l_Lean_Parser_Attr_extern_parenthesizer___closed__4() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Priority_numPrio_parenthesizer___closed__2; +x_1 = l_Lean_Parser_Attr_extern_parenthesizer___closed__3; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_optional_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Attr_extern_parenthesizer___closed__4() { +static lean_object* _init_l_Lean_Parser_Attr_extern_parenthesizer___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Attr_simple_parenthesizer___closed__5; +x_2 = l___regBuiltin_Lean_Parser_Attr_externEntry_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; +} +} +static lean_object* _init_l_Lean_Parser_Attr_extern_parenthesizer___closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___regBuiltin_Lean_Parser_Attr_externEntry_parenthesizer___closed__2; +x_1 = l_Lean_Parser_Attr_extern_parenthesizer___closed__5; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_many_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Attr_extern_parenthesizer___closed__5() { +static lean_object* _init_l_Lean_Parser_Attr_extern_parenthesizer___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Attr_extern_parenthesizer___closed__3; -x_2 = l_Lean_Parser_Attr_extern_parenthesizer___closed__4; +x_1 = l_Lean_Parser_Attr_extern_parenthesizer___closed__4; +x_2 = l_Lean_Parser_Attr_extern_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; } } -static lean_object* _init_l_Lean_Parser_Attr_extern_parenthesizer___closed__6() { +static lean_object* _init_l_Lean_Parser_Attr_extern_parenthesizer___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Attr_extern_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Attr_extern_parenthesizer___closed__5; +x_2 = l_Lean_Parser_Attr_extern_parenthesizer___closed__7; 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; } } -static lean_object* _init_l_Lean_Parser_Attr_extern_parenthesizer___closed__7() { +static lean_object* _init_l_Lean_Parser_Attr_extern_parenthesizer___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Attr_extern___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Attr_extern_parenthesizer___closed__6; +x_3 = l_Lean_Parser_Attr_extern_parenthesizer___closed__8; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -5482,7 +5707,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Attr_extern_parenthesizer(lean_object* x_ { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Attr_extern_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Attr_extern_parenthesizer___closed__7; +x_7 = l_Lean_Parser_Attr_extern_parenthesizer___closed__9; 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; } @@ -5665,6 +5890,8 @@ l_Lean_Parser_Attr_simple___closed__9 = _init_l_Lean_Parser_Attr_simple___closed lean_mark_persistent(l_Lean_Parser_Attr_simple___closed__9); l_Lean_Parser_Attr_simple___closed__10 = _init_l_Lean_Parser_Attr_simple___closed__10(); lean_mark_persistent(l_Lean_Parser_Attr_simple___closed__10); +l_Lean_Parser_Attr_simple___closed__11 = _init_l_Lean_Parser_Attr_simple___closed__11(); +lean_mark_persistent(l_Lean_Parser_Attr_simple___closed__11); l_Lean_Parser_Attr_simple = _init_l_Lean_Parser_Attr_simple(); lean_mark_persistent(l_Lean_Parser_Attr_simple); res = l___regBuiltin_Lean_Parser_Attr_simple(lean_io_mk_world()); @@ -5701,6 +5928,10 @@ l_Lean_Parser_Attr_simple_formatter___closed__6 = _init_l_Lean_Parser_Attr_simpl lean_mark_persistent(l_Lean_Parser_Attr_simple_formatter___closed__6); l_Lean_Parser_Attr_simple_formatter___closed__7 = _init_l_Lean_Parser_Attr_simple_formatter___closed__7(); lean_mark_persistent(l_Lean_Parser_Attr_simple_formatter___closed__7); +l_Lean_Parser_Attr_simple_formatter___closed__8 = _init_l_Lean_Parser_Attr_simple_formatter___closed__8(); +lean_mark_persistent(l_Lean_Parser_Attr_simple_formatter___closed__8); +l_Lean_Parser_Attr_simple_formatter___closed__9 = _init_l_Lean_Parser_Attr_simple_formatter___closed__9(); +lean_mark_persistent(l_Lean_Parser_Attr_simple_formatter___closed__9); l___regBuiltin_Lean_Parser_Attr_simple_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Attr_simple_formatter___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Attr_simple_formatter___closed__1); l___regBuiltin_Lean_Parser_Attr_simple_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Attr_simple_formatter___closed__2(); @@ -5726,6 +5957,10 @@ l_Lean_Parser_Attr_simple_parenthesizer___closed__6 = _init_l_Lean_Parser_Attr_s lean_mark_persistent(l_Lean_Parser_Attr_simple_parenthesizer___closed__6); l_Lean_Parser_Attr_simple_parenthesizer___closed__7 = _init_l_Lean_Parser_Attr_simple_parenthesizer___closed__7(); lean_mark_persistent(l_Lean_Parser_Attr_simple_parenthesizer___closed__7); +l_Lean_Parser_Attr_simple_parenthesizer___closed__8 = _init_l_Lean_Parser_Attr_simple_parenthesizer___closed__8(); +lean_mark_persistent(l_Lean_Parser_Attr_simple_parenthesizer___closed__8); +l_Lean_Parser_Attr_simple_parenthesizer___closed__9 = _init_l_Lean_Parser_Attr_simple_parenthesizer___closed__9(); +lean_mark_persistent(l_Lean_Parser_Attr_simple_parenthesizer___closed__9); l___regBuiltin_Lean_Parser_Attr_simple_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Attr_simple_parenthesizer___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Attr_simple_parenthesizer___closed__1); l___regBuiltin_Lean_Parser_Attr_simple_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Attr_simple_parenthesizer___closed__2(); @@ -6027,6 +6262,8 @@ l_Lean_Parser_Attr_instance___closed__8 = _init_l_Lean_Parser_Attr_instance___cl lean_mark_persistent(l_Lean_Parser_Attr_instance___closed__8); l_Lean_Parser_Attr_instance___closed__9 = _init_l_Lean_Parser_Attr_instance___closed__9(); lean_mark_persistent(l_Lean_Parser_Attr_instance___closed__9); +l_Lean_Parser_Attr_instance___closed__10 = _init_l_Lean_Parser_Attr_instance___closed__10(); +lean_mark_persistent(l_Lean_Parser_Attr_instance___closed__10); l_Lean_Parser_Attr_instance = _init_l_Lean_Parser_Attr_instance(); lean_mark_persistent(l_Lean_Parser_Attr_instance); res = l___regBuiltin_Lean_Parser_Attr_instance(lean_io_mk_world()); @@ -6059,6 +6296,8 @@ l_Lean_Parser_Attr_instance_formatter___closed__4 = _init_l_Lean_Parser_Attr_ins lean_mark_persistent(l_Lean_Parser_Attr_instance_formatter___closed__4); l_Lean_Parser_Attr_instance_formatter___closed__5 = _init_l_Lean_Parser_Attr_instance_formatter___closed__5(); lean_mark_persistent(l_Lean_Parser_Attr_instance_formatter___closed__5); +l_Lean_Parser_Attr_instance_formatter___closed__6 = _init_l_Lean_Parser_Attr_instance_formatter___closed__6(); +lean_mark_persistent(l_Lean_Parser_Attr_instance_formatter___closed__6); l___regBuiltin_Lean_Parser_Attr_instance_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Attr_instance_formatter___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Attr_instance_formatter___closed__1); l___regBuiltin_Lean_Parser_Attr_instance_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Attr_instance_formatter___closed__2(); @@ -6076,6 +6315,8 @@ l_Lean_Parser_Attr_instance_parenthesizer___closed__4 = _init_l_Lean_Parser_Attr lean_mark_persistent(l_Lean_Parser_Attr_instance_parenthesizer___closed__4); l_Lean_Parser_Attr_instance_parenthesizer___closed__5 = _init_l_Lean_Parser_Attr_instance_parenthesizer___closed__5(); lean_mark_persistent(l_Lean_Parser_Attr_instance_parenthesizer___closed__5); +l_Lean_Parser_Attr_instance_parenthesizer___closed__6 = _init_l_Lean_Parser_Attr_instance_parenthesizer___closed__6(); +lean_mark_persistent(l_Lean_Parser_Attr_instance_parenthesizer___closed__6); l___regBuiltin_Lean_Parser_Attr_instance_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Attr_instance_parenthesizer___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Attr_instance_parenthesizer___closed__1); l___regBuiltin_Lean_Parser_Attr_instance_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Attr_instance_parenthesizer___closed__2(); @@ -6099,8 +6340,6 @@ l_Lean_Parser_Attr_default__instance___closed__7 = _init_l_Lean_Parser_Attr_defa lean_mark_persistent(l_Lean_Parser_Attr_default__instance___closed__7); l_Lean_Parser_Attr_default__instance___closed__8 = _init_l_Lean_Parser_Attr_default__instance___closed__8(); lean_mark_persistent(l_Lean_Parser_Attr_default__instance___closed__8); -l_Lean_Parser_Attr_default__instance___closed__9 = _init_l_Lean_Parser_Attr_default__instance___closed__9(); -lean_mark_persistent(l_Lean_Parser_Attr_default__instance___closed__9); l_Lean_Parser_Attr_default__instance = _init_l_Lean_Parser_Attr_default__instance(); lean_mark_persistent(l_Lean_Parser_Attr_default__instance); res = l___regBuiltin_Lean_Parser_Attr_default__instance(lean_io_mk_world()); @@ -6173,6 +6412,8 @@ l_Lean_Parser_Attr_specialize___closed__9 = _init_l_Lean_Parser_Attr_specialize_ lean_mark_persistent(l_Lean_Parser_Attr_specialize___closed__9); l_Lean_Parser_Attr_specialize___closed__10 = _init_l_Lean_Parser_Attr_specialize___closed__10(); lean_mark_persistent(l_Lean_Parser_Attr_specialize___closed__10); +l_Lean_Parser_Attr_specialize___closed__11 = _init_l_Lean_Parser_Attr_specialize___closed__11(); +lean_mark_persistent(l_Lean_Parser_Attr_specialize___closed__11); l_Lean_Parser_Attr_specialize = _init_l_Lean_Parser_Attr_specialize(); lean_mark_persistent(l_Lean_Parser_Attr_specialize); res = l___regBuiltin_Lean_Parser_Attr_specialize(lean_io_mk_world()); @@ -6207,6 +6448,8 @@ l_Lean_Parser_Attr_specialize_formatter___closed__5 = _init_l_Lean_Parser_Attr_s lean_mark_persistent(l_Lean_Parser_Attr_specialize_formatter___closed__5); l_Lean_Parser_Attr_specialize_formatter___closed__6 = _init_l_Lean_Parser_Attr_specialize_formatter___closed__6(); lean_mark_persistent(l_Lean_Parser_Attr_specialize_formatter___closed__6); +l_Lean_Parser_Attr_specialize_formatter___closed__7 = _init_l_Lean_Parser_Attr_specialize_formatter___closed__7(); +lean_mark_persistent(l_Lean_Parser_Attr_specialize_formatter___closed__7); l___regBuiltin_Lean_Parser_Attr_specialize_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Attr_specialize_formatter___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Attr_specialize_formatter___closed__1); l___regBuiltin_Lean_Parser_Attr_specialize_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Attr_specialize_formatter___closed__2(); @@ -6226,6 +6469,8 @@ l_Lean_Parser_Attr_specialize_parenthesizer___closed__5 = _init_l_Lean_Parser_At lean_mark_persistent(l_Lean_Parser_Attr_specialize_parenthesizer___closed__5); l_Lean_Parser_Attr_specialize_parenthesizer___closed__6 = _init_l_Lean_Parser_Attr_specialize_parenthesizer___closed__6(); lean_mark_persistent(l_Lean_Parser_Attr_specialize_parenthesizer___closed__6); +l_Lean_Parser_Attr_specialize_parenthesizer___closed__7 = _init_l_Lean_Parser_Attr_specialize_parenthesizer___closed__7(); +lean_mark_persistent(l_Lean_Parser_Attr_specialize_parenthesizer___closed__7); l___regBuiltin_Lean_Parser_Attr_specialize_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Attr_specialize_parenthesizer___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Attr_specialize_parenthesizer___closed__1); l___regBuiltin_Lean_Parser_Attr_specialize_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Attr_specialize_parenthesizer___closed__2(); @@ -6257,6 +6502,8 @@ l_Lean_Parser_Attr_externEntry___closed__11 = _init_l_Lean_Parser_Attr_externEnt lean_mark_persistent(l_Lean_Parser_Attr_externEntry___closed__11); l_Lean_Parser_Attr_externEntry___closed__12 = _init_l_Lean_Parser_Attr_externEntry___closed__12(); lean_mark_persistent(l_Lean_Parser_Attr_externEntry___closed__12); +l_Lean_Parser_Attr_externEntry___closed__13 = _init_l_Lean_Parser_Attr_externEntry___closed__13(); +lean_mark_persistent(l_Lean_Parser_Attr_externEntry___closed__13); l_Lean_Parser_Attr_externEntry = _init_l_Lean_Parser_Attr_externEntry(); lean_mark_persistent(l_Lean_Parser_Attr_externEntry); l_Lean_Parser_Attr_extern___closed__1 = _init_l_Lean_Parser_Attr_extern___closed__1(); @@ -6283,6 +6530,8 @@ l_Lean_Parser_Attr_extern___closed__11 = _init_l_Lean_Parser_Attr_extern___close lean_mark_persistent(l_Lean_Parser_Attr_extern___closed__11); l_Lean_Parser_Attr_extern___closed__12 = _init_l_Lean_Parser_Attr_extern___closed__12(); lean_mark_persistent(l_Lean_Parser_Attr_extern___closed__12); +l_Lean_Parser_Attr_extern___closed__13 = _init_l_Lean_Parser_Attr_extern___closed__13(); +lean_mark_persistent(l_Lean_Parser_Attr_extern___closed__13); l_Lean_Parser_Attr_extern = _init_l_Lean_Parser_Attr_extern(); lean_mark_persistent(l_Lean_Parser_Attr_extern); res = l___regBuiltin_Lean_Parser_Attr_extern(lean_io_mk_world()); @@ -6321,6 +6570,8 @@ l_Lean_Parser_Attr_externEntry_formatter___closed__7 = _init_l_Lean_Parser_Attr_ lean_mark_persistent(l_Lean_Parser_Attr_externEntry_formatter___closed__7); l_Lean_Parser_Attr_externEntry_formatter___closed__8 = _init_l_Lean_Parser_Attr_externEntry_formatter___closed__8(); lean_mark_persistent(l_Lean_Parser_Attr_externEntry_formatter___closed__8); +l_Lean_Parser_Attr_externEntry_formatter___closed__9 = _init_l_Lean_Parser_Attr_externEntry_formatter___closed__9(); +lean_mark_persistent(l_Lean_Parser_Attr_externEntry_formatter___closed__9); l___regBuiltin_Lean_Parser_Attr_externEntry_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Attr_externEntry_formatter___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Attr_externEntry_formatter___closed__1); l___regBuiltin_Lean_Parser_Attr_externEntry_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Attr_externEntry_formatter___closed__2(); @@ -6342,6 +6593,10 @@ l_Lean_Parser_Attr_extern_formatter___closed__6 = _init_l_Lean_Parser_Attr_exter lean_mark_persistent(l_Lean_Parser_Attr_extern_formatter___closed__6); l_Lean_Parser_Attr_extern_formatter___closed__7 = _init_l_Lean_Parser_Attr_extern_formatter___closed__7(); lean_mark_persistent(l_Lean_Parser_Attr_extern_formatter___closed__7); +l_Lean_Parser_Attr_extern_formatter___closed__8 = _init_l_Lean_Parser_Attr_extern_formatter___closed__8(); +lean_mark_persistent(l_Lean_Parser_Attr_extern_formatter___closed__8); +l_Lean_Parser_Attr_extern_formatter___closed__9 = _init_l_Lean_Parser_Attr_extern_formatter___closed__9(); +lean_mark_persistent(l_Lean_Parser_Attr_extern_formatter___closed__9); l___regBuiltin_Lean_Parser_Attr_extern_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Attr_extern_formatter___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Attr_extern_formatter___closed__1); l___regBuiltin_Lean_Parser_Attr_extern_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Attr_extern_formatter___closed__2(); @@ -6365,6 +6620,8 @@ l_Lean_Parser_Attr_externEntry_parenthesizer___closed__7 = _init_l_Lean_Parser_A lean_mark_persistent(l_Lean_Parser_Attr_externEntry_parenthesizer___closed__7); l_Lean_Parser_Attr_externEntry_parenthesizer___closed__8 = _init_l_Lean_Parser_Attr_externEntry_parenthesizer___closed__8(); lean_mark_persistent(l_Lean_Parser_Attr_externEntry_parenthesizer___closed__8); +l_Lean_Parser_Attr_externEntry_parenthesizer___closed__9 = _init_l_Lean_Parser_Attr_externEntry_parenthesizer___closed__9(); +lean_mark_persistent(l_Lean_Parser_Attr_externEntry_parenthesizer___closed__9); l___regBuiltin_Lean_Parser_Attr_externEntry_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Attr_externEntry_parenthesizer___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Attr_externEntry_parenthesizer___closed__1); l___regBuiltin_Lean_Parser_Attr_externEntry_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Attr_externEntry_parenthesizer___closed__2(); @@ -6386,6 +6643,10 @@ l_Lean_Parser_Attr_extern_parenthesizer___closed__6 = _init_l_Lean_Parser_Attr_e lean_mark_persistent(l_Lean_Parser_Attr_extern_parenthesizer___closed__6); l_Lean_Parser_Attr_extern_parenthesizer___closed__7 = _init_l_Lean_Parser_Attr_extern_parenthesizer___closed__7(); lean_mark_persistent(l_Lean_Parser_Attr_extern_parenthesizer___closed__7); +l_Lean_Parser_Attr_extern_parenthesizer___closed__8 = _init_l_Lean_Parser_Attr_extern_parenthesizer___closed__8(); +lean_mark_persistent(l_Lean_Parser_Attr_extern_parenthesizer___closed__8); +l_Lean_Parser_Attr_extern_parenthesizer___closed__9 = _init_l_Lean_Parser_Attr_extern_parenthesizer___closed__9(); +lean_mark_persistent(l_Lean_Parser_Attr_extern_parenthesizer___closed__9); l___regBuiltin_Lean_Parser_Attr_extern_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Attr_extern_parenthesizer___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Attr_extern_parenthesizer___closed__1); l___regBuiltin_Lean_Parser_Attr_extern_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Attr_extern_parenthesizer___closed__2(); diff --git a/stage0/stdlib/Lean/Parser/Basic.c b/stage0/stdlib/Lean/Parser/Basic.c index 629febed3731..17436d4da356 100644 --- a/stage0/stdlib/Lean/Parser/Basic.c +++ b/stage0/stdlib/Lean/Parser/Basic.c @@ -44,7 +44,6 @@ LEAN_EXPORT lean_object* l_Lean_Parser_ParserState_keepTop(lean_object*, lean_ob LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Parser_indexed___spec__4(lean_object*); static lean_object* l_Lean_Parser_nameLitFn___closed__3; uint32_t lean_string_utf8_get(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9274_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_identFnAux(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_pushNone; LEAN_EXPORT lean_object* l_Lean_Parser_runLongestMatchParser(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -64,19 +63,19 @@ LEAN_EXPORT lean_object* l_Lean_Parser_instCoeStringParser; LEAN_EXPORT lean_object* l_Lean_Parser_checkColGt___elambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_sepByElemParser(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_sepByInfo___elambda__1(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_setLhsPrecFn___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_longestMatchFnAux___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__7; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_forArgsM___spec__2___rarg___lambda__1(size_t, lean_object*, lean_object*, lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_checkColGt___elambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_orelseFn(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_errorAtSavedPos___elambda__1(lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_withoutInfo___elambda__1(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__12; LEAN_EXPORT lean_object* l_Lean_Parser_skip___elambda__1___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_LeadingIdentBehavior_toCtorIdx___boxed(lean_object*); static lean_object* l_Lean_Parser_invalidLongestMatchParser___closed__1; static lean_object* l_Lean_Parser_strLitFnAux___closed__2; +static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__9; static lean_object* l_Lean_Parser_mkAntiquot___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_leadingParserAux(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); 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*); @@ -91,7 +90,6 @@ LEAN_EXPORT lean_object* l_Lean_Parser_LeadingIdentBehavior_toCtorIdx(uint8_t); static lean_object* l_Lean_Parser_dbgTraceStateFn___closed__3; static lean_object* l_Lean_Parser_decimalNumberFn___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_unicodeSymbolInfo___elambda__1(lean_object*); -static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_TokenMap_instForInTokenMapProdNameList(lean_object*, lean_object*); static lean_object* l_Lean_Parser_OrElseOnAntiquotBehavior_noConfusion___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_leadingParserAux___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -128,6 +126,7 @@ LEAN_EXPORT uint8_t l_Lean_Parser_binNumberFn___lambda__1(uint32_t); static lean_object* l_Lean_Parser_tokenAntiquotFn___lambda__2___closed__8; LEAN_EXPORT lean_object* l_Lean_Parser_withPositionAfterLinebreak___elambda__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_sepBy1Fn___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_array_push(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_identFnAux_parse(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_tokenAntiquotFn___lambda__2___closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_manyAux(lean_object*, lean_object*, lean_object*); @@ -145,8 +144,8 @@ LEAN_EXPORT lean_object* l_Lean_Parser_hexDigitFn___boxed(lean_object*, lean_obj uint8_t lean_usize_dec_eq(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Parser_checkWsBeforeFn(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_fieldIdx___closed__2; -static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__14; LEAN_EXPORT lean_object* l_Lean_Parser_sepBy1NoAntiquot(lean_object*, lean_object*, uint8_t); +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9301____lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_mkAtomicInfo___elambda__2___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_checkStackTop___elambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_identFn___closed__2; @@ -163,11 +162,13 @@ LEAN_EXPORT lean_object* l_Lean_Parser_epsilonInfo___elambda__2___boxed(lean_obj LEAN_EXPORT lean_object* l_Lean_Parser_fieldIdxFn(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_mkAtomicInfo___elambda__2(lean_object*); lean_object* l_String_trim(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777_(uint8_t, lean_object*); static lean_object* l_Lean_Parser_skip___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_nonReservedSymbol(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_foldArgs___spec__2___rarg(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_leadingParser___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_hexNumberFn___closed__1; +static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__8; static lean_object* l_Lean_Parser_finishCommentBlock_eoi___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_LeadingIdentBehavior_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Parser_octalNumberFn___lambda__1(uint32_t); @@ -202,7 +203,6 @@ LEAN_EXPORT lean_object* l_Lean_Parser_longestMatchStep(lean_object*, lean_objec LEAN_EXPORT lean_object* l_Lean_Parser_checkColGe___elambda__1___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_ParserState_restore(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_forArgsM___spec__2(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____boxed(lean_object*, lean_object*); lean_object* l_Lean_Parser_withCacheFn(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Parser_symbolFnAux___lambda__1(lean_object*, lean_object*); uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); @@ -217,14 +217,12 @@ LEAN_EXPORT lean_object* l_Lean_Parser_withoutInfo(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_isIdCont___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_foldArgsM(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_unicodeSymbol___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__16; static lean_object* l_Lean_Parser_epsilonInfo___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_fieldIdx; static lean_object* l_Lean_Parser_decimalNumberFn_parseOptExp___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_checkNoWsBefore___elambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_longestMatchMkResult(lean_object*, lean_object*); static lean_object* l_Lean_Parser_antiquotNestedExpr___closed__6; -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9236____lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_foldArgsM___at_Lean_Syntax_forArgsM___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_manyAux___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_many1Fn(lean_object*, lean_object*, lean_object*); @@ -242,9 +240,12 @@ static lean_object* l_Lean_Parser_checkNoImmediateColon___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_sepByElemParser___boxed(lean_object*, lean_object*); lean_object* l_Lean_Parser_ParserState_stackSize(lean_object*); static lean_object* l_Lean_Parser_strLitFn___closed__2; +static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__2; +static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_categoryParserFnExtension; LEAN_EXPORT lean_object* l_Lean_Parser_rawCh(uint32_t, uint8_t); LEAN_EXPORT lean_object* l_Lean_Parser_PrattParsingTables_trailingParsers___default; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9301____closed__1; static lean_object* l_Lean_Parser_nonReservedSymbolInfo___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_OrElseOnAntiquotBehavior_noConfusion___rarg___lambda__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_withResultOfInfo___elambda__2(lean_object*, lean_object*); @@ -256,6 +257,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_sepByNoAntiquot(lean_object*, lean_object static lean_object* l_Lean_Parser_instReprLeadingIdentBehavior___closed__1; lean_object* lean_string_push(lean_object*, uint32_t); static lean_object* l_Lean_Parser_longestMatchFn___closed__1; +static lean_object* l_Lean_Parser_hygieneInfoFn___closed__3; static lean_object* l_Lean_Parser_whitespace___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_ParserState_keepLatest___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_longestMatchFnAux_parse___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -285,7 +287,6 @@ LEAN_EXPORT lean_object* l_Lean_Parser_errorAtSavedPos(lean_object*, uint8_t); static lean_object* l_Lean_Parser_instBEqOrElseOnAntiquotBehavior___closed__1; static lean_object* l_Lean_Parser_mkAntiquot___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_updateTokenCache(lean_object*, lean_object*); -static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__13; LEAN_EXPORT lean_object* l_Lean_Parser_withAntiquotFn___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_antiquotNestedExpr___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_orelseInfo(lean_object*, lean_object*); @@ -297,7 +298,6 @@ LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at_Lean_Parser_TokenMap_insert___sp lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Parser_SyntaxNodeKindSet_insert___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_lookaheadFn(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_eoiFn___closed__1; -static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__15; LEAN_EXPORT lean_object* l_Lean_Parser_nonReservedSymbolInfo___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_trailingLoop___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_isAntiquot(lean_object*); @@ -311,7 +311,7 @@ static lean_object* l_Lean_Parser_strLitFn___closed__1; static lean_object* l_Lean_Parser_antiquotNestedExpr___closed__10; LEAN_EXPORT lean_object* l_Lean_Parser_orelse___elambda__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_ParserState_popSyntax(lean_object*); -static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__9; +LEAN_EXPORT lean_object* l_Lean_Parser_hygieneInfoFn(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_epsilonInfo; LEAN_EXPORT lean_object* l_Lean_Parser_symbolNoAntiquot(lean_object*); static lean_object* l_Lean_Parser_dbgTraceStateFn___closed__7; @@ -319,8 +319,8 @@ LEAN_EXPORT lean_object* l_Lean_Parser_eoi; static lean_object* l_Lean_Parser_quotedCharCoreFn___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_withPosition(lean_object*); static lean_object* l_Lean_Parser_incQuotDepth___closed__1; -static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__8; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9274____lambda__1___closed__1; +static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__10; +LEAN_EXPORT lean_object* l_Lean_Parser_hygieneInfoNoAntiquot; static lean_object* l_Lean_Parser_errorAtSavedPos___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_pickNonNone(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_dbgTraceStateFn___lambda__1(lean_object*, lean_object*); @@ -382,6 +382,7 @@ static lean_object* l_Lean_Parser_fieldIdx___closed__3; lean_object* l_Lean_Name_quickCmp___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_numberFnAux(lean_object*, lean_object*); uint8_t l_Substring_contains(lean_object*, uint32_t); +static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__11; LEAN_EXPORT lean_object* l_Lean_Parser_dbgTraceStateFn___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_mergeOrElseErrors(lean_object*, lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_foldArgsM___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); @@ -396,6 +397,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_ParserState_keepTop___boxed(lean_object*, LEAN_EXPORT lean_object* l_Lean_Parser_nonReservedSymbolInfo___elambda__1___boxed(lean_object*); lean_object* l_Lean_Parser_ParserState_mkErrorAt(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_mkAntiquotSplice___closed__5; +static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__19; static lean_object* l_Lean_Parser_checkPrecFn___closed__1; static lean_object* l_Lean_Parser_fieldIdx___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_OrElseOnAntiquotBehavior_noConfusion___rarg(uint8_t, uint8_t, lean_object*); @@ -412,7 +414,6 @@ LEAN_EXPORT lean_object* l_Lean_Parser_instReprLeadingIdentBehavior; uint8_t l_Lean_Parser_ParserState_hasError(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_withAntiquotFn(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_manyFn(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9236____closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_trailingNode(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_setLhsPrec___elambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_foldArgsM___at_Lean_Syntax_foldArgs___spec__1___rarg(lean_object*, lean_object*, lean_object*); @@ -420,6 +421,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_tokenFn(lean_object*, lean_object*, lean_ static lean_object* l_Lean_Parser_mkAntiquotSplice___closed__1; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_foldArgsM___spec__1___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Parser_Basic_0__Lean_Parser_beqLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8753_(uint8_t, uint8_t); LEAN_EXPORT lean_object* l_Lean_Parser_strLitFnAux(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_decimalNumberFn_parseOptDot___closed__1; static lean_object* l_Lean_Parser_sepByElemParser___closed__3; @@ -439,6 +441,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_takeWhileFn___boxed(lean_object*, lean_ob lean_object* lean_array_to_list(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_unicodeSymbolInfo___elambda__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_symbol(lean_object*); +static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_satisfyFn(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_leadingParserAux___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_optionaInfo(lean_object*); @@ -450,9 +453,11 @@ LEAN_EXPORT lean_object* l_Lean_Syntax_forArgsM(lean_object*); static lean_object* l_Lean_Parser_termParser___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_sepByNoAntiquot___elambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ParserState_mergeErrors___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_hygieneInfoFn___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_setLhsPrec___elambda__1___boxed(lean_object*, lean_object*, lean_object*); uint8_t lean_string_utf8_at_end(lean_object*, lean_object*); static lean_object* l_Lean_Parser_orelseFnCore___closed__1; +static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__17; LEAN_EXPORT lean_object* l_Lean_Parser_symbolInfo___elambda__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_checkNoImmediateColon; LEAN_EXPORT lean_object* l_Lean_Parser_many1Unbox___lambda__1(lean_object*); @@ -464,10 +469,10 @@ LEAN_EXPORT lean_object* l_Lean_Parser_manyAux___lambda__1(lean_object*, lean_ob LEAN_EXPORT lean_object* l_Lean_Parser_checkLhsPrecFn___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_sepBy___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ParserState_keepPrevError(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_8777____closed__15; LEAN_EXPORT lean_object* l_Lean_Parser_atomic___elambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_checkNoWsBefore(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_takeUntilFn(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__5; static lean_object* l_Lean_Parser_mkAntiquot___closed__8; LEAN_EXPORT lean_object* l_Lean_Parser_optionalNoAntiquot___elambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_sepByFnAux_parse___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -531,22 +536,26 @@ LEAN_EXPORT lean_object* l_Lean_Parser_node(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_chFn(uint32_t, uint8_t, lean_object*, lean_object*); uint8_t l_Lean_Syntax_isIdent(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_foldArgs(lean_object*); +static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__5; LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_sepByFnAux(lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Parser_TokenMap_insert___spec__1___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_numLitNoAntiquot; +static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__16; LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Parser_indexed___spec__5___rarg___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Parser_mkAntiquot___closed__14; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9339____lambda__1___closed__1; static lean_object* l_Lean_Parser_TokenMap_instForInTokenMapProdNameList___closed__1; static lean_object* l_Lean_Parser_numLitNoAntiquot___closed__1; +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9301_(lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Parser_Basic_0__Lean_Parser_isIdFirstOrBeginEscape(uint32_t); static lean_object* l_Lean_Parser_anyOfFn___closed__1; lean_object* l_Lean_Parser_ParserState_mkError(lean_object*, lean_object*); static lean_object* l_Lean_Parser_decimalNumberFn___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_checkStackTopFn___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_indexed___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_hygieneInfoNoAntiquot___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_withForbidden(lean_object*, lean_object*); lean_object* l_Lean_Parser_Error_merge(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9274____lambda__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_categoryParser___elambda__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_charLitFn___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_andthen(lean_object*, lean_object*); @@ -567,11 +576,12 @@ LEAN_EXPORT lean_object* l_Lean_Parser_checkColGe(lean_object*); static lean_object* l_Lean_Parser_instCoeStringParser___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_trailingNodeAux___elambda__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_whitespace___closed__3; -static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_sepByInfo___elambda__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Parser_whitespace___lambda__1(uint32_t); lean_object* l_Array_append___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_beqLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8753____boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_withAntiquotSuffixSpliceFn___lambda__1___closed__2; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9339____closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_checkLhsPrec___elambda__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_categoryParserFn___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_atomic(lean_object*); @@ -583,6 +593,7 @@ LEAN_EXPORT uint8_t l_Lean_Parser_unicodeSymbolFnAux___lambda__1(lean_object*, l LEAN_EXPORT lean_object* l_Lean_Parser_errorFn(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_sepBy1Info___elambda__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at_Lean_Parser_TokenMap_insert___spec__2(lean_object*); +static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__12; static lean_object* l_Lean_Parser_binNumberFn___closed__2; static lean_object* l_Lean_Parser_many1Unbox___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_octalNumberFn___lambda__1___boxed(lean_object*); @@ -597,18 +608,21 @@ LEAN_EXPORT lean_object* l_Lean_Parser_sepBy1NoAntiquot___elambda__1(uint8_t, le uint8_t l_Char_isDigit(uint32_t); static lean_object* l_Lean_Parser_indexed___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_sepBy1Info___elambda__1(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__13; +static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__4; static lean_object* l_Lean_Parser_tokenAntiquotFn___lambda__2___closed__6; static lean_object* l_Lean_Parser_dbgTraceStateFn___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_withPosition___elambda__1___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_symbolNoAntiquot___elambda__1(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__19; LEAN_EXPORT lean_object* l_Lean_Parser_nodeInfo___elambda__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Subarray_findSomeRevM_x3f_find___at___private_Lean_Parser_Basic_0__Lean_Parser_pickNonNone___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_unicodeSymbolFn(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_PrattParsingTables_trailingTable___default; static lean_object* l_Lean_Parser_checkNoImmediateColon___closed__2; LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Parser_indexed___spec__7___rarg___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Parser_mkTokenAndFixPos___closed__2; +static lean_object* l_Lean_Parser_hygieneInfoNoAntiquot___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_trailingLoop___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_mkAntiquotSplice___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_unicodeSymbolNoAntiquot___boxed(lean_object*, lean_object*); @@ -632,7 +646,6 @@ LEAN_EXPORT lean_object* l_Lean_Parser_suppressInsideQuot___lambda__1(lean_objec LEAN_EXPORT lean_object* l_Lean_Parser_errorAtSavedPos___boxed(lean_object*, lean_object*); lean_object* l_lexOrd___elambda__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_symbolFnAux(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_8712____closed__10; lean_object* l_Lean_Parser_FirstTokens_toOptional(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ParserCategory_kinds___default; LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at_Lean_Parser_TokenMap_insert___spec__5___rarg(lean_object*, lean_object*, lean_object*); @@ -703,7 +716,6 @@ static lean_object* l_Lean_Parser_strLitFnAux___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_symbolInfo___elambda__2(lean_object*, lean_object*); uint8_t l_Lean_Syntax_isNone(lean_object*); lean_object* l_UInt32_decEq___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9236_(lean_object*); static lean_object* l_Lean_Parser_orelseFnCore___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_checkColEq___elambda__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_chFn___closed__1; @@ -716,6 +728,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_withResultOf(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_unicodeSymbolFnAux(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_checkNoWsBeforeFn___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_checkLineEq(lean_object*); +static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__7; static lean_object* l_Lean_Parser_whitespace___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_orelseInfo___elambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_noFirstTokenInfo___elambda__1(lean_object*, lean_object*); @@ -735,11 +748,11 @@ static lean_object* l_Lean_Parser_charLitFnAux___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_identNoAntiquot; LEAN_EXPORT lean_object* l_Lean_Parser_error___elambda__1(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); -static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__18; static lean_object* l_Lean_Parser_pushNone___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_checkNoWsBeforeFn(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_FirstTokens_seq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_checkWsBefore(lean_object*); +static lean_object* l_Lean_Parser_hygieneInfoFn___closed__1; lean_object* lean_string_utf8_next_fast(lean_object*, lean_object*); static lean_object* l_Lean_Parser_epsilonInfo___closed__1; static lean_object* l_Lean_Parser_octalNumberFn___closed__2; @@ -749,11 +762,12 @@ static lean_object* l_Lean_Parser_quotedCharCoreFn___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_notFollowedByFn___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_strAux_parse___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_epsilonInfo___elambda__1___boxed(lean_object*); +static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__14; +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9339_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_hygieneInfoFn___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_pushNone___elambda__1___boxed(lean_object*); -static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_identEqFn(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_identFnAux_parse___closed__2; -static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__17; static lean_object* l_Lean_Parser_ParserCategory_kinds___default___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_lookahead(lean_object*); uint8_t l_Lean_Name_quickCmp(lean_object*, lean_object*); @@ -775,13 +789,11 @@ LEAN_EXPORT lean_object* l_Lean_Parser_TokenMap_insert___rarg(lean_object*, lean LEAN_EXPORT lean_object* l_Lean_Parser_OrElseOnAntiquotBehavior_toCtorIdx(uint8_t); LEAN_EXPORT lean_object* l_Lean_Parser_ParserState_mergeErrors(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Parser_checkTailLinebreak(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_beqLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8688____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_many1NoAntiquot___elambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_mkAntiquot(lean_object*, lean_object*, uint8_t, uint8_t); LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_withAntiquotSuffixSpliceFn(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_sepBy1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); -LEAN_EXPORT uint8_t l___private_Lean_Parser_Basic_0__Lean_Parser_beqLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8688_(uint8_t, uint8_t); static lean_object* l_Lean_Parser_scientificLitNoAntiquot___closed__2; static lean_object* l_Lean_Parser_whitespace___closed__5; static lean_object* l_Lean_Parser_tokenAntiquotFn___lambda__2___closed__5; @@ -805,7 +817,6 @@ LEAN_EXPORT lean_object* l_Lean_Parser_takeWhileFn(lean_object*, lean_object*, l LEAN_EXPORT lean_object* l_Lean_Parser_checkTailLinebreak___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_sepByFnAux_parse___lambda__3(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_checkTailNoWs___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9236____lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_setExpectedFn(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_forArgsM___spec__2___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_mkAntiquot___closed__5; @@ -824,6 +835,7 @@ static lean_object* l_Lean_Parser_antiquotExpr___closed__3; uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_1116____at___private_Lean_Parser_Types_0__Lean_Parser_beqCacheableParserContext____x40_Lean_Parser_Types___hyg_235____spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_dbgTraceState___elambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_beqOrElseOnAntiquotBehavior____x40_Lean_Parser_Basic___hyg_871____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9301____lambda__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_mkAntiquotSplice___closed__2; lean_object* lean_array_get_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_eoiFn(lean_object*, lean_object*); @@ -832,6 +844,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_orelseFnCore___boxed(lean_object*, lean_o LEAN_EXPORT lean_object* l_Lean_Syntax_foldArgs___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_whitespace(lean_object*, lean_object*); lean_object* lean_int_neg(lean_object*); +static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__18; LEAN_EXPORT lean_object* l_Lean_Parser_setLhsPrec(lean_object*); LEAN_EXPORT uint8_t l_Lean_Parser_checkTailNoWs(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_TokenMap_instInhabitedTokenMap(lean_object*); @@ -864,13 +877,11 @@ LEAN_EXPORT lean_object* l_Lean_Parser_tokenWithAntiquot___elambda__1(lean_objec LEAN_EXPORT lean_object* l_Lean_Parser_checkPrec___elambda__1___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_ParserState_mkUnexpectedErrorAt(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_withAntiquot___elambda__1(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9274____closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_takeWhileFn___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_instOrElseParser(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_categoryParserFnRef; LEAN_EXPORT lean_object* l_Lean_Parser_pushNone___elambda__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_skip___elambda__1___boxed(lean_object*); -static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_checkLineEq___elambda__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_mkAntiquot___closed__16; LEAN_EXPORT lean_object* l_Lean_Parser_leadingParserAux___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -879,7 +890,6 @@ LEAN_EXPORT lean_object* l_Lean_Parser_withoutPosition(lean_object*); static lean_object* l_Lean_Parser_nonReservedSymbolInfo___closed__1; uint8_t l_Lean_isIdFirst(uint32_t); LEAN_EXPORT lean_object* l_Lean_Parser_errorAtSavedPos___elambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712_(uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_forArgsM___rarg___lambda__1(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_isAntiquots(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_sepByNoAntiquot___boxed(lean_object*, lean_object*, lean_object*); @@ -893,7 +903,6 @@ LEAN_EXPORT lean_object* l_Lean_Parser_octalNumberFn(lean_object*, lean_object*, LEAN_EXPORT lean_object* l_Lean_Parser_categoryParserFn(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_antiquotNestedExpr___closed__9; LEAN_EXPORT lean_object* l_Lean_Parser_trailingLoop___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__4; lean_object* l_Nat_repr(lean_object*); static lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_withAntiquotSuffixSpliceFn___lambda__1___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_sepByFnAux_parse(lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*); @@ -905,12 +914,13 @@ static lean_object* l_Lean_Parser_tokenAntiquotFn___lambda__2___closed__9; LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at_Lean_Parser_TokenMap_insert___spec__3___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_mkAntiquot___closed__10; lean_object* l_Lean_Parser_ParserState_mkErrorsAt(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_8712____closed__11; static lean_object* l_Lean_Parser_quotedCharFn___closed__1; +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9339____lambda__1(lean_object*); static lean_object* l_Lean_Parser_antiquotNestedExpr___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_optionalFn(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_pushNone___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_decimalNumberFn_parseOptExp___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_hygieneInfoNoAntiquot___closed__2; static lean_object* l_Lean_Parser_epsilonInfo___closed__3; static lean_object* l_Lean_Parser_nameLitNoAntiquot___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_lookahead___elambda__1(lean_object*, lean_object*, lean_object*); @@ -9523,6 +9533,120 @@ lean_ctor_set(x_4, 1, x_2); return x_4; } } +static lean_object* _init_l_Lean_Parser_hygieneInfoFn___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_hygieneInfoFn___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("hygieneInfo", 11); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_hygieneInfoFn___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_hygieneInfoFn___closed__2; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Parser_hygieneInfoFn(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; +x_3 = lean_ctor_get(x_1, 0); +x_4 = lean_ctor_get(x_3, 0); +x_5 = lean_ctor_get(x_2, 2); +lean_inc(x_5); +lean_inc_n(x_5, 2); +lean_inc(x_4); +x_6 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_6, 0, x_4); +lean_ctor_set(x_6, 1, x_5); +lean_ctor_set(x_6, 2, x_5); +lean_inc(x_5); +lean_inc_n(x_6, 2); +x_7 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_5); +lean_ctor_set(x_7, 2, x_6); +lean_ctor_set(x_7, 3, x_5); +x_8 = lean_box(0); +x_9 = lean_box(0); +x_10 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_10, 0, x_7); +lean_ctor_set(x_10, 1, x_6); +lean_ctor_set(x_10, 2, x_9); +lean_ctor_set(x_10, 3, x_8); +x_11 = l_Lean_Parser_hygieneInfoFn___closed__1; +x_12 = lean_array_push(x_11, x_10); +x_13 = lean_box(2); +x_14 = l_Lean_Parser_hygieneInfoFn___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 = l_Lean_Parser_ParserState_pushSyntax(x_2, x_15); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Lean_Parser_hygieneInfoFn___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_Parser_hygieneInfoFn(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_hygieneInfoNoAntiquot___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_hygieneInfoFn___closed__3; +x_2 = l_Lean_Parser_epsilonInfo; +x_3 = l_Lean_Parser_nodeInfo(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_hygieneInfoNoAntiquot___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Parser_hygieneInfoFn___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_hygieneInfoNoAntiquot___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_hygieneInfoNoAntiquot___closed__1; +x_2 = l_Lean_Parser_hygieneInfoNoAntiquot___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); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_hygieneInfoNoAntiquot() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Parser_hygieneInfoNoAntiquot___closed__3; +return x_1; +} +} LEAN_EXPORT lean_object* l_Lean_Parser_ParserState_keepTop(lean_object* x_1, lean_object* x_2) { _start: { @@ -21113,7 +21237,7 @@ x_1 = 0; return x_1; } } -LEAN_EXPORT uint8_t l___private_Lean_Parser_Basic_0__Lean_Parser_beqLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8688_(uint8_t x_1, uint8_t x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Parser_Basic_0__Lean_Parser_beqLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8753_(uint8_t x_1, uint8_t x_2) { _start: { lean_object* x_3; lean_object* x_4; uint8_t x_5; @@ -21125,7 +21249,7 @@ lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_beqLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8688____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_beqLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8753____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; @@ -21133,7 +21257,7 @@ x_3 = lean_unbox(x_1); lean_dec(x_1); x_4 = lean_unbox(x_2); lean_dec(x_2); -x_5 = l___private_Lean_Parser_Basic_0__Lean_Parser_beqLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8688_(x_3, x_4); +x_5 = l___private_Lean_Parser_Basic_0__Lean_Parser_beqLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8753_(x_3, x_4); x_6 = lean_box(x_5); return x_6; } @@ -21142,7 +21266,7 @@ static lean_object* _init_l_Lean_Parser_instBEqLeadingIdentBehavior___closed__1( _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Parser_Basic_0__Lean_Parser_beqLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8688____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Parser_Basic_0__Lean_Parser_beqLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8753____boxed), 2, 0); return x_1; } } @@ -21154,7 +21278,7 @@ x_1 = l_Lean_Parser_instBEqLeadingIdentBehavior___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__1() { +static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__1() { _start: { lean_object* x_1; @@ -21162,17 +21286,17 @@ x_1 = lean_mk_string_from_bytes("Lean.Parser.LeadingIdentBehavior.default", 40); return x_1; } } -static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__2() { +static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__1; +x_1 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__1; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__3() { +static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__3() { _start: { lean_object* x_1; lean_object* x_2; @@ -21181,23 +21305,23 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__4() { +static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__3; -x_2 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__2; +x_1 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__3; +x_2 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____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_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__5() { +static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__5() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__4; +x_1 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__4; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -21205,23 +21329,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__6() { +static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_incQuotDepth___closed__1; -x_2 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__2; +x_2 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____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_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__7() { +static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__7() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__6; +x_1 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__6; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -21229,7 +21353,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__8() { +static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__8() { _start: { lean_object* x_1; @@ -21237,33 +21361,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Parser.LeadingIdentBehavior.symbol", 39); return x_1; } } -static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__9() { +static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__9() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__8; +x_1 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__8; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__10() { +static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__3; -x_2 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__9; +x_1 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__3; +x_2 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__9; 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_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__11() { +static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__11() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__10; +x_1 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__10; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -21271,23 +21395,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__12() { +static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_incQuotDepth___closed__1; -x_2 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__9; +x_2 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__9; 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_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__13() { +static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__13() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__12; +x_1 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__12; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -21295,7 +21419,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__14() { +static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__14() { _start: { lean_object* x_1; @@ -21303,33 +21427,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Parser.LeadingIdentBehavior.both", 37); return x_1; } } -static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__15() { +static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__15() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__14; +x_1 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__14; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__16() { +static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__3; -x_2 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__15; +x_1 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__3; +x_2 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__15; 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_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__17() { +static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__17() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__16; +x_1 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__16; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -21337,23 +21461,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__18() { +static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_incQuotDepth___closed__1; -x_2 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__15; +x_2 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__15; 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_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__19() { +static lean_object* _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__19() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__18; +x_1 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__18; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -21361,7 +21485,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712_(uint8_t x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777_(uint8_t x_1, lean_object* x_2) { _start: { switch (x_1) { @@ -21373,14 +21497,14 @@ x_4 = lean_nat_dec_le(x_3, x_2); if (x_4 == 0) { lean_object* x_5; lean_object* x_6; -x_5 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__5; +x_5 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__5; x_6 = l_Repr_addAppParen(x_5, x_2); return x_6; } else { lean_object* x_7; lean_object* x_8; -x_7 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__7; +x_7 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__7; x_8 = l_Repr_addAppParen(x_7, x_2); return x_8; } @@ -21393,14 +21517,14 @@ x_10 = lean_nat_dec_le(x_9, x_2); if (x_10 == 0) { lean_object* x_11; lean_object* x_12; -x_11 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__11; +x_11 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__11; x_12 = l_Repr_addAppParen(x_11, x_2); return x_12; } else { lean_object* x_13; lean_object* x_14; -x_13 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__13; +x_13 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__13; x_14 = l_Repr_addAppParen(x_13, x_2); return x_14; } @@ -21413,14 +21537,14 @@ x_16 = lean_nat_dec_le(x_15, x_2); if (x_16 == 0) { lean_object* x_17; lean_object* x_18; -x_17 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__17; +x_17 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__17; x_18 = l_Repr_addAppParen(x_17, x_2); return x_18; } else { lean_object* x_19; lean_object* x_20; -x_19 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__19; +x_19 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__19; x_20 = l_Repr_addAppParen(x_19, x_2); return x_20; } @@ -21428,13 +21552,13 @@ return x_20; } } } -LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; x_3 = lean_unbox(x_1); lean_dec(x_1); -x_4 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712_(x_3, x_2); +x_4 = l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777_(x_3, x_2); lean_dec(x_2); return x_4; } @@ -21443,7 +21567,7 @@ static lean_object* _init_l_Lean_Parser_instReprLeadingIdentBehavior___closed__1 _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____boxed), 2, 0); return x_1; } } @@ -22541,7 +22665,7 @@ lean_dec(x_1); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9236____lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9301____lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; @@ -22549,19 +22673,19 @@ x_4 = l_Lean_Parser_whitespace(x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9236____closed__1() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9301____closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9236____lambda__1___boxed), 3, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9301____lambda__1___boxed), 3, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9236_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9301_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; uint8_t x_4; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9236____closed__1; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9301____closed__1; x_3 = lean_st_mk_ref(x_2, x_1); x_4 = !lean_is_exclusive(x_3); if (x_4 == 0) @@ -22583,16 +22707,16 @@ return x_7; } } } -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9236____lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9301____lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9236____lambda__1(x_1, x_2, x_3); +x_4 = l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9301____lambda__1(x_1, x_2, x_3); lean_dec(x_1); return x_4; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9274____lambda__1___closed__1() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9339____lambda__1___closed__1() { _start: { lean_object* x_1; @@ -22600,11 +22724,11 @@ x_1 = l_Lean_Parser_categoryParserFnRef; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9274____lambda__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9339____lambda__1(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; uint8_t x_4; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9274____lambda__1___closed__1; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9339____lambda__1___closed__1; x_3 = lean_st_ref_get(x_2, x_1); x_4 = !lean_is_exclusive(x_3); if (x_4 == 0) @@ -22626,19 +22750,19 @@ return x_7; } } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9274____closed__1() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9339____closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9274____lambda__1), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9339____lambda__1), 1, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9274_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9339_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9274____closed__1; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9339____closed__1; x_3 = l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg(x_2, x_1); return x_3; } @@ -25593,6 +25717,20 @@ l_Lean_Parser_rawIdentNoAntiquot = _init_l_Lean_Parser_rawIdentNoAntiquot(); lean_mark_persistent(l_Lean_Parser_rawIdentNoAntiquot); l_Lean_Parser_identEqFn___closed__1 = _init_l_Lean_Parser_identEqFn___closed__1(); lean_mark_persistent(l_Lean_Parser_identEqFn___closed__1); +l_Lean_Parser_hygieneInfoFn___closed__1 = _init_l_Lean_Parser_hygieneInfoFn___closed__1(); +lean_mark_persistent(l_Lean_Parser_hygieneInfoFn___closed__1); +l_Lean_Parser_hygieneInfoFn___closed__2 = _init_l_Lean_Parser_hygieneInfoFn___closed__2(); +lean_mark_persistent(l_Lean_Parser_hygieneInfoFn___closed__2); +l_Lean_Parser_hygieneInfoFn___closed__3 = _init_l_Lean_Parser_hygieneInfoFn___closed__3(); +lean_mark_persistent(l_Lean_Parser_hygieneInfoFn___closed__3); +l_Lean_Parser_hygieneInfoNoAntiquot___closed__1 = _init_l_Lean_Parser_hygieneInfoNoAntiquot___closed__1(); +lean_mark_persistent(l_Lean_Parser_hygieneInfoNoAntiquot___closed__1); +l_Lean_Parser_hygieneInfoNoAntiquot___closed__2 = _init_l_Lean_Parser_hygieneInfoNoAntiquot___closed__2(); +lean_mark_persistent(l_Lean_Parser_hygieneInfoNoAntiquot___closed__2); +l_Lean_Parser_hygieneInfoNoAntiquot___closed__3 = _init_l_Lean_Parser_hygieneInfoNoAntiquot___closed__3(); +lean_mark_persistent(l_Lean_Parser_hygieneInfoNoAntiquot___closed__3); +l_Lean_Parser_hygieneInfoNoAntiquot = _init_l_Lean_Parser_hygieneInfoNoAntiquot(); +lean_mark_persistent(l_Lean_Parser_hygieneInfoNoAntiquot); l_Lean_Parser_invalidLongestMatchParser___closed__1 = _init_l_Lean_Parser_invalidLongestMatchParser___closed__1(); lean_mark_persistent(l_Lean_Parser_invalidLongestMatchParser___closed__1); l_Lean_Parser_longestMatchStep___closed__1 = _init_l_Lean_Parser_longestMatchStep___closed__1(); @@ -25636,44 +25774,44 @@ l_Lean_Parser_instBEqLeadingIdentBehavior___closed__1 = _init_l_Lean_Parser_inst lean_mark_persistent(l_Lean_Parser_instBEqLeadingIdentBehavior___closed__1); l_Lean_Parser_instBEqLeadingIdentBehavior = _init_l_Lean_Parser_instBEqLeadingIdentBehavior(); lean_mark_persistent(l_Lean_Parser_instBEqLeadingIdentBehavior); -l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__1 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__1(); -lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__1); -l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__2 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__2(); -lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__2); -l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__3 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__3(); -lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__3); -l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__4 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__4(); -lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__4); -l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__5 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__5(); -lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__5); -l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__6 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__6(); -lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__6); -l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__7 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__7(); -lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__7); -l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__8 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__8(); -lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__8); -l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__9 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__9(); -lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__9); -l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__10 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__10(); -lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__10); -l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__11 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__11(); -lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__11); -l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__12 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__12(); -lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__12); -l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__13 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__13(); -lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__13); -l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__14 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__14(); -lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__14); -l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__15 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__15(); -lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__15); -l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__16 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__16(); -lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__16); -l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__17 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__17(); -lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__17); -l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__18 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__18(); -lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__18); -l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__19 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__19(); -lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8712____closed__19); +l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__1 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__1(); +lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__1); +l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__2 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__2(); +lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__2); +l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__3 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__3(); +lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__3); +l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__4 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__4(); +lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__4); +l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__5 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__5(); +lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__5); +l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__6 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__6(); +lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__6); +l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__7 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__7(); +lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__7); +l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__8 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__8(); +lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__8); +l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__9 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__9(); +lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__9); +l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__10 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__10(); +lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__10); +l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__11 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__11(); +lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__11); +l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__12 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__12(); +lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__12); +l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__13 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__13(); +lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__13); +l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__14 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__14(); +lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__14); +l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__15 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__15(); +lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__15); +l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__16 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__16(); +lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__16); +l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__17 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__17(); +lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__17); +l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__18 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__18(); +lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__18); +l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__19 = _init_l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__19(); +lean_mark_persistent(l___private_Lean_Parser_Basic_0__Lean_Parser_reprLeadingIdentBehavior____x40_Lean_Parser_Basic___hyg_8777____closed__19); l_Lean_Parser_instReprLeadingIdentBehavior___closed__1 = _init_l_Lean_Parser_instReprLeadingIdentBehavior___closed__1(); lean_mark_persistent(l_Lean_Parser_instReprLeadingIdentBehavior___closed__1); l_Lean_Parser_instReprLeadingIdentBehavior = _init_l_Lean_Parser_instReprLeadingIdentBehavior(); @@ -25694,18 +25832,18 @@ l_Lean_Parser_instInhabitedParserCategory = _init_l_Lean_Parser_instInhabitedPar lean_mark_persistent(l_Lean_Parser_instInhabitedParserCategory); l_Lean_Parser_indexed___rarg___closed__1 = _init_l_Lean_Parser_indexed___rarg___closed__1(); lean_mark_persistent(l_Lean_Parser_indexed___rarg___closed__1); -l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9236____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9236____closed__1(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9236____closed__1); -if (builtin) {res = l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9236_(lean_io_mk_world()); +l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9301____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9301____closed__1(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9301____closed__1); +if (builtin) {res = l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9301_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Parser_categoryParserFnRef = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Parser_categoryParserFnRef); lean_dec_ref(res); -}l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9274____lambda__1___closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9274____lambda__1___closed__1(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9274____lambda__1___closed__1); -l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9274____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9274____closed__1(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9274____closed__1); -if (builtin) {res = l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9274_(lean_io_mk_world()); +}l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9339____lambda__1___closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9339____lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9339____lambda__1___closed__1); +l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9339____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9339____closed__1(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9339____closed__1); +if (builtin) {res = l_Lean_Parser_initFn____x40_Lean_Parser_Basic___hyg_9339_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Parser_categoryParserFnExtension = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Parser_categoryParserFnExtension); diff --git a/stage0/stdlib/Lean/Parser/Command.c b/stage0/stdlib/Lean/Parser/Command.c index 9a06eea3f8d2..bd3cee5a1b77 100644 --- a/stage0/stdlib/Lean/Parser/Command.c +++ b/stage0/stdlib/Lean/Parser/Command.c @@ -13,7 +13,6 @@ #ifdef __cplusplus extern "C" { #endif -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_Command_openRenamingItem; static lean_object* l_Lean_Parser_Command_declModifiers___closed__2; static lean_object* l_Lean_Parser_Command_attribute_formatter___closed__3; @@ -39,19 +38,21 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_declModifiersT_formatter(lean_obj static lean_object* l_Lean_Parser_Command_openHiding_formatter___closed__4; static lean_object* l_Lean_Parser_Command_declId_formatter___closed__11; static lean_object* l_Lean_Parser_Command_nonrec_parenthesizer___closed__2; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__14; static lean_object* l_Lean_Parser_Command_deriving_formatter___closed__8; static lean_object* l___regBuiltin_Lean_Parser_Tactic_open_declRange___closed__6; static lean_object* l_Lean_Parser_Command_noncomputableSection_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Command_protected_formatter___closed__3; static lean_object* l_Lean_Parser_Command_structure_parenthesizer___closed__13; static lean_object* l_Lean_Parser_Command_ctor_parenthesizer___closed__6; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__15; static lean_object* l_Lean_Parser_Command_structFields_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Tactic_open_formatter___closed__5; static lean_object* l_Lean_Parser_Command_structure___closed__11; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__13; static lean_object* l_Lean_Parser_Command_terminationByElement_formatter___closed__1; static lean_object* l_Lean_Parser_Command_optDefDeriving_parenthesizer___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_Command_declId_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__22; static lean_object* l_Lean_Parser_Command_openSimple___closed__6; static lean_object* l_Lean_Parser_Command_declSig_formatter___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Command_terminationHint1(lean_object*); @@ -91,20 +92,21 @@ static lean_object* l_Lean_Parser_Command_classInductive_formatter___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_Command_instance_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Command_initializeKeyword; static lean_object* l_Lean_Parser_Command_theorem___closed__8; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__8; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__41; LEAN_EXPORT lean_object* l_Lean_Parser_Command_openSimple_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_unsafe___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__8; static lean_object* l_Lean_Parser_Command_optDeriving_formatter___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_set__option_formatter___closed__2; static lean_object* l_Lean_Parser_Command_attribute___closed__3; lean_object* l_Lean_Parser_rawIdent_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__29; static lean_object* l_Lean_Parser_Command_partial_formatter___closed__3; static lean_object* l_Lean_Parser_Command_protected___closed__1; static lean_object* l_Lean_Parser_Command_terminationHintMany___closed__5; static lean_object* l_Lean_Parser_Tactic_set__option_formatter___closed__6; static lean_object* l_Lean_Parser_Command_declSig___closed__8; static lean_object* l___regBuiltin_Lean_Parser_Term_quot_formatter___closed__2; +static lean_object* l_Lean_Parser_Command_openOnly_formatter___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Command_partial_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_openHiding_parenthesizer___closed__8; lean_object* l_Lean_Parser_optional_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -120,7 +122,6 @@ static lean_object* l_Lean_Parser_Command_instance_formatter___closed__7; static lean_object* l_Lean_Parser_Term_set__option_formatter___closed__6; static lean_object* l_Lean_Parser_Command_openRenamingItem___closed__6; static lean_object* l_Lean_Parser_Command_decreasingBy_parenthesizer___closed__2; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__37; static lean_object* l___regBuiltin_Lean_Parser_Command_structureTk_parenthesizer___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Tactic_open_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_Command_synth___closed__1; @@ -176,7 +177,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_open_declRange___closed__6; static lean_object* l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Command_addDocString_declRange___closed__2; static lean_object* l_Lean_Parser_Command_structExplicitBinder___closed__11; -static lean_object* l_Lean_Parser_Command_export___closed__14; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__29; static lean_object* l_Lean_Parser_Command_exit___closed__4; static lean_object* l_Lean_Parser_Command_opaque_formatter___closed__6; static lean_object* l_Lean_Parser_Command_set__option_parenthesizer___closed__2; @@ -217,6 +218,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Tactic_set__option_docString___cl static lean_object* l_Lean_Parser_Command_openScoped_formatter___closed__1; static lean_object* l_Lean_Parser_Tactic_open_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_structInstBinder___closed__2; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__50; static lean_object* l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__12; static lean_object* l_Lean_Parser_Command_def___closed__6; static lean_object* l_Lean_Parser_Command_classInductive_parenthesizer___closed__10; @@ -228,19 +230,19 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_abbrev_formatter___closed LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_private_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_Command_declModifiers___closed__7; static lean_object* l_Lean_Parser_Command_declModifiers___closed__14; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__32; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_private_formatter(lean_object*); +static lean_object* l_Lean_Parser_Tactic_open___closed__7; static lean_object* l_Lean_Parser_Command_terminationByElement_formatter___closed__6; static lean_object* l_Lean_Parser_Command_noncomputableSection___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_reduce_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_instance_parenthesizer___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Command_structImplicitBinder_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Command_optNamedPrio_formatter___closed__1; static lean_object* l_Lean_Parser_Command_declModifiers___closed__17; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Tactic_set__option(lean_object*); static lean_object* l_Lean_Parser_Command_declId_parenthesizer___closed__8; static lean_object* l___regBuiltin_Lean_Parser_Command_quot_declRange___closed__6; static lean_object* l_Lean_Parser_Command_namedPrio___closed__6; -static lean_object* l_Lean_Parser_Command_mutual___closed__16; static lean_object* l_Lean_Parser_Command_structImplicitBinder___closed__11; static lean_object* l_Lean_Parser_Command_universe_formatter___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_set__option_formatter(lean_object*); @@ -259,7 +261,9 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_structFields_parenthesize static lean_object* l___regBuiltin_Lean_Parser_Command_printAxioms_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_noncomputable_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_axiom_parenthesizer___closed__5; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__47; static lean_object* l_Lean_Parser_Command_openOnly_parenthesizer___closed__2; +static lean_object* l_Lean_Parser_Command_structExplicitBinder___closed__18; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_precheckedQuot_declRange(lean_object*); static lean_object* l_Lean_Parser_Command_set__option___closed__4; static lean_object* l_Lean_Parser_Command_print___closed__4; @@ -277,6 +281,7 @@ static lean_object* l_Lean_Parser_Command_eval_parenthesizer___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Command_axiom_parenthesizer___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_import(lean_object*); 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_terminationByElement_formatter___closed__14; static lean_object* l_Lean_Parser_Command_inductive_formatter___closed__14; static lean_object* l_Lean_Parser_Command_structure_parenthesizer___closed__9; static lean_object* l_Lean_Parser_Command_declModifiers_formatter___closed__15; @@ -316,7 +321,6 @@ static lean_object* l_Lean_Parser_Command_abbrev_formatter___closed__2; static lean_object* l_Lean_Parser_Command_declValSimple_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_declModifiers_formatter___closed__21; static lean_object* l_Lean_Parser_Command_instance_parenthesizer___closed__7; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__18; static lean_object* l_Lean_Parser_Command_check__failure___closed__2; static lean_object* l_Lean_Parser_Command_set__option___closed__11; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_optDeriving_formatter(lean_object*); @@ -330,13 +334,13 @@ static lean_object* l_Lean_Parser_Command_structureTk___closed__2; static lean_object* l_Lean_Parser_Command_openSimple_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_openScoped___closed__1; static lean_object* l_Lean_Parser_Command_computedField_formatter___closed__2; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__50; static lean_object* l_Lean_Parser_Command_structExplicitBinder___closed__15; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_declaration_declRange(lean_object*); static lean_object* l_Lean_Parser_Command_quot___closed__10; static lean_object* l_Lean_Parser_Command_structImplicitBinder_parenthesizer___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Command_partial_formatter___closed__2; static lean_object* l_Lean_Parser_Command_declSig___closed__9; +static lean_object* l_Lean_Parser_Command_optDefDeriving_formatter___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Command_open_formatter___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Tactic_set__option_formatter(lean_object*); static lean_object* l_Lean_Parser_Command_def_parenthesizer___closed__4; @@ -395,6 +399,7 @@ static lean_object* l_Lean_Parser_Command_initializeKeyword___closed__3; static lean_object* l_Lean_Parser_Command_declaration_formatter___closed__3; static lean_object* l_Lean_Parser_Command_terminationBy___closed__5; static lean_object* l_Lean_Parser_Command_partial___closed__2; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Command_declId_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_synth_formatter___closed__1; static lean_object* l_Lean_Parser_Command_declId___closed__11; @@ -425,6 +430,7 @@ static lean_object* l_Lean_Parser_Command_check_formatter___closed__2; static lean_object* l_Lean_Parser_Command_openOnly___closed__5; static lean_object* l_Lean_Parser_Command_attribute___closed__15; static lean_object* l_Lean_Parser_Command_decreasingBy_formatter___closed__3; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__6; static lean_object* l_Lean_Parser_Command_section___closed__3; lean_object* l_Lean_Parser_leadingNode(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_private_parenthesizer___closed__2; @@ -442,7 +448,6 @@ extern lean_object* l_Lean_Parser_Term_binderDefault; static lean_object* l_Lean_Parser_Command_classInductive_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_import_parenthesizer___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Command_optNamedPrio_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Command_attribute___closed__16; lean_object* l_Lean_Parser_categoryParser(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_terminationHintMany___closed__2; static lean_object* l_Lean_Parser_Command_terminationByElement_formatter___closed__7; @@ -465,6 +470,7 @@ static lean_object* l_Lean_Parser_Command_init__quot___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_declValEqns_parenthesizer(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Command_terminationHint1_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_group_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Command_openHiding___closed__16; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_quot(lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_example_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_Command_structureTk_formatter___closed__2; @@ -478,7 +484,6 @@ static lean_object* l_Lean_Parser_Command_openDecl___closed__1; static lean_object* l_Lean_Parser_Command_openOnly_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Command_extends___closed__3; static lean_object* l_Lean_Parser_Command_structure___closed__2; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__6; static lean_object* l_Lean_Parser_Command_eraseAttr___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Command_openHiding_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_decreasingBy_formatter___closed__5; @@ -488,6 +493,7 @@ static lean_object* l_Lean_Parser_Command_whereStructInst_parenthesizer___closed static lean_object* l_Lean_Parser_Command_check___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_protected_formatter___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_3161____closed__28; LEAN_EXPORT lean_object* l_Lean_Parser_Term_open; extern lean_object* l_Lean_Parser_Term_attributes; static lean_object* l_Lean_Parser_Command_declValSimple___closed__1; @@ -515,13 +521,13 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_deriving_parenthesizer___ 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_open___closed__1; static lean_object* l_Lean_Parser_Command_structInstBinder_parenthesizer___closed__4; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__39; static lean_object* l_Lean_Parser_Command_moduleDoc_formatter___closed__4; static lean_object* l_Lean_Parser_Command_print_formatter___closed__4; static lean_object* l_Lean_Parser_Command_optionValue___closed__6; static lean_object* l_Lean_Parser_Command_structure_parenthesizer___closed__18; 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___regBuiltin_Lean_Parser_Command_init__quot_parenthesizer___closed__1; +static lean_object* l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__15; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_section_declRange(lean_object*); static lean_object* l_Lean_Parser_Command_example_formatter___closed__4; static lean_object* l_Lean_Parser_Command_declValSimple_formatter___closed__9; @@ -541,6 +547,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_quot_parenthesizer___closed_ static lean_object* l_Lean_Parser_Command_structExplicitBinder_formatter___closed__2; static lean_object* l_Lean_Parser_Command_namedPrio_formatter___closed__7; static lean_object* l_Lean_Parser_Command_opaque_formatter___closed__7; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__31; static lean_object* l_Lean_Parser_Command_reduce___closed__8; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_namespace_formatter(lean_object*); static lean_object* l_Lean_Parser_Command_terminationHintMany___closed__1; @@ -561,12 +568,14 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_exit(lean_object*); static lean_object* l_Lean_Parser_Command_structSimpleBinder_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Command_moduleDoc___closed__4; static lean_object* l_Lean_Parser_Command_declModifiers_formatter___closed__10; +static lean_object* l_Lean_Parser_Command_noncomputableSection_formatter___closed__7; static lean_object* l_Lean_Parser_Command_open___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_noncomputableSection_formatter___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_structInstBinder_formatter___closed__1; static lean_object* l_Lean_Parser_Command_quot_parenthesizer___closed__6; static lean_object* l_Lean_Parser_Command_abbrev_formatter___closed__5; static lean_object* l_Lean_Parser_Command_print___closed__7; +static lean_object* l_Lean_Parser_Command_structExplicitBinder_formatter___closed__15; static lean_object* l_Lean_Parser_Command_structInstBinder___closed__9; static lean_object* l_Lean_Parser_Command_openRenaming___closed__4; static lean_object* l_Lean_Parser_Command_moduleDoc_formatter___closed__7; @@ -577,6 +586,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_eval_declRange(lean_ static lean_object* l_Lean_Parser_Command_end___closed__5; static lean_object* l_Lean_Parser_Command_initialize_parenthesizer___closed__10; static lean_object* l___regBuiltin_Lean_Parser_Command_openRenamingItem_formatter___closed__2; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__42; lean_object* l_Lean_Parser_optional(lean_object*); static lean_object* l_Lean_Parser_Command_variable___closed__8; static lean_object* l_Lean_Parser_Command_optionValue_formatter___closed__5; @@ -585,11 +595,13 @@ static lean_object* l_Lean_Parser_Command_openRenaming_formatter___closed__7; static lean_object* l_Lean_Parser_Command_end___closed__3; static lean_object* l_Lean_Parser_Command_terminationByCore_formatter___closed__2; static lean_object* l_Lean_Parser_Command_computedField___closed__1; +static lean_object* l_Lean_Parser_Command_decreasingBy_parenthesizer___closed__7; static lean_object* l_Lean_Parser_Command_nonrec___closed__6; static lean_object* l_Lean_Parser_Command_openRenamingItem_parenthesizer___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_declValSimple_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_deriving___closed__14; static lean_object* l_Lean_Parser_Command_theorem___closed__7; +static lean_object* l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__14; static lean_object* l_Lean_Parser_Command_structFields_parenthesizer___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_quot(lean_object*); static lean_object* l_Lean_Parser_Command_extends_formatter___closed__5; @@ -597,11 +609,11 @@ lean_object* l_Lean_Parser_nonReservedSymbol(lean_object*, uint8_t); static lean_object* l_Lean_Parser_Command_whereStructField_parenthesizer___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_whereStructInst_parenthesizer___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_quot_declRange___closed__7; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__3; static lean_object* l_Lean_Parser_Command_whereStructInst___closed__11; LEAN_EXPORT lean_object* l_Lean_Parser_Command_print; static lean_object* l_Lean_Parser_Term_precheckedQuot___closed__4; static lean_object* l_Lean_Parser_Command_abbrev_parenthesizer___closed__7; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__25; static lean_object* l___regBuiltin_Lean_Parser_Command_instance_formatter___closed__2; static lean_object* l_Lean_Parser_Command_optDeclSig_parenthesizer___closed__9; static lean_object* l_Lean_Parser_Command_set__option_formatter___closed__2; @@ -724,6 +736,7 @@ static lean_object* l_Lean_Parser_Command_computedFields_formatter___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_structure_formatter(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_private_formatter___closed__1; static lean_object* l_Lean_Parser_Command_declModifiers_parenthesizer___closed__21; +static lean_object* l_Lean_Parser_Command_initialize___closed__14; static lean_object* l_Lean_Parser_Command_instance___closed__6; static lean_object* l_Lean_Parser_Command_example_parenthesizer___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_noncomputableSection_parenthesizer(lean_object*); @@ -731,13 +744,13 @@ lean_object* l_Lean_Parser_orelse(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_terminationSuffix_formatter___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_declSig_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_declModifiers_formatter___closed__19; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__12; static lean_object* l_Lean_Parser_Command_section_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_openRenaming___closed__2; static lean_object* l_Lean_Parser_Command_example___closed__6; static lean_object* l_Lean_Parser_Command_classInductive_formatter___closed__4; static lean_object* l_Lean_Parser_Command_openScoped___closed__5; lean_object* l_Lean_addBuiltinDeclarationRanges(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Command_terminationByCore_formatter___closed__6; static lean_object* l_Lean_Parser_Command_declValSimple___closed__9; static lean_object* l_Lean_Parser_Command_openOnly_parenthesizer___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_openScoped_formatter(lean_object*); @@ -758,6 +771,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_set__option_declRange___c static lean_object* l___regBuiltin_Lean_Parser_Command_quot_docString___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_open_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_classTk_parenthesizer(lean_object*); +static lean_object* l_Lean_Parser_Command_openHiding___closed__17; LEAN_EXPORT lean_object* l_Lean_Parser_Command_structSimpleBinder_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_namedPrio___closed__2; static lean_object* l_Lean_Parser_Command_computedField___closed__8; @@ -773,7 +787,6 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_genInjectiveTheorems static lean_object* l___regBuiltin_Lean_Parser_Command_initialize_declRange___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Command_deriving_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_ctor_parenthesizer___closed__8; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__54; static lean_object* l_Lean_Parser_Term_set__option_formatter___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Command_namespace_declRange___closed__3; static lean_object* l_Lean_Parser_Command_declModifiers___closed__18; @@ -782,7 +795,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_opaque_parenthesizer___cl static lean_object* l___regBuiltin_Lean_Parser_Command_structImplicitBinder_formatter___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_section(lean_object*); static lean_object* l_Lean_Parser_Command_structExplicitBinder___closed__7; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__28; static lean_object* l_Lean_Parser_Command_structFields_parenthesizer___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Command_instance_formatter___closed__1; static lean_object* l_Lean_Parser_Command_example_formatter___closed__1; @@ -829,9 +841,9 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_example_parenthesizer___c static lean_object* l_Lean_Parser_Command_private_formatter___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_end_formatter___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_export_declRange___closed__7; +static lean_object* l_Lean_Parser_Command_openRenaming___closed__13; static lean_object* l_Lean_Parser_Command_structure___closed__16; static lean_object* l___regBuiltin_Lean_Parser_Command_exit_declRange___closed__3; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__48; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_addDocString_declRange(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Tactic_open_docString___closed__1; static lean_object* l_Lean_Parser_Command_declaration_formatter___closed__2; @@ -874,7 +886,6 @@ lean_object* l_Lean_Parser_Term_matchAltsWhereDecls_formatter(lean_object*, lean static lean_object* l_Lean_Parser_Command_reduce_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_init__quot___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Command_section_declRange___closed__1; -static lean_object* l_Lean_Parser_Command_open___closed__10; static lean_object* l___regBuiltin_Lean_Parser_Command_synth_formatter___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Command_inductive; static lean_object* l_Lean_Parser_Command_whereStructInst_parenthesizer___closed__1; @@ -890,11 +901,11 @@ static lean_object* l_Lean_Parser_Command_check__failure___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_Command_terminationBy_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_initializeKeyword_formatter___closed__2; static lean_object* l_Lean_Parser_Command_eoi___closed__5; +static lean_object* l_Lean_Parser_Command_decreasingBy_formatter___closed__7; static lean_object* l_Lean_Parser_Command_structExplicitBinder___closed__10; static lean_object* l_Lean_Parser_Command_initialize_formatter___closed__9; static lean_object* l___regBuiltin_Lean_Parser_Term_set__option_parenthesizer___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Command_whereStructField; -static lean_object* l_Lean_Parser_Command_attribute___closed__17; static lean_object* l_Lean_Parser_Command_terminationHintMany_formatter___closed__4; static lean_object* l_Lean_Parser_Command_optDeclSig___closed__2; static lean_object* l_Lean_Parser_Command_abbrev_formatter___closed__6; @@ -913,11 +924,13 @@ static lean_object* l_Lean_Parser_Command_open_formatter___closed__5; static lean_object* l_Lean_Parser_Term_precheckedQuot___closed__7; static lean_object* l_Lean_Parser_Command_quot___closed__11; static lean_object* l___regBuiltin_Lean_Parser_Command_initializeKeyword_formatter___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__46; LEAN_EXPORT lean_object* l_Lean_Parser_Command_terminationHint_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_printAxioms_parenthesizer___closed__3; lean_object* l_Lean_Parser_priorityParser_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_deriving___closed__8; static lean_object* l___regBuiltin_Lean_Parser_Command_declaration_declRange___closed__4; +static lean_object* l_Lean_Parser_Command_structure_parenthesizer___closed__19; static lean_object* l_Lean_Parser_Command_ctor___closed__13; static lean_object* l___regBuiltin_Lean_Parser_Command_decreasingBy_parenthesizer___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_initializeKeyword_parenthesizer___closed__1; @@ -933,7 +946,6 @@ static lean_object* l_Lean_Parser_Command_structFields___closed__5; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_check__failure_formatter(lean_object*); static lean_object* l_Lean_Parser_Command_print___closed__9; static lean_object* l_Lean_Parser_Command_computedField___closed__11; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__15; static lean_object* l___regBuiltin_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_mutual_parenthesizer___closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_Term_precheckedQuot_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -950,7 +962,6 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_structure; static lean_object* l_Lean_Parser_Command_whereStructInst_parenthesizer___closed__9; static lean_object* l_Lean_Parser_Command_structFields___closed__14; static lean_object* l_Lean_Parser_Command_openHiding___closed__1; -static lean_object* l_Lean_Parser_Command_mutual___closed__17; static lean_object* l_Lean_Parser_Command_computedField___closed__12; LEAN_EXPORT lean_object* l_Lean_Parser_Command_import_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_private_parenthesizer___closed__1; @@ -959,6 +970,8 @@ static lean_object* l_Lean_Parser_Command_terminationByElement_parenthesizer___c static lean_object* l___regBuiltin_Lean_Parser_Command_nonrec_formatter___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_mutual_declRange___closed__2; static lean_object* l_Lean_Parser_Command_export_parenthesizer___closed__1; +static lean_object* l_Lean_Parser_Command_terminationByCore_parenthesizer___closed__6; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_export_formatter___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Command_extends_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_terminationByCore_parenthesizer___closed__1; @@ -986,6 +999,7 @@ static lean_object* l_Lean_Parser_Command_computedField_formatter___closed__6; static lean_object* l_Lean_Parser_Command_set__option___closed__3; static lean_object* l_Lean_Parser_Command_decreasingBy___closed__2; static lean_object* l_Lean_Parser_Command_protected_formatter___closed__1; +static lean_object* l_Lean_Parser_Tactic_open___closed__6; static lean_object* l_Lean_Parser_Command_attribute___closed__6; static lean_object* l_Lean_Parser_Command_openHiding_parenthesizer___closed__9; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_opaque_parenthesizer(lean_object*); @@ -994,7 +1008,6 @@ static lean_object* l_Lean_Parser_Command_structure___closed__23; static lean_object* l___regBuiltin_Lean_Parser_Command_eval_declRange___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Command_synth_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_openScoped_parenthesizer___closed__2; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__5; static lean_object* l_Lean_Parser_Command_openScoped___closed__7; static lean_object* l_Lean_Parser_Command_classTk_formatter___closed__2; static lean_object* l_Lean_Parser_Command_computedField_parenthesizer___closed__3; @@ -1015,6 +1028,7 @@ static lean_object* l_Lean_Parser_Command_universe___closed__7; static lean_object* l_Lean_Parser_Command_structSimpleBinder___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Command_optDeclSig_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_axiom_formatter___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__21; static lean_object* l_Lean_Parser_Command_eval___closed__1; static lean_object* l_Lean_Parser_Command_terminationHintMany___closed__12; static lean_object* l_Lean_Parser_Command_inductive___closed__5; @@ -1074,13 +1088,13 @@ static lean_object* l_Lean_Parser_Command_namedPrio___closed__16; static lean_object* l___regBuiltin_Lean_Parser_Command_universe_declRange___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Command_structInstBinder; 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_3161____closed__5; static lean_object* l_Lean_Parser_Command_openRenaming___closed__7; static lean_object* l_Lean_Parser_Command_classTk_formatter___closed__1; static lean_object* l_Lean_Parser_Term_set__option___closed__1; static lean_object* l_Lean_Parser_Command_universe___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_precheckedQuot_parenthesizer___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_universe_parenthesizer___closed__1; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__31; static lean_object* l_Lean_Parser_Command_whereStructInst___closed__21; LEAN_EXPORT lean_object* l_Lean_Parser_Command_set__option_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_universe___closed__8; @@ -1110,11 +1124,13 @@ static lean_object* l_Lean_Parser_Command_partial___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_Command_terminationHint1_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_precheckedQuot_declRange___closed__2; lean_object* l_Lean_Parser_Term_structInst_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Command_openHiding_parenthesizer___closed__10; static lean_object* l___regBuiltin_Lean_Parser_Command_mutual_declRange___closed__1; static lean_object* l_Lean_Parser_Command_declValEqns_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_declModifiers_parenthesizer___closed__23; static lean_object* l___regBuiltin_Lean_Parser_Command_inductive_parenthesizer___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_unsafe_parenthesizer(lean_object*); +static lean_object* l_Lean_Parser_Command_derivingClasses_formatter___closed__7; static lean_object* l_Lean_Parser_Term_quot___closed__17; static lean_object* l_Lean_Parser_Command_inductive_parenthesizer___closed__8; lean_object* l_Lean_Parser_Command_docComment_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1143,6 +1159,7 @@ static lean_object* l_Lean_Parser_Command_terminationByCore_parenthesizer___clos static lean_object* l_Lean_Parser_Command_optDeclSig_parenthesizer___closed__8; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_declaration_formatter(lean_object*); lean_object* l_Lean_Parser_strLit_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__38; static lean_object* l_Lean_Parser_Command_deriving_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Command_openRenaming___closed__9; static lean_object* l_Lean_Parser_Command_eoi_formatter___closed__4; @@ -1160,10 +1177,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_end_formatter(lean_object*, lean_ static lean_object* l_Lean_Parser_Command_terminationByCore_formatter___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_mutual_declRange___closed__5; static lean_object* l_Lean_Parser_Command_structExplicitBinder_formatter___closed__4; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__33; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__52; static lean_object* l___regBuiltin_Lean_Parser_Command_initialize_declRange___closed__5; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Command_check__failure; lean_object* l_Lean_Parser_registerBuiltinNodeKind(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_inductive___closed__11; @@ -1196,7 +1210,6 @@ static lean_object* l_Lean_Parser_Command_openRenaming_parenthesizer___closed__4 static lean_object* l___regBuiltin_Lean_Parser_Command_exit_declRange___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Command_declSig; static lean_object* l___regBuiltin_Lean_Parser_Command_mutual_declRange___closed__6; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__20; lean_object* l_Lean_Parser_withPosition(lean_object*); static lean_object* l_Lean_Parser_Command_set__option___closed__9; static lean_object* l___regBuiltin_Lean_Parser_Command_printAxioms_declRange___closed__3; @@ -1264,22 +1277,22 @@ static lean_object* l_Lean_Parser_Command_abbrev___closed__12; static lean_object* l_Lean_Parser_Command_import_parenthesizer___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_check__failure_declRange___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Tactic_set__option_declRange___closed__2; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__54; static lean_object* l_Lean_Parser_Command_structSimpleBinder_formatter___closed__5; static lean_object* l_Lean_Parser_Term_open___closed__1; static lean_object* l_Lean_Parser_Command_whereStructField_formatter___closed__1; static lean_object* l_Lean_Parser_Command_structure___closed__20; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_synth_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_Command_genInjectiveTheorems_parenthesizer___closed__1; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__45; static lean_object* l_Lean_Parser_Command_export_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Term_quot_parenthesizer___closed__6; static lean_object* l_Lean_Parser_Command_initializeKeyword___closed__4; static lean_object* l_Lean_Parser_Command_whereStructInst_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Command_openHiding___closed__7; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__41; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_end_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_Command_declSig___closed__10; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__9; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__18; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__25; static lean_object* l_Lean_Parser_Command_optionValue_formatter___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_axiom_formatter(lean_object*); static lean_object* l_Lean_Parser_Command_in_formatter___closed__3; @@ -1287,7 +1300,7 @@ static lean_object* l_Lean_Parser_Command_structure_parenthesizer___closed__17; static lean_object* l___regBuiltin_Lean_Parser_Command_quot_declRange___closed__2; static lean_object* l_Lean_Parser_Command_terminationBy___closed__1; static lean_object* l_Lean_Parser_Command_structCtor_parenthesizer___closed__4; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__53; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__35; static lean_object* l___regBuiltin_Lean_Parser_Command_export_declRange___closed__4; static lean_object* l_Lean_Parser_Command_terminationByCore___closed__1; static lean_object* l_Lean_Parser_Command_openDecl___closed__6; @@ -1295,6 +1308,7 @@ static lean_object* l_Lean_Parser_Command_partial___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Command_eval_declRange___closed__4; static lean_object* l_Lean_Parser_Term_precheckedQuot_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Command_abbrev___closed__4; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__27; LEAN_EXPORT lean_object* l_Lean_Parser_Command_structSimpleBinder_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_instance_formatter___closed__6; lean_object* l_Lean_Parser_many1_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1321,12 +1335,15 @@ static lean_object* l_Lean_Parser_Command_decreasingBy_formatter___closed__2; static lean_object* l_Lean_Parser_Command_classTk_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_protected___closed__2; static lean_object* l_Lean_Parser_Command_addDocString_formatter___closed__3; +static lean_object* l_Lean_Parser_Command_terminationByCore___closed__11; static lean_object* l_Lean_Parser_Command_universe_parenthesizer___closed__3; +static lean_object* l_Lean_Parser_Command_openOnly___closed__10; 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_Command_eval_declRange___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Command_optDefDeriving_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_addDocString___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_eval_declRange___closed__7; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__9; static lean_object* l_Lean_Parser_Command_def_formatter___closed__1; static lean_object* l_Lean_Parser_Command_openOnly___closed__8; static lean_object* l_Lean_Parser_Command_extends___closed__8; @@ -1335,6 +1352,8 @@ static lean_object* l_Lean_Parser_Command_classInductive___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_import_declRange___closed__6; static lean_object* l_Lean_Parser_Command_inductive_formatter___closed__9; static lean_object* l_Lean_Parser_Command_noncomputableSection_formatter___closed__5; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__26; +static lean_object* l_Lean_Parser_Command_optDefDeriving_parenthesizer___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Command_section_declRange___closed__3; static lean_object* l_Lean_Parser_Term_set__option_formatter___closed__7; static lean_object* l_Lean_Parser_Command_optDefDeriving___closed__5; @@ -1392,6 +1411,7 @@ static lean_object* l_Lean_Parser_Command_declSig___closed__6; static lean_object* l_Lean_Parser_Command_moduleDoc_formatter___closed__5; static lean_object* l_Lean_Parser_Command_whereStructInst___closed__13; lean_object* l_Lean_Parser_withOpen_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_3161____closed__43; static lean_object* l_Lean_Parser_Command_structure_formatter___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_mutual_declRange___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Command_check_declRange___closed__4; @@ -1456,6 +1476,7 @@ static lean_object* l_Lean_Parser_Command_private_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_terminationByElement_formatter___closed__9; static lean_object* l_Lean_Parser_Command_classInductive_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Command_declValSimple_parenthesizer___closed__6; +static lean_object* l_Lean_Parser_Command_universe_formatter___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Command_universe_declRange___closed__2; static lean_object* l_Lean_Parser_Command_namedPrio___closed__5; static lean_object* l_Lean_Parser_Command_instance___closed__10; @@ -1485,6 +1506,7 @@ static lean_object* l_Lean_Parser_Command_mutual_formatter___closed__10; extern lean_object* l_Lean_PrettyPrinter_parenthesizerAttribute; static lean_object* l_Lean_Parser_Command_structExplicitBinder_formatter___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_openScoped_parenthesizer___closed__1; +LEAN_EXPORT lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3138_(lean_object*); static lean_object* l_Lean_Parser_Command_structSimpleBinder___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_declaration(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_classTk_formatter___closed__2; @@ -1531,7 +1553,6 @@ static lean_object* l_Lean_Parser_Command_openSimple___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Command_init__quot_declRange___closed__4; static lean_object* l_Lean_Parser_Command_classInductive_formatter___closed__6; static lean_object* l_Lean_Parser_Command_variable_parenthesizer___closed__2; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__49; static lean_object* l_Lean_Parser_Command_initialize_parenthesizer___closed__9; static lean_object* l_Lean_Parser_Command_theorem___closed__11; static lean_object* l_Lean_Parser_Command_printAxioms_formatter___closed__3; @@ -1545,14 +1566,12 @@ static lean_object* l_Lean_Parser_Command_structCtor_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Command_eval___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Command_structImplicitBinder_parenthesizer___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Command_computedField; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__32; static lean_object* l_Lean_Parser_Command_structSimpleBinder_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_declModifiers_parenthesizer___closed__13; static lean_object* l_Lean_Parser_Command_namedPrio___closed__3; static lean_object* l_Lean_Parser_Command_export___closed__12; static lean_object* l_Lean_Parser_Command_moduleDoc___closed__9; static lean_object* l_Lean_Parser_Command_whereStructInst___closed__14; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__13; static lean_object* l_Lean_Parser_Command_whereStructInst_formatter___closed__8; static lean_object* l___regBuiltin_Lean_Parser_Command_terminationByCore_formatter___closed__2; static lean_object* l_Lean_Parser_Command_classInductive_formatter___closed__9; @@ -1562,7 +1581,6 @@ static lean_object* l_Lean_Parser_Command_computedFields___closed__8; static lean_object* l___regBuiltin_Lean_Parser_Command_mutual_declRange___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_computedField_formatter(lean_object*); static lean_object* l_Lean_Parser_Command_opaque___closed__1; -static lean_object* l_Lean_Parser_Command_attribute_parenthesizer___closed__12; static lean_object* l_Lean_Parser_Command_classInductive_formatter___closed__2; static lean_object* l_Lean_Parser_Command_namespace___closed__9; static lean_object* l_Lean_Parser_Command_check___closed__1; @@ -1586,16 +1604,16 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_whereStructInst; lean_object* l_Lean_PrettyPrinter_Formatter_many1Unbox_formatter(lean_object*, 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_genInjectiveTheorems___closed__7; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__23; static lean_object* l_Lean_Parser_Command_structureTk_formatter___closed__1; static lean_object* l_Lean_Parser_Command_structSimpleBinder_formatter___closed__4; static lean_object* l_Lean_Parser_Command_noncomputableSection_parenthesizer___closed__2; +static lean_object* l_Lean_Parser_Command_initialize_parenthesizer___closed__12; static lean_object* l___regBuiltin_Lean_Parser_Command_reduce_declRange___closed__2; -LEAN_EXPORT lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063_(lean_object*); static lean_object* l_Lean_Parser_Command_variable_formatter___closed__1; extern lean_object* l_Lean_Parser_Term_binderIdent; static lean_object* l_Lean_Parser_Command_reduce___closed__6; static lean_object* l_Lean_Parser_Command_declModifiers_parenthesizer___closed__3; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__23; static lean_object* l_Lean_Parser_Command_openRenaming_formatter___closed__4; static lean_object* l_Lean_Parser_Command_declModifiers_parenthesizer___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_variable_formatter(lean_object*); @@ -1680,11 +1698,9 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_exit_formatter___closed__ static lean_object* l_Lean_Parser_Command_declValSimple_formatter___closed__8; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_whereStructField_formatter(lean_object*); static lean_object* l_Lean_Parser_Command_namedPrio___closed__18; -static lean_object* l_Lean_Parser_Command_section___closed__6; static lean_object* l_Lean_Parser_Command_declVal___closed__5; static lean_object* l_Lean_Parser_Command_axiom_formatter___closed__5; lean_object* l_Lean_Parser_symbol(lean_object*); -static lean_object* l_Lean_Parser_Command_mutual_parenthesizer___closed__13; static lean_object* l_Lean_Parser_Command_printAxioms_parenthesizer___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Command_structFields_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_namedPrio___closed__11; @@ -1696,11 +1712,13 @@ static lean_object* l_Lean_Parser_Command_terminationHintMany___closed__15; static lean_object* l_Lean_Parser_Term_precheckedQuot_formatter___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Command_initializeKeyword_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_structFields___closed__13; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__10; LEAN_EXPORT lean_object* l_Lean_Parser_Command_terminationByElement_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_declVal___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Command_addDocString; static lean_object* l___regBuiltin_Lean_Parser_Command_declaration_formatter___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_synth(lean_object*); +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__24; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_structCtor_formatter(lean_object*); static lean_object* l_Lean_Parser_Command_structExplicitBinder_formatter___closed__7; static lean_object* l_Lean_Parser_Command_set__option_formatter___closed__1; @@ -1709,6 +1727,7 @@ static lean_object* l_Lean_Parser_Command_inductive_formatter___closed__11; static lean_object* l_Lean_Parser_Command_genInjectiveTheorems___closed__5; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_check_declRange(lean_object*); static lean_object* l_Lean_Parser_Command_declValSimple___closed__5; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__22; static lean_object* l_Lean_Parser_Command_axiom___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Command_check__failure_declRange___closed__4; static lean_object* l_Lean_Parser_Command_protected___closed__3; @@ -1748,6 +1767,7 @@ static lean_object* l_Lean_Parser_Command_structImplicitBinder___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Command_noncomputableSection_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_section_formatter___closed__1; static lean_object* l_Lean_Parser_Command_declValSimple_parenthesizer___closed__8; +static lean_object* l_Lean_Parser_Command_derivingClasses_parenthesizer___closed__7; static lean_object* l_Lean_Parser_Command_eraseAttr_formatter___closed__1; static lean_object* l_Lean_Parser_Command_computedField___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Command_open_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1788,12 +1808,13 @@ static lean_object* l_Lean_Parser_Command_print_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_optDeclSig___closed__4; static lean_object* l_Lean_Parser_Tactic_open___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Term_set__option_parenthesizer___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__39; +static lean_object* l_Lean_Parser_Command_openHiding_formatter___closed__10; static lean_object* l_Lean_Parser_Command_derivingClasses_parenthesizer___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Command_protected_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_partial___closed__3; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_set__option_docString(lean_object*); static lean_object* l_Lean_Parser_Command_moduleDoc___closed__2; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__36; static lean_object* l_Lean_Parser_Command_initializeKeyword___closed__8; static lean_object* l_Lean_Parser_Command_attribute_formatter___closed__4; static lean_object* l_Lean_Parser_Command_structure_parenthesizer___closed__14; @@ -1812,6 +1833,7 @@ static lean_object* l_Lean_Parser_Command_variable___closed__2; static lean_object* l_Lean_Parser_Command_instance_parenthesizer___closed__11; static lean_object* l_Lean_Parser_Command_openRenamingItem_formatter___closed__2; static lean_object* l_Lean_Parser_Command_mutual_parenthesizer___closed__9; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__30; LEAN_EXPORT lean_object* l_Lean_Parser_Command_variable_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_export_formatter___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Command_reduce_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1833,6 +1855,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_check__failure(lean_ static lean_object* l_Lean_Parser_Command_openRenaming___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Command_structCtor; 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_initFn____x40_Lean_Parser_Command___hyg_3161____closed__49; extern lean_object* l_Lean_Parser_Command_docComment; static lean_object* l_Lean_Parser_Command_openRenaming___closed__12; static lean_object* l_Lean_Parser_Command_deriving_parenthesizer___closed__2; @@ -1871,6 +1894,7 @@ static lean_object* l_Lean_Parser_Command_eoi_parenthesizer___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_section_formatter(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_noncomputable_formatter___closed__2; static lean_object* l_Lean_Parser_Command_unsafe_parenthesizer___closed__2; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__11; static lean_object* l_Lean_Parser_Term_quot_formatter___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Command_deriving_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_terminationByElement___closed__11; @@ -1899,7 +1923,6 @@ static lean_object* l_Lean_Parser_Command_whereStructInst___closed__19; static lean_object* l_Lean_Parser_Command_structure___closed__18; LEAN_EXPORT lean_object* l_Lean_Parser_Command_inductive_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_exit_declRange___closed__5; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__24; static lean_object* l_Lean_Parser_Command_terminationBy___closed__10; static lean_object* l_Lean_Parser_Command_synth___closed__9; static lean_object* l_Lean_Parser_Command_optDefDeriving___closed__9; @@ -1918,7 +1941,6 @@ static lean_object* l_Lean_Parser_Command_terminationByElement_parenthesizer___c lean_object* l_Lean_Parser_withoutPosition___lambda__1(lean_object*); static lean_object* l_Lean_Parser_Command_private___closed__7; static lean_object* l_Lean_Parser_Command_private___closed__5; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__27; static lean_object* l_Lean_Parser_Tactic_open_formatter___closed__3; static lean_object* l_Lean_Parser_Command_structure_parenthesizer___closed__15; static lean_object* l_Lean_Parser_Command_attribute___closed__4; @@ -1989,6 +2011,7 @@ static lean_object* l_Lean_Parser_Command_declModifiers_parenthesizer___closed__ LEAN_EXPORT lean_object* l_Lean_Parser_Command_whereStructInst_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_declModifiers_parenthesizer___closed__20; 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_Command_noncomputableSection_parenthesizer___closed__7; lean_object* l_Lean_Parser_Term_attributes_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Command_structExplicitBinder; static lean_object* l___regBuiltin_Lean_Parser_Command_eval_declRange___closed__3; @@ -1996,6 +2019,7 @@ static lean_object* l_Lean_Parser_Command_initialize_formatter___closed__5; static lean_object* l_Lean_Parser_Command_optDeclSig_formatter___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_Command_declModifiersT; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Tactic_open_declRange(lean_object*); +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__20; static lean_object* l_Lean_Parser_Command_reduce___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Command_addDocString_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_mutual_formatter___closed__1; @@ -2029,6 +2053,7 @@ static lean_object* l_Lean_Parser_Command_initialize___closed__6; static lean_object* l_Lean_Parser_Command_import___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Command_end; static lean_object* l_Lean_Parser_Command_printAxioms___closed__8; +static lean_object* l_Lean_Parser_Command_openRenaming_formatter___closed__8; static lean_object* l_Lean_Parser_Command_inductive_formatter___closed__13; static lean_object* l_Lean_Parser_Command_attribute___closed__13; static lean_object* l_Lean_Parser_Command_inductive_formatter___closed__7; @@ -2081,7 +2106,6 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_private; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_set__option(lean_object*); static lean_object* l_Lean_Parser_Command_declModifiers___closed__15; static lean_object* l_Lean_Parser_Command_structImplicitBinder_formatter___closed__7; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__16; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_initialize_formatter(lean_object*); static lean_object* l_Lean_Parser_Command_whereStructField___closed__3; static lean_object* l_Lean_Parser_Command_eval_formatter___closed__2; @@ -2089,6 +2113,7 @@ static lean_object* l_Lean_Parser_Command_declModifiers_parenthesizer___closed__ static lean_object* l_Lean_Parser_Command_deriving_formatter___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Term_open_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_printAxioms_declRange___closed__5; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__4; static lean_object* l_Lean_Parser_Command_structImplicitBinder___closed__1; static lean_object* l_Lean_Parser_Command_init__quot___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_Command_unsafe; @@ -2102,15 +2127,16 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_structFields_parenth static lean_object* l___regBuiltin_Lean_Parser_Command_end_declRange___closed__1; static lean_object* l_Lean_Parser_Command_partial_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_structInstBinder___closed__10; +static lean_object* l_Lean_Parser_Command_openHiding_parenthesizer___closed__11; static lean_object* l___regBuiltin_Lean_Parser_Term_open_formatter___closed__2; static lean_object* l_Lean_Parser_Command_check___closed__4; static lean_object* l_Lean_Parser_Command_mutual_parenthesizer___closed__11; static lean_object* l_Lean_Parser_Command_set__option_formatter___closed__7; +static lean_object* l_Lean_Parser_Command_structCtor_parenthesizer___closed__7; static lean_object* l_Lean_Parser_Command_check__failure___closed__9; static lean_object* l_Lean_Parser_Command_theorem___closed__10; static lean_object* l_Lean_Parser_Command_moduleDoc___closed__10; static lean_object* l_Lean_Parser_Command_whereStructInst___closed__5; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__40; static lean_object* l_Lean_Parser_Command_check__failure___closed__6; static lean_object* l_Lean_Parser_Command_openRenamingItem___closed__9; static lean_object* l_Lean_Parser_Command_declaration___closed__6; @@ -2118,6 +2144,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_section_parenthesizer___c static lean_object* l_Lean_Parser_Term_set__option___closed__9; static lean_object* l_Lean_Parser_Tactic_set__option_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Command_openRenamingItem_parenthesizer___closed__1; +static lean_object* l_Lean_Parser_Command_structExplicitBinder___closed__17; LEAN_EXPORT lean_object* l_Lean_Parser_Command_optDeclSig_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Command_section_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_variable_declRange___closed__1; @@ -2147,6 +2174,7 @@ lean_object* l_Lean_Parser_andthen(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_terminationByElement___closed__9; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_genInjectiveTheorems_docString(lean_object*); static lean_object* l_Lean_Parser_Command_declVal___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__52; LEAN_EXPORT lean_object* l_Lean_Parser_Command_eoi_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_variable_parenthesizer___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_classInductive_parenthesizer___closed__1; @@ -2182,7 +2210,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_quot___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Command_declaration_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_declModifiers_formatter___closed__14; lean_object* l_Lean_Parser_checkColGe(lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__1; static lean_object* l_Lean_Parser_Command_declaration_formatter___closed__9; static lean_object* l_Lean_Parser_Command_eval___closed__5; static lean_object* l_Lean_Parser_Command_init__quot___closed__6; @@ -2200,7 +2227,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_in_declRange___closed__5; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_deriving_declRange(lean_object*); static lean_object* l_Lean_Parser_Term_quot___closed__6; static lean_object* l_Lean_Parser_Term_open_formatter___closed__6; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__38; static lean_object* l_Lean_Parser_Command_structExplicitBinder_formatter___closed__10; static lean_object* l_Lean_Parser_Command_check___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Command_ctor_parenthesizer___closed__2; @@ -2234,14 +2260,12 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_set__option_parenthesizer(lean_obj static lean_object* l_Lean_Parser_Command_decreasingBy___closed__10; LEAN_EXPORT lean_object* l_Lean_Parser_Command_declModifiers___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Command_structure_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__3; static lean_object* l___regBuiltin_Lean_Parser_Tactic_set__option_formatter___closed__1; static lean_object* l_Lean_Parser_Command_init__quot___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_declValEqns_formatter___closed__2; static lean_object* l_Lean_Parser_Command_quot_parenthesizer___closed__4; extern lean_object* l_Lean_Parser_rawIdent; static lean_object* l___regBuiltin_Lean_Parser_Command_openSimple_formatter___closed__2; -static lean_object* l_Lean_Parser_Command_export_parenthesizer___closed__8; static lean_object* l_Lean_Parser_Command_declModifiers___closed__3; static lean_object* l_Lean_Parser_Command_init__quot_formatter___closed__1; static lean_object* l_Lean_Parser_Tactic_set__option_parenthesizer___closed__6; @@ -2266,6 +2290,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_inductive_formatter___clo static lean_object* l___regBuiltin_Lean_Parser_Command_check_declRange___closed__6; static lean_object* l_Lean_Parser_Command_optDeriving_formatter___closed__5; static lean_object* l_Lean_Parser_Command_printAxioms_parenthesizer___closed__4; +LEAN_EXPORT lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Command_open; static lean_object* l_Lean_Parser_Command_def___closed__2; static lean_object* l_Lean_Parser_Command_computedFields_parenthesizer___closed__1; @@ -2279,7 +2304,6 @@ static lean_object* l_Lean_Parser_Term_quot_formatter___closed__8; static lean_object* l_Lean_Parser_Command_declId_parenthesizer___closed__11; static lean_object* l_Lean_Parser_Command_classInductive___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_export_parenthesizer(lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__51; static lean_object* l_Lean_Parser_Command_declaration___closed__10; static lean_object* l___regBuiltin_Lean_Parser_Command_quot_declRange___closed__1; static lean_object* l_Lean_Parser_Command_moduleDoc_formatter___closed__1; @@ -2309,6 +2333,7 @@ static lean_object* l_Lean_Parser_Command_eraseAttr_formatter___closed__3; static lean_object* l_Lean_Parser_Command_openHiding_parenthesizer___closed__7; static lean_object* l_Lean_Parser_Command_check__failure_formatter___closed__3; lean_object* l_Lean_Parser_many(lean_object*); +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__44; static lean_object* l___regBuiltin_Lean_Parser_Command_inductive_formatter___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Command_synth_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_opaque_formatter___closed__1; @@ -2335,11 +2360,11 @@ static lean_object* l_Lean_Parser_Command_init__quot_formatter___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_universe_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_terminationByElement___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Command_variable_formatter___closed__1; -static lean_object* l_Lean_Parser_Command_export___closed__13; static lean_object* l_Lean_Parser_Command_opaque___closed__12; static lean_object* l___regBuiltin_Lean_Parser_Tactic_open_declRange___closed__3; static lean_object* l_Lean_Parser_Command_derivingClasses_formatter___closed__4; static lean_object* l_Lean_Parser_Command_openOnly___closed__4; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__37; static lean_object* l_Lean_Parser_Term_precheckedQuot___closed__5; static lean_object* l_Lean_Parser_Command_initialize___closed__3; static lean_object* l_Lean_Parser_Command_deriving___closed__4; @@ -2360,6 +2385,7 @@ static lean_object* l_Lean_Parser_Command_noncomputableSection___closed__11; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_quot_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_Command_axiom___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_import_formatter(lean_object*); +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__51; static lean_object* l_Lean_Parser_Command_terminationByCore_formatter___closed__3; static lean_object* l_Lean_Parser_Command_def_formatter___closed__9; static lean_object* l_Lean_Parser_Command_structFields_formatter___closed__7; @@ -2381,6 +2407,7 @@ static lean_object* l_Lean_Parser_Command_check___closed__6; static lean_object* l_Lean_Parser_Command_structure_formatter___closed__9; static lean_object* l_Lean_Parser_Command_openDecl_formatter___closed__2; static lean_object* l_Lean_Parser_Command_openHiding_formatter___closed__7; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__12; static lean_object* l___regBuiltin_Lean_Parser_Command_openRenamingItem_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_structSimpleBinder___closed__7; static lean_object* l_Lean_Parser_Command_exit_formatter___closed__3; @@ -2417,7 +2444,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_printAxioms_declRange___c static lean_object* l___regBuiltin_Lean_Parser_Term_open_declRange___closed__2; lean_object* l_Lean_Parser_withPosition_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_declaration_parenthesizer(lean_object*); -static lean_object* l_Lean_Parser_Command_export_formatter___closed__8; 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_Lean_Parser_Command_opaque_formatter___closed__1; static lean_object* l_Lean_Parser_Command_terminationHintMany___closed__4; @@ -2441,9 +2467,7 @@ static lean_object* l_Lean_Parser_Term_quot___closed__19; LEAN_EXPORT lean_object* l_Lean_Parser_Command_universe_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_whereStructField_parenthesizer(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_classInductive_parenthesizer___closed__2; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__42; static lean_object* l_Lean_Parser_Command_openScoped_parenthesizer___closed__4; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__11; static lean_object* l_Lean_Parser_Command_inductive___closed__4; static lean_object* l_Lean_Parser_Command_attribute___closed__5; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_universe(lean_object*); @@ -2466,7 +2490,6 @@ lean_object* l_Lean_Parser_unicodeSymbol_formatter(lean_object*, lean_object*, l LEAN_EXPORT lean_object* l_Lean_Parser_Command_structInstBinder_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_def_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_declId_parenthesizer___closed__2; -LEAN_EXPORT lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3040_(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_printAxioms_formatter___closed__1; static lean_object* l_Lean_Parser_Command_classInductive___closed__12; LEAN_EXPORT lean_object* l_Lean_Parser_Command_namespace_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -2493,9 +2516,9 @@ lean_object* l_Lean_Parser_manyIndent_formatter(lean_object*, lean_object*, lean static lean_object* l_Lean_Parser_Command_private_parenthesizer___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Command_whereStructInst_formatter___closed__2; static lean_object* l_Lean_Parser_Command_open___closed__5; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__26; static lean_object* l_Lean_Parser_Command_computedFields___closed__11; static lean_object* l_Lean_Parser_Command_protected___closed__8; +static lean_object* l_Lean_Parser_Command_openHiding_formatter___closed__11; static lean_object* l_Lean_Parser_Command_openRenaming___closed__1; static lean_object* l_Lean_Parser_Term_quot___closed__11; static lean_object* l_Lean_Parser_Command_theorem_parenthesizer___closed__1; @@ -2524,7 +2547,6 @@ static lean_object* l_Lean_Parser_Command_inductive_formatter___closed__12; static lean_object* l___regBuiltin_Lean_Parser_Command_terminationByCore_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_declId_formatter___closed__2; static lean_object* l_Lean_Parser_Command_namespace___closed__6; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__46; static lean_object* l___regBuiltin_Lean_Parser_Command_namespace_declRange___closed__1; static lean_object* l_Lean_Parser_Command_openRenamingItem_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Command_terminationByElement_formatter___closed__4; @@ -2533,7 +2555,6 @@ static lean_object* l_Lean_Parser_Command_def_formatter___closed__5; static lean_object* l_Lean_Parser_Command_universe_formatter___closed__1; static lean_object* l_Lean_Parser_Command_mutual___closed__8; LEAN_EXPORT lean_object* l_Lean_Parser_Command_computedField_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__44; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_inductive_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_Command_optDefDeriving_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_eoi_formatter___closed__3; @@ -2541,7 +2562,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_namedPrio_formatter___clo static lean_object* l_Lean_Parser_Command_structImplicitBinder___closed__5; static lean_object* l_Lean_Parser_Command_mutual___closed__14; static lean_object* l_Lean_Parser_Command_ctor_parenthesizer___closed__2; -static lean_object* l_Lean_Parser_Command_optNamedPrio___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Command_eval; static lean_object* l_Lean_Parser_Command_whereStructField___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_Command_moduleDoc_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -2600,7 +2620,6 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_openRenaming_parenth static lean_object* l_Lean_Parser_Command_optionValue_formatter___closed__4; static lean_object* l_Lean_Parser_Command_theorem_formatter___closed__4; static lean_object* l_Lean_Parser_Command_visibility___closed__1; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__14; static lean_object* l_Lean_Parser_Command_check_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_optionValue_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_instance___closed__11; @@ -2622,7 +2641,6 @@ static lean_object* l_Lean_Parser_Command_instance_parenthesizer___closed__10; static lean_object* l___regBuiltin_Lean_Parser_Command_reduce_parenthesizer___closed__1; lean_object* l_Lean_Parser_atomic_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_Term_typeSpec_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__19; static lean_object* l_Lean_Parser_Term_quot_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_optDefDeriving_parenthesizer___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Tactic_open___closed__2; @@ -2653,7 +2671,6 @@ static lean_object* l_Lean_Parser_Command_print_parenthesizer___closed__3; 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_Command_noncomputable___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Command_genInjectiveTheorems_declRange___closed__3; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__43; static lean_object* l_Lean_Parser_Command_declValSimple_formatter___closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_Command_namedPrio; static lean_object* l_Lean_Parser_Command_declaration_parenthesizer___closed__10; @@ -2667,6 +2684,7 @@ static lean_object* l_Lean_Parser_Command_optDefDeriving___closed__1; static lean_object* l_Lean_Parser_Command_derivingClasses_parenthesizer___closed__5; lean_object* l_Lean_Parser_withoutPosition_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_check__failure_formatter___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__19; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_reduce(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_section_formatter___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_extends_formatter___closed__2; @@ -2698,6 +2716,7 @@ static lean_object* l_Lean_Parser_Command_optDeclSig_formatter___closed__7; lean_object* l_Lean_Parser_unicodeSymbol(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_quot_declRange___closed__7; lean_object* l_Lean_Parser_Term_binderDefault_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Command_structure_formatter___closed__19; LEAN_EXPORT lean_object* l_Lean_Parser_Command_decreasingBy_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_nonrec_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_optDeclSig_formatter___closed__2; @@ -2741,6 +2760,7 @@ static lean_object* l_Lean_Parser_Command_ctor_formatter___closed__1; static lean_object* l_Lean_Parser_Command_example___closed__8; static lean_object* l_Lean_Parser_Command_structExplicitBinder___closed__14; static lean_object* l___regBuiltin_Lean_Parser_Tactic_open_declRange___closed__1; +static lean_object* l_Lean_Parser_Tactic_open_parenthesizer___closed__7; static lean_object* l_Lean_Parser_Command_structExplicitBinder_formatter___closed__5; static lean_object* l_Lean_Parser_Command_instance___closed__5; static lean_object* l_Lean_Parser_Command_quot_formatter___closed__3; @@ -2766,6 +2786,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_private_formatter(lean_object*, l static lean_object* l___regBuiltin_Lean_Parser_Term_set__option_declRange___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Command_genInjectiveTheorems_docString___closed__1; static lean_object* l_Lean_Parser_Command_structFields___closed__16; +static lean_object* l_Lean_Parser_Command_in_formatter___closed__4; static lean_object* l_Lean_Parser_Command_exit___closed__5; static lean_object* l_Lean_Parser_Command_declId___closed__10; LEAN_EXPORT lean_object* l_Lean_Parser_Command_openRenaming; @@ -2830,10 +2851,13 @@ static lean_object* l_Lean_Parser_Command_deriving_parenthesizer___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Command_export_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_declVal_parenthesizer___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Command_moduleDoc_parenthesizer___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__53; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__45; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_namedPrio_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_Command_unsafe___closed__6; static lean_object* l_Lean_Parser_Command_declId___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Command_nonrec_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__34; extern lean_object* l_Lean_Parser_Command_commentBody; static lean_object* l_Lean_Parser_Command_structCtor_formatter___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Term_quot_declRange___closed__4; @@ -2851,6 +2875,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_structSimpleBinder_parent LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_terminationByElement_formatter(lean_object*); static lean_object* l_Lean_Parser_Command_declModifiers_formatter___closed__3; 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_Command_decreasingBy___closed__11; LEAN_EXPORT lean_object* l_Lean_Parser_Command_set__option; static lean_object* l_Lean_Parser_Command_export___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_moduleDoc_declRange___closed__7; @@ -2862,7 +2887,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_init__quot_declRange___cl static lean_object* l___regBuiltin_Lean_Parser_Command_import_declRange___closed__1; static lean_object* l_Lean_Parser_Command_inductive_parenthesizer___closed__14; static lean_object* l_Lean_Parser_Command_computedFields_formatter___closed__6; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__47; LEAN_EXPORT lean_object* l_Lean_Parser_Command_moduleDoc; static lean_object* l_Lean_Parser_Command_axiom_formatter___closed__3; static lean_object* l_Lean_Parser_Command_synth___closed__4; @@ -2893,7 +2917,6 @@ static lean_object* l_Lean_Parser_Command_opaque_parenthesizer___closed__2; lean_object* l_Lean_Parser_Command_commentBody_formatter(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___regBuiltin_Lean_Parser_Tactic_open_declRange___closed__7; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__34; static lean_object* l_Lean_Parser_Command_declaration_parenthesizer___closed__3; lean_object* l_Lean_Parser_Tactic_tacticSeq_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_structure_parenthesizer___closed__8; @@ -2917,6 +2940,7 @@ static lean_object* l_Lean_Parser_Command_instance___closed__3; static lean_object* l_Lean_Parser_Command_open_parenthesizer___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_universe_formatter(lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_open_parenthesizer(lean_object*); +static lean_object* l_Lean_Parser_Command_openOnly_parenthesizer___closed__6; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_end(lean_object*); static lean_object* l_Lean_Parser_Command_structFields___closed__1; static lean_object* l_Lean_Parser_Command_whereStructField_parenthesizer___closed__3; @@ -2950,15 +2974,12 @@ lean_object* l_Lean_PrettyPrinter_Parenthesizer_trailingNode_parenthesizer(lean_ static lean_object* l_Lean_Parser_Command_structInstBinder___closed__12; static lean_object* l_Lean_Parser_Command_synth_formatter___closed__3; static lean_object* l_Lean_Parser_Command_optDeclSig_parenthesizer___closed__3; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__35; static lean_object* l_Lean_Parser_Command_print_formatter___closed__6; static lean_object* l_Lean_Parser_Command_openHiding_parenthesizer___closed__6; static lean_object* l_Lean_Parser_Command_terminationBy___closed__7; static lean_object* l_Lean_Parser_Command_def_parenthesizer___closed__8; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_check__failure_declRange(lean_object*); -static lean_object* l_Lean_Parser_Command_mutual_formatter___closed__13; static lean_object* l_Lean_Parser_Command_derivingClasses_formatter___closed__5; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__17; static lean_object* l___regBuiltin_Lean_Parser_Command_optDeclSig_formatter___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_def_formatter___closed__1; static lean_object* l_Lean_Parser_Command_genInjectiveTheorems___closed__8; @@ -2995,7 +3016,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_openScoped_formatter___cl lean_object* l_Lean_Parser_Term_bracketedBinder_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Command_partial; static lean_object* l_Lean_Parser_Command_abbrev___closed__11; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__4; static lean_object* l_Lean_Parser_Command_namedPrio___closed__14; static lean_object* l_Lean_Parser_Command_whereStructField___closed__6; static lean_object* l_Lean_Parser_Command_terminationBy_formatter___closed__2; @@ -3013,6 +3033,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_structExplicitBinder_form static lean_object* l___regBuiltin_Lean_Parser_Command_structCtor_parenthesizer___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Command_attribute_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_extends_parenthesizer___closed__2; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__33; static lean_object* l_Lean_Parser_Command_theorem___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Command_theorem_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_deriving_parenthesizer___closed__1; @@ -3053,7 +3074,6 @@ lean_object* l_Lean_Parser_lookahead(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_variable_formatter___closed__2; static lean_object* l_Lean_Parser_Command_structFields_parenthesizer___closed__6; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_theorem_parenthesizer(lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__30; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_terminationByElement_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_Command_structFields_formatter___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Term_quot_declRange___closed__5; @@ -3061,7 +3081,6 @@ static lean_object* l_Lean_Parser_Command_export_parenthesizer___closed__6; static lean_object* l_Lean_Parser_Command_openDecl_formatter___closed__1; static lean_object* l_Lean_Parser_Command_terminationHintMany___closed__11; static lean_object* l_Lean_Parser_Command_synth___closed__6; -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__10; static lean_object* l_Lean_Parser_Command_structFields_parenthesizer___closed__7; static lean_object* l_Lean_Parser_Command_initializeKeyword_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Term_quot___closed__9; @@ -3072,10 +3091,13 @@ static lean_object* l_Lean_Parser_Command_ctor_formatter___closed__4; static lean_object* l_Lean_Parser_Command_def_parenthesizer___closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_Command_inductive_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_initializeKeyword_formatter___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__16; static lean_object* l_Lean_Parser_Command_genInjectiveTheorems___closed__4; +static lean_object* l_Lean_Parser_Command_noncomputableSection___closed__12; static lean_object* l_Lean_Parser_Command_classInductive_formatter___closed__10; static lean_object* l_Lean_Parser_Command_mutual_formatter___closed__3; static lean_object* l_Lean_Parser_Command_extends___closed__7; +static lean_object* l_Lean_Parser_Command_in_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Command_declModifiers_parenthesizer___closed__7; static lean_object* l_Lean_Parser_Command_whereStructInst_parenthesizer___closed__8; static lean_object* l_Lean_Parser_Command_whereStructInst___closed__22; @@ -3099,6 +3121,7 @@ static lean_object* l_Lean_Parser_Command_structure___closed__22; static lean_object* l___regBuiltin_Lean_Parser_Command_noncomputableSection_declRange___closed__5; static lean_object* l_Lean_Parser_Command_optDeclSig_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_optionValue___closed__2; +static lean_object* l_Lean_Parser_Command_terminationByElement___closed__14; static lean_object* l___regBuiltin_Lean_Parser_Command_declaration_formatter___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Tactic_open(lean_object*); static lean_object* l_Lean_Parser_Command_declaration_formatter___closed__11; @@ -3178,7 +3201,6 @@ lean_object* l_Lean_Parser_incQuotDepth_parenthesizer(lean_object*, lean_object* static lean_object* l___regBuiltin_Lean_Parser_Command_in_declRange___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Command_check_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Command_whereStructInst_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__21; static lean_object* l_Lean_Parser_Command_export_formatter___closed__3; static lean_object* l_Lean_Parser_Command_abbrev___closed__2; static lean_object* l_Lean_Parser_Command_whereStructInst___closed__8; @@ -3212,6 +3234,8 @@ static lean_object* l_Lean_Parser_Command_def_formatter___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Command_initialize_parenthesizer___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_open_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_initialize_parenthesizer___closed__8; +static lean_object* l_Lean_Parser_Command_universe_parenthesizer___closed__5; +static lean_object* l_Lean_Parser_Command_structCtor_formatter___closed__7; static lean_object* l_Lean_Parser_Command_export_formatter___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Command_computedField_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_mutual_parenthesizer___closed__5; @@ -3230,6 +3254,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_example_parenthesizer(lean_object LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_terminationByCore_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_Command_addDocString___closed__7; static lean_object* l_Lean_Parser_Tactic_set__option___closed__3; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__48; static lean_object* l_Lean_Parser_Command_namedPrio___closed__12; static lean_object* l_Lean_Parser_Command_instance___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_set__option_declRange(lean_object*); @@ -3269,7 +3294,6 @@ static lean_object* l_Lean_Parser_Command_decreasingBy___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_ctor_parenthesizer(lean_object*); lean_object* l_Lean_Parser_group_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_open_parenthesizer___closed__6; -static lean_object* l_Lean_Parser_Command_optNamedPrio_parenthesizer___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_open_declRange___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Command_terminationHintMany_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_check_declRange___closed__3; @@ -3303,6 +3327,7 @@ static lean_object* l_Lean_Parser_Command_classInductive___closed__4; static lean_object* l_Lean_Parser_Command_noncomputable___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Command_ctor_parenthesizer___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_structCtor_parenthesizer___closed__2; +static lean_object* l_Lean_Parser_Command_initialize_formatter___closed__12; static lean_object* l_Lean_Parser_Command_classInductive___closed__6; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_genInjectiveTheorems(lean_object*); static lean_object* l_Lean_Parser_Command_structure_formatter___closed__7; @@ -3365,6 +3390,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_deriving_parenthesizer___ static lean_object* l_Lean_Parser_Command_computedField_parenthesizer___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Command_partial_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_structure_formatter___closed__6; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__7; static lean_object* l___regBuiltin_Lean_Parser_Command_universe_formatter___closed__2; static lean_object* l_Lean_Parser_Term_open_formatter___closed__1; lean_object* l_Lean_Parser_withOpenDecl_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -3389,7 +3415,6 @@ static lean_object* l_Lean_Parser_Command_noncomputable___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_precheckedQuot_formatter(lean_object*); static lean_object* l_Lean_Parser_Term_open_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Command_optDeriving_parenthesizer___closed__4; -static lean_object* l_Lean_Parser_Command_attribute_formatter___closed__12; static lean_object* l_Lean_Parser_Command_variable___closed__9; static lean_object* l_Lean_Parser_Command_exit___closed__8; static lean_object* l_Lean_Parser_Command_theorem_parenthesizer___closed__6; @@ -3404,6 +3429,7 @@ static lean_object* l_Lean_Parser_Command_classInductive___closed__16; static lean_object* l_Lean_Parser_Command_computedFields_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Command_inductive_formatter___closed__2; static lean_object* l_Lean_Parser_Command_declSig___closed__2; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__17; static lean_object* l_Lean_Parser_Command_declaration_parenthesizer___closed__13; static lean_object* l_Lean_Parser_Command_openOnly___closed__6; static lean_object* l_Lean_Parser_Command_structInstBinder_formatter___closed__8; @@ -3418,6 +3444,7 @@ static lean_object* l_Lean_Parser_Command_example_parenthesizer___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Command_decreasingBy_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_open_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Command_openRenamingItem_formatter___closed__4; +static lean_object* l_Lean_Parser_Command_optDefDeriving___closed__10; static lean_object* l_Lean_Parser_Command_terminationByElement_formatter___closed__10; static lean_object* l_Lean_Parser_Command_derivingClasses_formatter___closed__1; static lean_object* l_Lean_Parser_Command_declaration___closed__17; @@ -3431,6 +3458,7 @@ static lean_object* l_Lean_Parser_Command_structure_parenthesizer___closed__7; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_export(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Command_whereStructInst_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_structCtor___closed__3; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__36; static lean_object* l_Lean_Parser_Command_declModifiers_formatter___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_structExplicitBinder_parenthesizer(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Command_visibility_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -3454,6 +3482,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_computedField_formatter__ static lean_object* l___regBuiltin_Lean_Parser_Command_partial_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_terminationByElement_formatter___closed__5; static lean_object* l_Lean_Parser_Command_openOnly_formatter___closed__5; +static lean_object* l_Lean_Parser_Tactic_open_formatter___closed__7; static lean_object* l_Lean_Parser_Term_precheckedQuot_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_attribute___closed__11; static lean_object* l___regBuiltin_Lean_Parser_Command_optDeclSig_formatter___closed__2; @@ -3523,7 +3552,6 @@ static lean_object* l_Lean_Parser_Command_check__failure_formatter___closed__2; static lean_object* l_Lean_Parser_Command_def___closed__3; static lean_object* l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_synth___closed__8; -static lean_object* l_Lean_Parser_Command_end___closed__9; static lean_object* l_Lean_Parser_Command_terminationBy_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_variable_formatter___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Command_decreasingBy_parenthesizer___closed__2; @@ -3596,6 +3624,7 @@ static lean_object* l_Lean_Parser_Command_openHiding_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Command_openHiding___closed__4; static lean_object* l_Lean_Parser_Command_moduleDoc_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_structure_formatter___closed__12; +static lean_object* l_Lean_Parser_Command_openRenaming_parenthesizer___closed__8; static lean_object* l_Lean_Parser_Command_export___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Term_precheckedQuot_formatter___closed__2; static lean_object* l_Lean_Parser_Command_quot_parenthesizer___closed__8; @@ -3640,6 +3669,7 @@ static lean_object* l_Lean_Parser_Command_terminationBy___closed__4; static lean_object* l_Lean_Parser_Command_genInjectiveTheorems_formatter___closed__2; static lean_object* l_Lean_Parser_Command_noncomputableSection_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Command_initializeKeyword_parenthesizer___closed__1; +static lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__40; static lean_object* l_Lean_Parser_Command_private___closed__6; static lean_object* l_Lean_Parser_Command_exit___closed__2; static lean_object* l_Lean_Parser_Command_terminationByElement_formatter___closed__12; @@ -5584,30 +5614,40 @@ return x_3; static lean_object* _init_l_Lean_Parser_Command_terminationByCore___closed__8() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_skip; +x_2 = l_Lean_Parser_Command_terminationByCore___closed__7; +x_3 = l_Lean_Parser_andthen(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Command_terminationByCore___closed__9() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_terminationByCore___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_terminationByCore___closed__7; +x_3 = l_Lean_Parser_Command_terminationByCore___closed__8; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Command_terminationByCore___closed__9() { +static lean_object* _init_l_Lean_Parser_Command_terminationByCore___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_terminationByCore___closed__3; -x_2 = l_Lean_Parser_Command_terminationByCore___closed__8; +x_2 = l_Lean_Parser_Command_terminationByCore___closed__9; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_terminationByCore___closed__10() { +static lean_object* _init_l_Lean_Parser_Command_terminationByCore___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_terminationByCore___closed__2; -x_2 = l_Lean_Parser_Command_terminationByCore___closed__9; +x_2 = l_Lean_Parser_Command_terminationByCore___closed__10; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -5616,7 +5656,7 @@ static lean_object* _init_l_Lean_Parser_Command_terminationByCore() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Command_terminationByCore___closed__10; +x_1 = l_Lean_Parser_Command_terminationByCore___closed__11; return x_1; } } @@ -5691,30 +5731,40 @@ return x_3; static lean_object* _init_l_Lean_Parser_Command_decreasingBy___closed__8() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_skip; +x_2 = l_Lean_Parser_Command_decreasingBy___closed__7; +x_3 = l_Lean_Parser_andthen(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Command_decreasingBy___closed__9() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_decreasingBy___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_decreasingBy___closed__7; +x_3 = l_Lean_Parser_Command_decreasingBy___closed__8; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Command_decreasingBy___closed__9() { +static lean_object* _init_l_Lean_Parser_Command_decreasingBy___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_decreasingBy___closed__3; -x_2 = l_Lean_Parser_Command_decreasingBy___closed__8; +x_2 = l_Lean_Parser_Command_decreasingBy___closed__9; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_decreasingBy___closed__10() { +static lean_object* _init_l_Lean_Parser_Command_decreasingBy___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_decreasingBy___closed__2; -x_2 = l_Lean_Parser_Command_decreasingBy___closed__9; +x_2 = l_Lean_Parser_Command_decreasingBy___closed__10; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -5723,7 +5773,7 @@ static lean_object* _init_l_Lean_Parser_Command_decreasingBy() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Command_decreasingBy___closed__10; +x_1 = l_Lean_Parser_Command_decreasingBy___closed__11; return x_1; } } @@ -5772,13 +5822,23 @@ return x_3; static lean_object* _init_l_Lean_Parser_Command_terminationByElement___closed__5() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_skip; +x_2 = l_Lean_Parser_Command_terminationByElement___closed__4; +x_3 = l_Lean_Parser_andthen(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Command_terminationByElement___closed__6() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_terminationByElement___closed__4; +x_1 = l_Lean_Parser_Command_terminationByElement___closed__5; x_2 = l_Lean_Parser_many(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_terminationByElement___closed__6() { +static lean_object* _init_l_Lean_Parser_Command_terminationByElement___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -5788,73 +5848,73 @@ x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_terminationByElement___closed__7() { +static lean_object* _init_l_Lean_Parser_Command_terminationByElement___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_terminationHintMany___closed__5; -x_2 = l_Lean_Parser_Command_terminationByElement___closed__6; +x_2 = l_Lean_Parser_Command_terminationByElement___closed__7; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_terminationByElement___closed__8() { +static lean_object* _init_l_Lean_Parser_Command_terminationByElement___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_terminationByElement___closed__5; -x_2 = l_Lean_Parser_Command_terminationByElement___closed__7; +x_1 = l_Lean_Parser_Command_terminationByElement___closed__6; +x_2 = l_Lean_Parser_Command_terminationByElement___closed__8; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_terminationByElement___closed__9() { +static lean_object* _init_l_Lean_Parser_Command_terminationByElement___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_terminationByElement___closed__4; -x_2 = l_Lean_Parser_Command_terminationByElement___closed__8; +x_2 = l_Lean_Parser_Command_terminationByElement___closed__9; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_terminationByElement___closed__10() { +static lean_object* _init_l_Lean_Parser_Command_terminationByElement___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_skip; -x_2 = l_Lean_Parser_Command_terminationByElement___closed__9; +x_2 = l_Lean_Parser_Command_terminationByElement___closed__10; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_terminationByElement___closed__11() { +static lean_object* _init_l_Lean_Parser_Command_terminationByElement___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_terminationByElement___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_terminationByElement___closed__10; +x_3 = l_Lean_Parser_Command_terminationByElement___closed__11; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Command_terminationByElement___closed__12() { +static lean_object* _init_l_Lean_Parser_Command_terminationByElement___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_terminationByElement___closed__3; -x_2 = l_Lean_Parser_Command_terminationByElement___closed__11; +x_2 = l_Lean_Parser_Command_terminationByElement___closed__12; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_terminationByElement___closed__13() { +static lean_object* _init_l_Lean_Parser_Command_terminationByElement___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_terminationByElement___closed__2; -x_2 = l_Lean_Parser_Command_terminationByElement___closed__12; +x_2 = l_Lean_Parser_Command_terminationByElement___closed__13; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -5863,7 +5923,7 @@ static lean_object* _init_l_Lean_Parser_Command_terminationByElement() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Command_terminationByElement___closed__13; +x_1 = l_Lean_Parser_Command_terminationByElement___closed__14; return x_1; } } @@ -5903,7 +5963,7 @@ static lean_object* _init_l_Lean_Parser_Command_terminationBy___closed__4() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("termination_by ", 15); +x_1 = lean_mk_string_from_bytes("termination_by", 14); return x_1; } } @@ -6578,7 +6638,7 @@ static lean_object* _init_l_Lean_Parser_Command_namedPrio___closed__4() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("(", 1); +x_1 = lean_mk_string_from_bytes(" (", 2); return x_1; } } @@ -6755,18 +6815,8 @@ return x_1; static lean_object* _init_l_Lean_Parser_Command_optNamedPrio___closed__1() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_skip; -x_2 = l_Lean_Parser_Command_namedPrio; -x_3 = l_Lean_Parser_andthen(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Parser_Command_optNamedPrio___closed__2() { -_start: -{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_optNamedPrio___closed__1; +x_1 = l_Lean_Parser_Command_namedPrio; x_2 = l_Lean_Parser_optional(x_1); return x_2; } @@ -6775,7 +6825,7 @@ static lean_object* _init_l_Lean_Parser_Command_optNamedPrio() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Command_optNamedPrio___closed__2; +x_1 = l_Lean_Parser_Command_optNamedPrio___closed__1; return x_1; } } @@ -8703,8 +8753,18 @@ return x_3; static lean_object* _init_l_Lean_Parser_Command_optDefDeriving___closed__9() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_skip; +x_2 = l_Lean_Parser_Command_optDefDeriving___closed__8; +x_3 = l_Lean_Parser_andthen(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Command_optDefDeriving___closed__10() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_optDefDeriving___closed__8; +x_1 = l_Lean_Parser_Command_optDefDeriving___closed__9; x_2 = l_Lean_Parser_optional(x_1); return x_2; } @@ -8713,7 +8773,7 @@ static lean_object* _init_l_Lean_Parser_Command_optDefDeriving() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Command_optDefDeriving___closed__9; +x_1 = l_Lean_Parser_Command_optDefDeriving___closed__10; return x_1; } } @@ -10439,23 +10499,40 @@ return x_5; static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder___closed__4() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("(", 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Command_structExplicitBinder___closed__4; +x_2 = l_Lean_Parser_symbol(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder___closed__6() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_ctor___closed__8; -x_2 = l_Lean_Parser_Command_namedPrio___closed__5; +x_2 = l_Lean_Parser_Command_structExplicitBinder___closed__5; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder___closed__5() { +static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder___closed__7() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_structExplicitBinder___closed__4; +x_1 = l_Lean_Parser_Command_structExplicitBinder___closed__6; x_2 = l_Lean_Parser_atomic(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder___closed__6() { +static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder___closed__8() { _start: { lean_object* x_1; lean_object* x_2; @@ -10464,7 +10541,7 @@ x_2 = l_Lean_Parser_many1(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder___closed__7() { +static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -10474,92 +10551,92 @@ x_3 = l_Lean_Parser_orelse(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder___closed__8() { +static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder___closed__10() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_structExplicitBinder___closed__7; +x_1 = l_Lean_Parser_Command_structExplicitBinder___closed__9; x_2 = l_Lean_Parser_optional(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder___closed__9() { +static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_optDeclSig; -x_2 = l_Lean_Parser_Command_structExplicitBinder___closed__8; +x_2 = l_Lean_Parser_Command_structExplicitBinder___closed__10; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder___closed__10() { +static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_structExplicitBinder___closed__6; -x_2 = l_Lean_Parser_Command_structExplicitBinder___closed__9; +x_1 = l_Lean_Parser_Command_structExplicitBinder___closed__8; +x_2 = l_Lean_Parser_Command_structExplicitBinder___closed__11; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder___closed__11() { +static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_quot___closed__11; -x_2 = l_Lean_Parser_Command_structExplicitBinder___closed__10; +x_2 = l_Lean_Parser_Command_structExplicitBinder___closed__12; x_3 = l_Lean_Parser_adaptCacheableContext(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder___closed__12() { +static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_structExplicitBinder___closed__11; +x_1 = l_Lean_Parser_Command_structExplicitBinder___closed__13; x_2 = l_Lean_Parser_Term_quot___closed__14; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder___closed__13() { +static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder___closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_structExplicitBinder___closed__5; -x_2 = l_Lean_Parser_Command_structExplicitBinder___closed__12; +x_1 = l_Lean_Parser_Command_structExplicitBinder___closed__7; +x_2 = l_Lean_Parser_Command_structExplicitBinder___closed__14; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder___closed__14() { +static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder___closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_structExplicitBinder___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_structExplicitBinder___closed__13; +x_3 = l_Lean_Parser_Command_structExplicitBinder___closed__15; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder___closed__15() { +static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder___closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_structExplicitBinder___closed__3; -x_2 = l_Lean_Parser_Command_structExplicitBinder___closed__14; +x_2 = l_Lean_Parser_Command_structExplicitBinder___closed__16; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder___closed__16() { +static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder___closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_structExplicitBinder___closed__2; -x_2 = l_Lean_Parser_Command_structExplicitBinder___closed__15; +x_2 = l_Lean_Parser_Command_structExplicitBinder___closed__17; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -10568,7 +10645,7 @@ static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Command_structExplicitBinder___closed__16; +x_1 = l_Lean_Parser_Command_structExplicitBinder___closed__18; return x_1; } } @@ -10644,7 +10721,7 @@ static lean_object* _init_l_Lean_Parser_Command_structImplicitBinder___closed__8 _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_structExplicitBinder___closed__6; +x_1 = l_Lean_Parser_Command_structExplicitBinder___closed__8; x_2 = l_Lean_Parser_Command_declSig; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; @@ -10919,7 +10996,7 @@ static lean_object* _init_l_Lean_Parser_Command_structSimpleBinder___closed__6() { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_structSimpleBinder___closed__5; -x_2 = l_Lean_Parser_Command_structExplicitBinder___closed__9; +x_2 = l_Lean_Parser_Command_structExplicitBinder___closed__11; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -11588,36 +11665,37 @@ return x_2; static lean_object* _init_l_Lean_Parser_Command_structure___closed__8() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes(" where ", 7); -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Command_structure___closed__7; +x_2 = l_Lean_Parser_Term_optType; +x_3 = l_Lean_Parser_andthen(x_1, x_2); +return x_3; } } static lean_object* _init_l_Lean_Parser_Command_structure___closed__9() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_structure___closed__8; -x_2 = l_Lean_Parser_symbol(x_1); -return x_2; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Command_structure___closed__6; +x_2 = l_Lean_Parser_Command_structure___closed__8; +x_3 = l_Lean_Parser_andthen(x_1, x_2); +return x_3; } } static lean_object* _init_l_Lean_Parser_Command_structure___closed__10() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_namedPrio___closed__11; -x_2 = l_Lean_Parser_Command_structure___closed__9; -x_3 = l_Lean_Parser_orelse(x_1, x_2); -return x_3; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(" where ", 7); +return x_1; } } static lean_object* _init_l_Lean_Parser_Command_structure___closed__11() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_structCtor; -x_2 = l_Lean_Parser_optional(x_1); +x_1 = l_Lean_Parser_Command_structure___closed__10; +x_2 = l_Lean_Parser_symbol(x_1); return x_2; } } @@ -11625,37 +11703,37 @@ static lean_object* _init_l_Lean_Parser_Command_structure___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_structure___closed__11; -x_2 = l_Lean_Parser_Command_structFields; -x_3 = l_Lean_Parser_andthen(x_1, x_2); +x_1 = l_Lean_Parser_Command_namedPrio___closed__11; +x_2 = l_Lean_Parser_Command_structure___closed__11; +x_3 = l_Lean_Parser_orelse(x_1, x_2); return x_3; } } static lean_object* _init_l_Lean_Parser_Command_structure___closed__13() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_structure___closed__10; -x_2 = l_Lean_Parser_Command_structure___closed__12; -x_3 = l_Lean_Parser_andthen(x_1, x_2); -return x_3; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Command_structCtor; +x_2 = l_Lean_Parser_optional(x_1); +return x_2; } } static lean_object* _init_l_Lean_Parser_Command_structure___closed__14() { _start: { -lean_object* x_1; lean_object* x_2; +lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_structure___closed__13; -x_2 = l_Lean_Parser_optional(x_1); -return x_2; +x_2 = l_Lean_Parser_Command_structFields; +x_3 = l_Lean_Parser_andthen(x_1, x_2); +return x_3; } } static lean_object* _init_l_Lean_Parser_Command_structure___closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_structure___closed__14; -x_2 = l_Lean_Parser_Command_optDeriving; +x_1 = l_Lean_Parser_Command_structure___closed__12; +x_2 = l_Lean_Parser_Command_structure___closed__14; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -11663,19 +11741,18 @@ return x_3; static lean_object* _init_l_Lean_Parser_Command_structure___closed__16() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_optType; -x_2 = l_Lean_Parser_Command_structure___closed__15; -x_3 = l_Lean_Parser_andthen(x_1, x_2); -return x_3; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Command_structure___closed__15; +x_2 = l_Lean_Parser_optional(x_1); +return x_2; } } static lean_object* _init_l_Lean_Parser_Command_structure___closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_structure___closed__7; -x_2 = l_Lean_Parser_Command_structure___closed__16; +x_1 = l_Lean_Parser_Command_structure___closed__16; +x_2 = l_Lean_Parser_Command_optDeriving; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -11684,7 +11761,7 @@ static lean_object* _init_l_Lean_Parser_Command_structure___closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_structure___closed__6; +x_1 = l_Lean_Parser_Command_structure___closed__9; x_2 = l_Lean_Parser_Command_structure___closed__17; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; @@ -11946,7 +12023,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(188u); +x_1 = lean_unsigned_to_nat(187u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11958,7 +12035,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(191u); +x_1 = lean_unsigned_to_nat(190u); x_2 = lean_unsigned_to_nat(50u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11986,7 +12063,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(188u); +x_1 = lean_unsigned_to_nat(187u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11998,7 +12075,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(188u); +x_1 = lean_unsigned_to_nat(187u); x_2 = lean_unsigned_to_nat(41u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13775,14 +13852,26 @@ return x_3; static lean_object* _init_l_Lean_Parser_Command_terminationByElement_formatter___closed__4() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Command_optDeclSig_formatter___closed__5; +x_2 = l_Lean_Parser_Command_terminationByElement_formatter___closed__3; +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; +} +} +static lean_object* _init_l_Lean_Parser_Command_terminationByElement_formatter___closed__5() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_terminationByElement_formatter___closed__3; +x_1 = l_Lean_Parser_Command_terminationByElement_formatter___closed__4; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_many_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_terminationByElement_formatter___closed__5() { +static lean_object* _init_l_Lean_Parser_Command_terminationByElement_formatter___closed__6() { _start: { lean_object* x_1; lean_object* x_2; @@ -13792,7 +13881,7 @@ lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_terminationByElement_formatter___closed__6() { +static lean_object* _init_l_Lean_Parser_Command_terminationByElement_formatter___closed__7() { _start: { lean_object* x_1; lean_object* x_2; @@ -13802,83 +13891,83 @@ lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_terminationByElement_formatter___closed__7() { +static lean_object* _init_l_Lean_Parser_Command_terminationByElement_formatter___closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_terminationByElement_formatter___closed__6; +x_1 = l_Lean_Parser_Command_terminationByElement_formatter___closed__7; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_optional_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_terminationByElement_formatter___closed__8() { +static lean_object* _init_l_Lean_Parser_Command_terminationByElement_formatter___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_quot_formatter___closed__3; -x_2 = l_Lean_Parser_Command_terminationByElement_formatter___closed__7; +x_2 = l_Lean_Parser_Command_terminationByElement_formatter___closed__8; 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; } } -static lean_object* _init_l_Lean_Parser_Command_terminationByElement_formatter___closed__9() { +static lean_object* _init_l_Lean_Parser_Command_terminationByElement_formatter___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_terminationByElement_formatter___closed__5; -x_2 = l_Lean_Parser_Command_terminationByElement_formatter___closed__8; +x_1 = l_Lean_Parser_Command_terminationByElement_formatter___closed__6; +x_2 = l_Lean_Parser_Command_terminationByElement_formatter___closed__9; 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; } } -static lean_object* _init_l_Lean_Parser_Command_terminationByElement_formatter___closed__10() { +static lean_object* _init_l_Lean_Parser_Command_terminationByElement_formatter___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_terminationByElement_formatter___closed__4; -x_2 = l_Lean_Parser_Command_terminationByElement_formatter___closed__9; +x_1 = l_Lean_Parser_Command_terminationByElement_formatter___closed__5; +x_2 = l_Lean_Parser_Command_terminationByElement_formatter___closed__10; 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; } } -static lean_object* _init_l_Lean_Parser_Command_terminationByElement_formatter___closed__11() { +static lean_object* _init_l_Lean_Parser_Command_terminationByElement_formatter___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_terminationByElement_formatter___closed__3; -x_2 = l_Lean_Parser_Command_terminationByElement_formatter___closed__10; +x_2 = l_Lean_Parser_Command_terminationByElement_formatter___closed__11; 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; } } -static lean_object* _init_l_Lean_Parser_Command_terminationByElement_formatter___closed__12() { +static lean_object* _init_l_Lean_Parser_Command_terminationByElement_formatter___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_moduleDoc_formatter___closed__4; -x_2 = l_Lean_Parser_Command_terminationByElement_formatter___closed__11; +x_2 = l_Lean_Parser_Command_terminationByElement_formatter___closed__12; 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; } } -static lean_object* _init_l_Lean_Parser_Command_terminationByElement_formatter___closed__13() { +static lean_object* _init_l_Lean_Parser_Command_terminationByElement_formatter___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_terminationByElement___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_terminationByElement_formatter___closed__12; +x_3 = l_Lean_Parser_Command_terminationByElement_formatter___closed__13; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -13891,7 +13980,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_terminationByElement_formatter(le { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Command_terminationByElement_formatter___closed__1; -x_7 = l_Lean_Parser_Command_terminationByElement_formatter___closed__13; +x_7 = l_Lean_Parser_Command_terminationByElement_formatter___closed__14; 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; } @@ -13983,7 +14072,7 @@ static lean_object* _init_l_Lean_Parser_Command_terminationBy_formatter___closed _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_moduleDoc_formatter___closed__4; +x_1 = l_Lean_Parser_Command_declModifiers_formatter___closed__13; x_2 = l_Lean_Parser_Command_terminationBy_formatter___closed__4; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -14071,7 +14160,7 @@ static lean_object* _init_l_Lean_Parser_Command_terminationHintMany_formatter___ { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_declId_formatter___closed__4; -x_2 = l_Lean_Parser_Command_terminationByElement_formatter___closed__5; +x_2 = l_Lean_Parser_Command_terminationByElement_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); @@ -14101,12 +14190,12 @@ return x_2; LEAN_EXPORT lean_object* l_Lean_Parser_Command_terminationHintMany_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; 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; -x_7 = l_Lean_Parser_Command_terminationByElement_formatter___closed__7; +lean_object* x_7; 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; +x_7 = l_Lean_Parser_Command_terminationByElement_formatter___closed__8; x_8 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_8, 0, x_1); lean_closure_set(x_8, 1, x_7); -x_9 = l_Lean_Parser_Command_terminationByElement_formatter___closed__5; +x_9 = l_Lean_Parser_Command_terminationByElement_formatter___closed__6; x_10 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_10, 0, x_9); lean_closure_set(x_10, 1, x_8); @@ -14114,27 +14203,29 @@ x_11 = l_Lean_Parser_Command_declId_formatter___closed__4; x_12 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_12, 0, x_11); lean_closure_set(x_12, 1, x_10); -x_13 = l_Lean_Parser_Command_moduleDoc_formatter___closed__4; -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_12); -x_15 = lean_alloc_closure((void*)(l_Lean_Parser_group_formatter), 6, 1); +x_13 = lean_alloc_closure((void*)(l_Lean_ppIndent_formatter), 6, 1); +lean_closure_set(x_13, 0, x_12); +x_14 = l_Lean_Parser_Command_moduleDoc_formatter___closed__4; +x_15 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_15, 0, x_14); -x_16 = lean_alloc_closure((void*)(l_Lean_Parser_many1Indent_formatter), 6, 1); +lean_closure_set(x_15, 1, x_13); +x_16 = lean_alloc_closure((void*)(l_Lean_Parser_group_formatter), 6, 1); lean_closure_set(x_16, 0, x_15); -x_17 = l_Lean_Parser_Command_terminationHintMany_formatter___closed__4; -x_18 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); -lean_closure_set(x_18, 0, x_17); -lean_closure_set(x_18, 1, x_16); -x_19 = l_Lean_Parser_Command_terminationHintMany___closed__2; -x_20 = lean_unsigned_to_nat(1024u); -x_21 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); -lean_closure_set(x_21, 0, x_19); -lean_closure_set(x_21, 1, x_20); -lean_closure_set(x_21, 2, x_18); -x_22 = l_Lean_Parser_Command_terminationHintMany_formatter___closed__1; -x_23 = l_Lean_PrettyPrinter_Formatter_orelse_formatter(x_22, x_21, x_2, x_3, x_4, x_5, x_6); -return x_23; +x_17 = lean_alloc_closure((void*)(l_Lean_Parser_many1Indent_formatter), 6, 1); +lean_closure_set(x_17, 0, x_16); +x_18 = l_Lean_Parser_Command_terminationHintMany_formatter___closed__4; +x_19 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); +lean_closure_set(x_19, 0, x_18); +lean_closure_set(x_19, 1, x_17); +x_20 = l_Lean_Parser_Command_terminationHintMany___closed__2; +x_21 = lean_unsigned_to_nat(1024u); +x_22 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); +lean_closure_set(x_22, 0, x_20); +lean_closure_set(x_22, 1, x_21); +lean_closure_set(x_22, 2, x_19); +x_23 = l_Lean_Parser_Command_terminationHintMany_formatter___closed__1; +x_24 = l_Lean_PrettyPrinter_Formatter_orelse_formatter(x_23, x_22, x_2, x_3, x_4, x_5, x_6); +return x_24; } } static lean_object* _init_l_Lean_Parser_Command_terminationHint1_formatter___closed__1() { @@ -14236,10 +14327,22 @@ return x_3; static lean_object* _init_l_Lean_Parser_Command_terminationByCore_formatter___closed__5() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Command_declModifiers_formatter___closed__13; +x_2 = l_Lean_Parser_Command_terminationByCore_formatter___closed__4; +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; +} +} +static lean_object* _init_l_Lean_Parser_Command_terminationByCore_formatter___closed__6() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_terminationByCore___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_terminationByCore_formatter___closed__4; +x_3 = l_Lean_Parser_Command_terminationByCore_formatter___closed__5; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -14252,7 +14355,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_terminationByCore_formatter(lean_ { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Command_terminationByCore_formatter___closed__1; -x_7 = l_Lean_Parser_Command_terminationByCore_formatter___closed__5; +x_7 = l_Lean_Parser_Command_terminationByCore_formatter___closed__6; 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; } @@ -14351,10 +14454,22 @@ return x_3; static lean_object* _init_l_Lean_Parser_Command_decreasingBy_formatter___closed__6() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Command_declModifiers_formatter___closed__13; +x_2 = l_Lean_Parser_Command_decreasingBy_formatter___closed__5; +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; +} +} +static lean_object* _init_l_Lean_Parser_Command_decreasingBy_formatter___closed__7() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_decreasingBy___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_decreasingBy_formatter___closed__5; +x_3 = l_Lean_Parser_Command_decreasingBy_formatter___closed__6; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -14367,7 +14482,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_decreasingBy_formatter(lean_objec { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Command_decreasingBy_formatter___closed__1; -x_7 = l_Lean_Parser_Command_decreasingBy_formatter___closed__6; +x_7 = l_Lean_Parser_Command_decreasingBy_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; } @@ -14660,11 +14775,23 @@ lean_closure_set(x_3, 1, x_2); return x_3; } } +static lean_object* _init_l_Lean_Parser_Command_optDefDeriving_formatter___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Command_declModifiers_formatter___closed__13; +x_2 = l_Lean_Parser_Command_optDefDeriving_formatter___closed__5; +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_Command_optDefDeriving_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_Parser_Command_optDefDeriving_formatter___closed__5; +x_6 = l_Lean_Parser_Command_optDefDeriving_formatter___closed__6; x_7 = l_Lean_Parser_optional_formatter(x_6, x_1, x_2, x_3, x_4, x_5); return x_7; } @@ -15363,23 +15490,11 @@ 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_Command_optNamedPrio_formatter___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_optDeclSig_formatter___closed__5; -x_2 = l___regBuiltin_Lean_Parser_Command_namedPrio_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_Command_optNamedPrio_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_Parser_Command_optNamedPrio_formatter___closed__1; +x_6 = l___regBuiltin_Lean_Parser_Command_namedPrio_formatter___closed__2; x_7 = l_Lean_Parser_optional_formatter(x_6, x_1, x_2, x_3, x_4, x_5); return x_7; } @@ -16238,42 +16353,52 @@ return x_1; static lean_object* _init_l_Lean_Parser_Command_derivingClasses_formatter___closed__3() { _start: { +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Command_derivingClasses_formatter___closed__2; +x_2 = lean_alloc_closure((void*)(l_Lean_ppIndent_formatter), 6, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Command_derivingClasses_formatter___closed__4() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_derivingClasses_formatter___closed__1; -x_2 = l_Lean_Parser_Command_derivingClasses_formatter___closed__2; +x_2 = l_Lean_Parser_Command_derivingClasses_formatter___closed__3; 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; } } -static lean_object* _init_l_Lean_Parser_Command_derivingClasses_formatter___closed__4() { +static lean_object* _init_l_Lean_Parser_Command_derivingClasses_formatter___closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_derivingClasses_formatter___closed__3; +x_1 = l_Lean_Parser_Command_derivingClasses_formatter___closed__4; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_optional_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_derivingClasses_formatter___closed__5() { +static lean_object* _init_l_Lean_Parser_Command_derivingClasses_formatter___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_declId_formatter___closed__4; -x_2 = l_Lean_Parser_Command_derivingClasses_formatter___closed__4; +x_2 = l_Lean_Parser_Command_derivingClasses_formatter___closed__5; 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; } } -static lean_object* _init_l_Lean_Parser_Command_derivingClasses_formatter___closed__6() { +static lean_object* _init_l_Lean_Parser_Command_derivingClasses_formatter___closed__7() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_derivingClasses_formatter___closed__5; +x_1 = l_Lean_Parser_Command_derivingClasses_formatter___closed__6; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_group_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; @@ -16283,7 +16408,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_derivingClasses_formatter(lean_ob _start: { lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; lean_object* x_10; -x_6 = l_Lean_Parser_Command_derivingClasses_formatter___closed__6; +x_6 = l_Lean_Parser_Command_derivingClasses_formatter___closed__7; x_7 = l_Lean_Parser_Command_declId___closed__6; x_8 = l_Lean_Parser_Command_declId_formatter___closed__3; x_9 = 0; @@ -16540,7 +16665,7 @@ static lean_object* _init_l_Lean_Parser_Command_inductive_formatter___closed__12 _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Parser_Command_optDeclSig_formatter___closed__2; +x_1 = l_Lean_Parser_Command_abbrev_formatter___closed__3; x_2 = l_Lean_Parser_Command_inductive_formatter___closed__11; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -17137,7 +17262,7 @@ static lean_object* _init_l_Lean_Parser_Command_structCtor_formatter___closed__5 { lean_object* x_1; lean_object* x_2; x_1 = l_Lean_Parser_Command_structCtor_formatter___closed__4; -x_2 = lean_alloc_closure((void*)(l_Lean_Parser_atomic_formatter), 6, 1); +x_2 = lean_alloc_closure((void*)(l_Lean_ppIndent_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } @@ -17145,10 +17270,20 @@ return x_2; static lean_object* _init_l_Lean_Parser_Command_structCtor_formatter___closed__6() { _start: { +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Command_structCtor_formatter___closed__5; +x_2 = lean_alloc_closure((void*)(l_Lean_Parser_atomic_formatter), 6, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Command_structCtor_formatter___closed__7() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_structCtor___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_structCtor_formatter___closed__5; +x_3 = l_Lean_Parser_Command_structCtor_formatter___closed__6; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -17161,7 +17296,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_structCtor_formatter(lean_object* { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Command_structCtor_formatter___closed__1; -x_7 = l_Lean_Parser_Command_structCtor_formatter___closed__6; +x_7 = l_Lean_Parser_Command_structCtor_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; } @@ -17220,26 +17355,36 @@ return x_7; static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder_formatter___closed__2() { _start: { +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Command_structExplicitBinder___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; +} +} +static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder_formatter___closed__3() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_ctor_formatter___closed__5; -x_2 = l_Lean_Parser_Command_namedPrio_formatter___closed__2; +x_2 = l_Lean_Parser_Command_structExplicitBinder_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; } } -static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder_formatter___closed__3() { +static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder_formatter___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_structExplicitBinder_formatter___closed__2; +x_1 = l_Lean_Parser_Command_structExplicitBinder_formatter___closed__3; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_atomic_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder_formatter___closed__4() { +static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder_formatter___closed__5() { _start: { lean_object* x_1; lean_object* x_2; @@ -17249,7 +17394,7 @@ lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder_formatter___closed__5() { +static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder_formatter___closed__6() { _start: { lean_object* x_1; @@ -17257,7 +17402,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Term_binderTactic_formatter), 5, return x_1; } } -static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder_formatter___closed__6() { +static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder_formatter___closed__7() { _start: { lean_object* x_1; @@ -17265,67 +17410,67 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Term_binderDefault_formatter), 5, return x_1; } } -static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder_formatter___closed__7() { +static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder_formatter___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_structExplicitBinder_formatter___closed__5; -x_2 = l_Lean_Parser_Command_structExplicitBinder_formatter___closed__6; +x_1 = l_Lean_Parser_Command_structExplicitBinder_formatter___closed__6; +x_2 = l_Lean_Parser_Command_structExplicitBinder_formatter___closed__7; 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_Command_structExplicitBinder_formatter___closed__8() { +static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder_formatter___closed__9() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_structExplicitBinder_formatter___closed__7; +x_1 = l_Lean_Parser_Command_structExplicitBinder_formatter___closed__8; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_optional_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder_formatter___closed__9() { +static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder_formatter___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_abbrev_formatter___closed__3; -x_2 = l_Lean_Parser_Command_structExplicitBinder_formatter___closed__8; +x_2 = l_Lean_Parser_Command_structExplicitBinder_formatter___closed__9; 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; } } -static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder_formatter___closed__10() { +static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder_formatter___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_structExplicitBinder_formatter___closed__4; -x_2 = l_Lean_Parser_Command_structExplicitBinder_formatter___closed__9; +x_1 = l_Lean_Parser_Command_structExplicitBinder_formatter___closed__5; +x_2 = l_Lean_Parser_Command_structExplicitBinder_formatter___closed__10; 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; } } -static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder_formatter___closed__11() { +static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder_formatter___closed__12() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_structExplicitBinder_formatter___closed__10; +x_1 = l_Lean_Parser_Command_structExplicitBinder_formatter___closed__11; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_withoutPosition_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder_formatter___closed__12() { +static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder_formatter___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_structExplicitBinder_formatter___closed__11; +x_1 = l_Lean_Parser_Command_structExplicitBinder_formatter___closed__12; x_2 = l_Lean_Parser_Term_quot_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); @@ -17333,25 +17478,25 @@ lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder_formatter___closed__13() { +static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder_formatter___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_structExplicitBinder_formatter___closed__3; -x_2 = l_Lean_Parser_Command_structExplicitBinder_formatter___closed__12; +x_1 = l_Lean_Parser_Command_structExplicitBinder_formatter___closed__4; +x_2 = l_Lean_Parser_Command_structExplicitBinder_formatter___closed__13; 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; } } -static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder_formatter___closed__14() { +static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder_formatter___closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_structExplicitBinder___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_structExplicitBinder_formatter___closed__13; +x_3 = l_Lean_Parser_Command_structExplicitBinder_formatter___closed__14; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -17364,7 +17509,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_structExplicitBinder_formatter(le { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Command_structExplicitBinder_formatter___closed__1; -x_7 = l_Lean_Parser_Command_structExplicitBinder_formatter___closed__14; +x_7 = l_Lean_Parser_Command_structExplicitBinder_formatter___closed__15; 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; } @@ -17456,7 +17601,7 @@ static lean_object* _init_l_Lean_Parser_Command_structImplicitBinder_formatter__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_structExplicitBinder_formatter___closed__4; +x_1 = l_Lean_Parser_Command_structExplicitBinder_formatter___closed__5; x_2 = l___regBuiltin_Lean_Parser_Command_declSig_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); @@ -17741,7 +17886,7 @@ static lean_object* _init_l_Lean_Parser_Command_structSimpleBinder_formatter___c { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_Lean_Parser_Command_optDeclSig_formatter___closed__2; -x_2 = l_Lean_Parser_Command_structExplicitBinder_formatter___closed__8; +x_2 = l_Lean_Parser_Command_structExplicitBinder_formatter___closed__9; 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); @@ -18045,20 +18190,22 @@ return x_2; static lean_object* _init_l_Lean_Parser_Command_structure_formatter___closed__6() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_structure___closed__8; -x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_formatter), 6, 1); -lean_closure_set(x_2, 0, x_1); -return x_2; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Command_structure_formatter___closed__5; +x_2 = l_Lean_Parser_Command_optDeclSig_formatter___closed__8; +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; } } static lean_object* _init_l_Lean_Parser_Command_structure_formatter___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_namedPrio_formatter___closed__6; +x_1 = l_Lean_Parser_Command_structure_formatter___closed__4; x_2 = l_Lean_Parser_Command_structure_formatter___closed__6; -x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_orelse_formatter), 7, 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; @@ -18068,8 +18215,8 @@ static lean_object* _init_l_Lean_Parser_Command_structure_formatter___closed__8( _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___regBuiltin_Lean_Parser_Command_structCtor_formatter___closed__2; -x_2 = lean_alloc_closure((void*)(l_Lean_Parser_optional_formatter), 6, 1); +x_1 = l_Lean_Parser_Command_structure_formatter___closed__7; +x_2 = lean_alloc_closure((void*)(l_Lean_ppIndent_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } @@ -18077,22 +18224,20 @@ return x_2; static lean_object* _init_l_Lean_Parser_Command_structure_formatter___closed__9() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_structure_formatter___closed__8; -x_2 = l___regBuiltin_Lean_Parser_Command_structFields_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_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Command_structure___closed__10; +x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_formatter), 6, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; } } static lean_object* _init_l_Lean_Parser_Command_structure_formatter___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_structure_formatter___closed__7; +x_1 = l_Lean_Parser_Command_namedPrio_formatter___closed__6; x_2 = l_Lean_Parser_Command_structure_formatter___closed__9; -x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 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; @@ -18102,7 +18247,7 @@ static lean_object* _init_l_Lean_Parser_Command_structure_formatter___closed__11 _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_structure_formatter___closed__10; +x_1 = l___regBuiltin_Lean_Parser_Command_structCtor_formatter___closed__2; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_optional_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; @@ -18113,7 +18258,7 @@ static lean_object* _init_l_Lean_Parser_Command_structure_formatter___closed__12 { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_structure_formatter___closed__11; -x_2 = l___regBuiltin_Lean_Parser_Command_optDeriving_formatter___closed__2; +x_2 = l___regBuiltin_Lean_Parser_Command_structFields_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); @@ -18124,7 +18269,7 @@ static lean_object* _init_l_Lean_Parser_Command_structure_formatter___closed__13 _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_optDeclSig_formatter___closed__8; +x_1 = l_Lean_Parser_Command_structure_formatter___closed__10; x_2 = l_Lean_Parser_Command_structure_formatter___closed__12; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -18135,58 +18280,68 @@ return x_3; static lean_object* _init_l_Lean_Parser_Command_structure_formatter___closed__14() { _start: { +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Command_structure_formatter___closed__13; +x_2 = lean_alloc_closure((void*)(l_Lean_Parser_optional_formatter), 6, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Command_structure_formatter___closed__15() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_structure_formatter___closed__5; -x_2 = l_Lean_Parser_Command_structure_formatter___closed__13; +x_1 = l_Lean_Parser_Command_structure_formatter___closed__14; +x_2 = l___regBuiltin_Lean_Parser_Command_optDeriving_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; } } -static lean_object* _init_l_Lean_Parser_Command_structure_formatter___closed__15() { +static lean_object* _init_l_Lean_Parser_Command_structure_formatter___closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_structure_formatter___closed__4; -x_2 = l_Lean_Parser_Command_structure_formatter___closed__14; +x_1 = l_Lean_Parser_Command_structure_formatter___closed__8; +x_2 = l_Lean_Parser_Command_structure_formatter___closed__15; 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; } } -static lean_object* _init_l_Lean_Parser_Command_structure_formatter___closed__16() { +static lean_object* _init_l_Lean_Parser_Command_structure_formatter___closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_Lean_Parser_Command_declId_formatter___closed__2; -x_2 = l_Lean_Parser_Command_structure_formatter___closed__15; +x_2 = l_Lean_Parser_Command_structure_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); lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_structure_formatter___closed__17() { +static lean_object* _init_l_Lean_Parser_Command_structure_formatter___closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_structure_formatter___closed__2; -x_2 = l_Lean_Parser_Command_structure_formatter___closed__16; +x_2 = l_Lean_Parser_Command_structure_formatter___closed__17; 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; } } -static lean_object* _init_l_Lean_Parser_Command_structure_formatter___closed__18() { +static lean_object* _init_l_Lean_Parser_Command_structure_formatter___closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_structure___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_structure_formatter___closed__17; +x_3 = l_Lean_Parser_Command_structure_formatter___closed__18; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -18199,7 +18354,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_structure_formatter(lean_object* { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Command_structure_formatter___closed__1; -x_7 = l_Lean_Parser_Command_structure_formatter___closed__18; +x_7 = l_Lean_Parser_Command_structure_formatter___closed__19; 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; } @@ -20177,14 +20332,26 @@ return x_3; static lean_object* _init_l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__4() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Command_optDeclSig_parenthesizer___closed__5; +x_2 = l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__3; +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; +} +} +static lean_object* _init_l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__5() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__3; +x_1 = l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__4; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_many_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__5() { +static lean_object* _init_l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__6() { _start: { lean_object* x_1; lean_object* x_2; @@ -20194,7 +20361,7 @@ lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__6() { +static lean_object* _init_l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__7() { _start: { lean_object* x_1; lean_object* x_2; @@ -20204,83 +20371,83 @@ lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__7() { +static lean_object* _init_l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__6; +x_1 = l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__7; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_optional_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__8() { +static lean_object* _init_l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_quot_parenthesizer___closed__3; -x_2 = l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__7; +x_2 = l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__8; 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; } } -static lean_object* _init_l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__9() { +static lean_object* _init_l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__5; -x_2 = l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__8; +x_1 = l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__6; +x_2 = l_Lean_Parser_Command_terminationByElement_parenthesizer___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); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__10() { +static lean_object* _init_l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__4; -x_2 = l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__9; +x_1 = l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__5; +x_2 = l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__10; 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; } } -static lean_object* _init_l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__11() { +static lean_object* _init_l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__3; -x_2 = l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__10; +x_2 = l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__11; 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; } } -static lean_object* _init_l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__12() { +static lean_object* _init_l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_moduleDoc_parenthesizer___closed__4; -x_2 = l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__11; +x_2 = l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__12; 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; } } -static lean_object* _init_l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__13() { +static lean_object* _init_l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_terminationByElement___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__12; +x_3 = l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__13; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -20293,7 +20460,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_terminationByElement_parenthesize { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__13; +x_7 = l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__14; 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; } @@ -20385,7 +20552,7 @@ static lean_object* _init_l_Lean_Parser_Command_terminationBy_parenthesizer___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_moduleDoc_parenthesizer___closed__4; +x_1 = l_Lean_Parser_Command_declModifiers_parenthesizer___closed__13; x_2 = l_Lean_Parser_Command_terminationBy_parenthesizer___closed__4; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -20473,7 +20640,7 @@ static lean_object* _init_l_Lean_Parser_Command_terminationHintMany_parenthesize { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_declId_parenthesizer___closed__4; -x_2 = l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__5; +x_2 = l_Lean_Parser_Command_terminationByElement_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); @@ -20503,12 +20670,12 @@ return x_2; LEAN_EXPORT lean_object* l_Lean_Parser_Command_terminationHintMany_parenthesizer(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; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_7 = l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__7; +lean_object* x_7; 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; +x_7 = l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__8; x_8 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_8, 0, x_1); lean_closure_set(x_8, 1, x_7); -x_9 = l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__5; +x_9 = l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__6; x_10 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_10, 0, x_9); lean_closure_set(x_10, 1, x_8); @@ -20516,27 +20683,29 @@ x_11 = l_Lean_Parser_Command_declId_parenthesizer___closed__4; x_12 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_12, 0, x_11); lean_closure_set(x_12, 1, x_10); -x_13 = l_Lean_Parser_Command_moduleDoc_parenthesizer___closed__4; -x_14 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); -lean_closure_set(x_14, 0, x_13); -lean_closure_set(x_14, 1, x_12); -x_15 = lean_alloc_closure((void*)(l_Lean_Parser_group_parenthesizer), 6, 1); +x_13 = lean_alloc_closure((void*)(l_Lean_Parser_ppIndent_parenthesizer), 6, 1); +lean_closure_set(x_13, 0, x_12); +x_14 = l_Lean_Parser_Command_moduleDoc_parenthesizer___closed__4; +x_15 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_15, 0, x_14); -x_16 = lean_alloc_closure((void*)(l_Lean_Parser_many1Indent_parenthesizer), 6, 1); +lean_closure_set(x_15, 1, x_13); +x_16 = lean_alloc_closure((void*)(l_Lean_Parser_group_parenthesizer), 6, 1); lean_closure_set(x_16, 0, x_15); -x_17 = l_Lean_Parser_Command_terminationHintMany_parenthesizer___closed__4; -x_18 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); -lean_closure_set(x_18, 0, x_17); -lean_closure_set(x_18, 1, x_16); -x_19 = l_Lean_Parser_Command_terminationHintMany___closed__2; -x_20 = lean_unsigned_to_nat(1024u); -x_21 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); -lean_closure_set(x_21, 0, x_19); -lean_closure_set(x_21, 1, x_20); -lean_closure_set(x_21, 2, x_18); -x_22 = l_Lean_Parser_Command_terminationHintMany_parenthesizer___closed__1; -x_23 = l_Lean_PrettyPrinter_Parenthesizer_withAntiquot_parenthesizer(x_22, x_21, x_2, x_3, x_4, x_5, x_6); -return x_23; +x_17 = lean_alloc_closure((void*)(l_Lean_Parser_many1Indent_parenthesizer), 6, 1); +lean_closure_set(x_17, 0, x_16); +x_18 = l_Lean_Parser_Command_terminationHintMany_parenthesizer___closed__4; +x_19 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); +lean_closure_set(x_19, 0, x_18); +lean_closure_set(x_19, 1, x_17); +x_20 = l_Lean_Parser_Command_terminationHintMany___closed__2; +x_21 = lean_unsigned_to_nat(1024u); +x_22 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); +lean_closure_set(x_22, 0, x_20); +lean_closure_set(x_22, 1, x_21); +lean_closure_set(x_22, 2, x_19); +x_23 = l_Lean_Parser_Command_terminationHintMany_parenthesizer___closed__1; +x_24 = l_Lean_PrettyPrinter_Parenthesizer_withAntiquot_parenthesizer(x_23, x_22, x_2, x_3, x_4, x_5, x_6); +return x_24; } } static lean_object* _init_l_Lean_Parser_Command_terminationHint1_parenthesizer___closed__1() { @@ -20638,10 +20807,22 @@ return x_3; static lean_object* _init_l_Lean_Parser_Command_terminationByCore_parenthesizer___closed__5() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Command_declModifiers_parenthesizer___closed__13; +x_2 = l_Lean_Parser_Command_terminationByCore_parenthesizer___closed__4; +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; +} +} +static lean_object* _init_l_Lean_Parser_Command_terminationByCore_parenthesizer___closed__6() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_terminationByCore___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_terminationByCore_parenthesizer___closed__4; +x_3 = l_Lean_Parser_Command_terminationByCore_parenthesizer___closed__5; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -20654,7 +20835,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_terminationByCore_parenthesizer(l { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Command_terminationByCore_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Command_terminationByCore_parenthesizer___closed__5; +x_7 = l_Lean_Parser_Command_terminationByCore_parenthesizer___closed__6; 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; } @@ -20753,10 +20934,22 @@ return x_3; static lean_object* _init_l_Lean_Parser_Command_decreasingBy_parenthesizer___closed__6() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Command_declModifiers_parenthesizer___closed__13; +x_2 = l_Lean_Parser_Command_decreasingBy_parenthesizer___closed__5; +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; +} +} +static lean_object* _init_l_Lean_Parser_Command_decreasingBy_parenthesizer___closed__7() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_decreasingBy___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_decreasingBy_parenthesizer___closed__5; +x_3 = l_Lean_Parser_Command_decreasingBy_parenthesizer___closed__6; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -20769,7 +20962,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_decreasingBy_parenthesizer(lean_o { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Command_decreasingBy_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Command_decreasingBy_parenthesizer___closed__6; +x_7 = l_Lean_Parser_Command_decreasingBy_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; } @@ -21062,11 +21255,23 @@ lean_closure_set(x_3, 1, x_2); return x_3; } } +static lean_object* _init_l_Lean_Parser_Command_optDefDeriving_parenthesizer___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Command_declModifiers_parenthesizer___closed__13; +x_2 = l_Lean_Parser_Command_optDefDeriving_parenthesizer___closed__5; +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_Command_optDefDeriving_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; -x_6 = l_Lean_Parser_Command_optDefDeriving_parenthesizer___closed__5; +x_6 = l_Lean_Parser_Command_optDefDeriving_parenthesizer___closed__6; x_7 = l_Lean_Parser_optional_parenthesizer(x_6, x_1, x_2, x_3, x_4, x_5); return x_7; } @@ -21767,23 +21972,11 @@ 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_Command_optNamedPrio_parenthesizer___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_optDeclSig_parenthesizer___closed__5; -x_2 = l___regBuiltin_Lean_Parser_Command_namedPrio_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_Command_optNamedPrio_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; -x_6 = l_Lean_Parser_Command_optNamedPrio_parenthesizer___closed__1; +x_6 = l___regBuiltin_Lean_Parser_Command_namedPrio_parenthesizer___closed__2; x_7 = l_Lean_Parser_optional_parenthesizer(x_6, x_1, x_2, x_3, x_4, x_5); return x_7; } @@ -22642,42 +22835,52 @@ return x_1; static lean_object* _init_l_Lean_Parser_Command_derivingClasses_parenthesizer___closed__3() { _start: { +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Command_derivingClasses_parenthesizer___closed__2; +x_2 = lean_alloc_closure((void*)(l_Lean_Parser_ppIndent_parenthesizer), 6, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Command_derivingClasses_parenthesizer___closed__4() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_derivingClasses_parenthesizer___closed__1; -x_2 = l_Lean_Parser_Command_derivingClasses_parenthesizer___closed__2; +x_2 = l_Lean_Parser_Command_derivingClasses_parenthesizer___closed__3; 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; } } -static lean_object* _init_l_Lean_Parser_Command_derivingClasses_parenthesizer___closed__4() { +static lean_object* _init_l_Lean_Parser_Command_derivingClasses_parenthesizer___closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_derivingClasses_parenthesizer___closed__3; +x_1 = l_Lean_Parser_Command_derivingClasses_parenthesizer___closed__4; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_optional_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_derivingClasses_parenthesizer___closed__5() { +static lean_object* _init_l_Lean_Parser_Command_derivingClasses_parenthesizer___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_declId_parenthesizer___closed__4; -x_2 = l_Lean_Parser_Command_derivingClasses_parenthesizer___closed__4; +x_2 = l_Lean_Parser_Command_derivingClasses_parenthesizer___closed__5; 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; } } -static lean_object* _init_l_Lean_Parser_Command_derivingClasses_parenthesizer___closed__6() { +static lean_object* _init_l_Lean_Parser_Command_derivingClasses_parenthesizer___closed__7() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_derivingClasses_parenthesizer___closed__5; +x_1 = l_Lean_Parser_Command_derivingClasses_parenthesizer___closed__6; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_group_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; @@ -22687,7 +22890,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_derivingClasses_parenthesizer(lea _start: { lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; lean_object* x_10; -x_6 = l_Lean_Parser_Command_derivingClasses_parenthesizer___closed__6; +x_6 = l_Lean_Parser_Command_derivingClasses_parenthesizer___closed__7; x_7 = l_Lean_Parser_Command_declId___closed__6; x_8 = l_Lean_Parser_Command_declId_parenthesizer___closed__3; x_9 = 0; @@ -22944,7 +23147,7 @@ static lean_object* _init_l_Lean_Parser_Command_inductive_parenthesizer___closed _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Parser_Command_optDeclSig_parenthesizer___closed__2; +x_1 = l_Lean_Parser_Command_abbrev_parenthesizer___closed__3; x_2 = l_Lean_Parser_Command_inductive_parenthesizer___closed__11; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -23541,7 +23744,7 @@ static lean_object* _init_l_Lean_Parser_Command_structCtor_parenthesizer___close { lean_object* x_1; lean_object* x_2; x_1 = l_Lean_Parser_Command_structCtor_parenthesizer___closed__4; -x_2 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_atomic_parenthesizer), 6, 1); +x_2 = lean_alloc_closure((void*)(l_Lean_Parser_ppIndent_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } @@ -23549,10 +23752,20 @@ return x_2; static lean_object* _init_l_Lean_Parser_Command_structCtor_parenthesizer___closed__6() { _start: { +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Command_structCtor_parenthesizer___closed__5; +x_2 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_atomic_parenthesizer), 6, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Command_structCtor_parenthesizer___closed__7() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_structCtor___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_structCtor_parenthesizer___closed__5; +x_3 = l_Lean_Parser_Command_structCtor_parenthesizer___closed__6; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -23565,7 +23778,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_structCtor_parenthesizer(lean_obj { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Command_structCtor_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Command_structCtor_parenthesizer___closed__6; +x_7 = l_Lean_Parser_Command_structCtor_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; } @@ -23624,26 +23837,36 @@ return x_7; static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__2() { _start: { +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Command_structExplicitBinder___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; +} +} +static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__3() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_ctor_parenthesizer___closed__5; -x_2 = l_Lean_Parser_Command_namedPrio_parenthesizer___closed__2; +x_2 = l_Lean_Parser_Command_structExplicitBinder_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; } } -static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__3() { +static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__2; +x_1 = l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__3; x_2 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_atomic_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__4() { +static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__5() { _start: { lean_object* x_1; lean_object* x_2; @@ -23653,7 +23876,7 @@ lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__5() { +static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__6() { _start: { lean_object* x_1; @@ -23661,7 +23884,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Term_binderTactic_parenthesizer), return x_1; } } -static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__6() { +static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__7() { _start: { lean_object* x_1; @@ -23669,67 +23892,67 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Term_binderDefault_parenthesizer) return x_1; } } -static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__7() { +static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__5; -x_2 = l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__6; +x_1 = l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__6; +x_2 = l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__7; 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_Command_structExplicitBinder_parenthesizer___closed__8() { +static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__9() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__7; +x_1 = l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__8; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_optional_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__9() { +static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_abbrev_parenthesizer___closed__3; -x_2 = l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__8; +x_2 = l_Lean_Parser_Command_structExplicitBinder_parenthesizer___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); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__10() { +static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__4; -x_2 = l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__9; +x_1 = l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__5; +x_2 = l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__10; 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; } } -static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__11() { +static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__12() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__10; +x_1 = l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__11; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_withoutPosition_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__12() { +static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__11; +x_1 = l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__12; x_2 = l_Lean_Parser_Term_quot_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); @@ -23737,25 +23960,25 @@ lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__13() { +static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__3; -x_2 = l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__12; +x_1 = l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__4; +x_2 = l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__13; 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; } } -static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__14() { +static lean_object* _init_l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_structExplicitBinder___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__13; +x_3 = l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__14; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -23768,7 +23991,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_structExplicitBinder_parenthesize { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__14; +x_7 = l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__15; 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; } @@ -23860,7 +24083,7 @@ static lean_object* _init_l_Lean_Parser_Command_structImplicitBinder_parenthesiz _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__4; +x_1 = l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__5; x_2 = l___regBuiltin_Lean_Parser_Command_declSig_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); @@ -24145,7 +24368,7 @@ static lean_object* _init_l_Lean_Parser_Command_structSimpleBinder_parenthesizer { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_Lean_Parser_Command_optDeclSig_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__8; +x_2 = l_Lean_Parser_Command_structExplicitBinder_parenthesizer___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); @@ -24449,20 +24672,22 @@ return x_2; static lean_object* _init_l_Lean_Parser_Command_structure_parenthesizer___closed__6() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_structure___closed__8; -x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_parenthesizer), 6, 1); -lean_closure_set(x_2, 0, x_1); -return x_2; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Command_structure_parenthesizer___closed__5; +x_2 = l_Lean_Parser_Command_optDeclSig_parenthesizer___closed__8; +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; } } static lean_object* _init_l_Lean_Parser_Command_structure_parenthesizer___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_namedPrio_parenthesizer___closed__6; +x_1 = l_Lean_Parser_Command_structure_parenthesizer___closed__4; x_2 = l_Lean_Parser_Command_structure_parenthesizer___closed__6; -x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_orelse_parenthesizer), 7, 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; @@ -24472,8 +24697,8 @@ static lean_object* _init_l_Lean_Parser_Command_structure_parenthesizer___closed _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___regBuiltin_Lean_Parser_Command_structCtor_parenthesizer___closed__2; -x_2 = lean_alloc_closure((void*)(l_Lean_Parser_optional_parenthesizer), 6, 1); +x_1 = l_Lean_Parser_Command_structure_parenthesizer___closed__7; +x_2 = lean_alloc_closure((void*)(l_Lean_Parser_ppIndent_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } @@ -24481,22 +24706,20 @@ return x_2; static lean_object* _init_l_Lean_Parser_Command_structure_parenthesizer___closed__9() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_structure_parenthesizer___closed__8; -x_2 = l___regBuiltin_Lean_Parser_Command_structFields_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_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Command_structure___closed__10; +x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_parenthesizer), 6, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; } } static lean_object* _init_l_Lean_Parser_Command_structure_parenthesizer___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_structure_parenthesizer___closed__7; +x_1 = l_Lean_Parser_Command_namedPrio_parenthesizer___closed__6; x_2 = l_Lean_Parser_Command_structure_parenthesizer___closed__9; -x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 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; @@ -24506,7 +24729,7 @@ static lean_object* _init_l_Lean_Parser_Command_structure_parenthesizer___closed _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_structure_parenthesizer___closed__10; +x_1 = l___regBuiltin_Lean_Parser_Command_structCtor_parenthesizer___closed__2; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_optional_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; @@ -24517,7 +24740,7 @@ static lean_object* _init_l_Lean_Parser_Command_structure_parenthesizer___closed { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_structure_parenthesizer___closed__11; -x_2 = l___regBuiltin_Lean_Parser_Command_optDeriving_parenthesizer___closed__2; +x_2 = l___regBuiltin_Lean_Parser_Command_structFields_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); @@ -24528,7 +24751,7 @@ static lean_object* _init_l_Lean_Parser_Command_structure_parenthesizer___closed _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_optDeclSig_parenthesizer___closed__8; +x_1 = l_Lean_Parser_Command_structure_parenthesizer___closed__10; x_2 = l_Lean_Parser_Command_structure_parenthesizer___closed__12; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -24539,58 +24762,68 @@ return x_3; static lean_object* _init_l_Lean_Parser_Command_structure_parenthesizer___closed__14() { _start: { +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Command_structure_parenthesizer___closed__13; +x_2 = lean_alloc_closure((void*)(l_Lean_Parser_optional_parenthesizer), 6, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Command_structure_parenthesizer___closed__15() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_structure_parenthesizer___closed__5; -x_2 = l_Lean_Parser_Command_structure_parenthesizer___closed__13; +x_1 = l_Lean_Parser_Command_structure_parenthesizer___closed__14; +x_2 = l___regBuiltin_Lean_Parser_Command_optDeriving_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; } } -static lean_object* _init_l_Lean_Parser_Command_structure_parenthesizer___closed__15() { +static lean_object* _init_l_Lean_Parser_Command_structure_parenthesizer___closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_structure_parenthesizer___closed__4; -x_2 = l_Lean_Parser_Command_structure_parenthesizer___closed__14; +x_1 = l_Lean_Parser_Command_structure_parenthesizer___closed__8; +x_2 = l_Lean_Parser_Command_structure_parenthesizer___closed__15; 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; } } -static lean_object* _init_l_Lean_Parser_Command_structure_parenthesizer___closed__16() { +static lean_object* _init_l_Lean_Parser_Command_structure_parenthesizer___closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_Lean_Parser_Command_declId_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Command_structure_parenthesizer___closed__15; +x_2 = l_Lean_Parser_Command_structure_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); lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_structure_parenthesizer___closed__17() { +static lean_object* _init_l_Lean_Parser_Command_structure_parenthesizer___closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_structure_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Command_structure_parenthesizer___closed__16; +x_2 = l_Lean_Parser_Command_structure_parenthesizer___closed__17; 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; } } -static lean_object* _init_l_Lean_Parser_Command_structure_parenthesizer___closed__18() { +static lean_object* _init_l_Lean_Parser_Command_structure_parenthesizer___closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_structure___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_structure_parenthesizer___closed__17; +x_3 = l_Lean_Parser_Command_structure_parenthesizer___closed__18; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -24603,7 +24836,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_structure_parenthesizer(lean_obje { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Command_structure_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Command_structure_parenthesizer___closed__18; +x_7 = l_Lean_Parser_Command_structure_parenthesizer___closed__19; 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; } @@ -25009,7 +25242,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(192u); +x_1 = lean_unsigned_to_nat(191u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -25021,7 +25254,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(193u); +x_1 = lean_unsigned_to_nat(192u); x_2 = lean_unsigned_to_nat(79u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -25049,7 +25282,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(192u); +x_1 = lean_unsigned_to_nat(191u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -25061,7 +25294,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(192u); +x_1 = lean_unsigned_to_nat(191u); x_2 = lean_unsigned_to_nat(40u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -25429,7 +25662,7 @@ static lean_object* _init_l_Lean_Parser_Command_noncomputableSection___closed__4 _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("section ", 8); +x_1 = lean_mk_string_from_bytes("section", 7); return x_1; } } @@ -25445,59 +25678,69 @@ return x_2; static lean_object* _init_l_Lean_Parser_Command_noncomputableSection___closed__6() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_skip; +x_2 = l_Lean_Parser_ident; +x_3 = l_Lean_Parser_andthen(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Command_noncomputableSection___closed__7() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_ident; +x_1 = l_Lean_Parser_Command_noncomputableSection___closed__6; x_2 = l_Lean_Parser_optional(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_noncomputableSection___closed__7() { +static lean_object* _init_l_Lean_Parser_Command_noncomputableSection___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_noncomputableSection___closed__5; -x_2 = l_Lean_Parser_Command_noncomputableSection___closed__6; +x_2 = l_Lean_Parser_Command_noncomputableSection___closed__7; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_noncomputableSection___closed__8() { +static lean_object* _init_l_Lean_Parser_Command_noncomputableSection___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_noncomputable___closed__5; -x_2 = l_Lean_Parser_Command_noncomputableSection___closed__7; +x_2 = l_Lean_Parser_Command_noncomputableSection___closed__8; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_noncomputableSection___closed__9() { +static lean_object* _init_l_Lean_Parser_Command_noncomputableSection___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_noncomputableSection___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_noncomputableSection___closed__8; +x_3 = l_Lean_Parser_Command_noncomputableSection___closed__9; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Command_noncomputableSection___closed__10() { +static lean_object* _init_l_Lean_Parser_Command_noncomputableSection___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_noncomputableSection___closed__3; -x_2 = l_Lean_Parser_Command_noncomputableSection___closed__9; +x_2 = l_Lean_Parser_Command_noncomputableSection___closed__10; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_noncomputableSection___closed__11() { +static lean_object* _init_l_Lean_Parser_Command_noncomputableSection___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_noncomputableSection___closed__2; -x_2 = l_Lean_Parser_Command_noncomputableSection___closed__10; +x_2 = l_Lean_Parser_Command_noncomputableSection___closed__11; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -25506,7 +25749,7 @@ static lean_object* _init_l_Lean_Parser_Command_noncomputableSection() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Command_noncomputableSection___closed__11; +x_1 = l_Lean_Parser_Command_noncomputableSection___closed__12; return x_1; } } @@ -25527,7 +25770,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(194u); +x_1 = lean_unsigned_to_nat(193u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -25539,8 +25782,8 @@ 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(195u); -x_2 = lean_unsigned_to_nat(50u); +x_1 = lean_unsigned_to_nat(194u); +x_2 = lean_unsigned_to_nat(62u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -25554,7 +25797,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Parser_Command_noncomputableSection_declRange___closed__1; x_2 = lean_unsigned_to_nat(26u); x_3 = l___regBuiltin_Lean_Parser_Command_noncomputableSection_declRange___closed__2; -x_4 = lean_unsigned_to_nat(50u); +x_4 = lean_unsigned_to_nat(62u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -25567,7 +25810,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(194u); +x_1 = lean_unsigned_to_nat(193u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -25579,7 +25822,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(194u); +x_1 = lean_unsigned_to_nat(193u); x_2 = lean_unsigned_to_nat(50u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -25656,44 +25899,56 @@ return x_2; static lean_object* _init_l_Lean_Parser_Command_noncomputableSection_formatter___closed__3() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Command_optDeclSig_formatter___closed__5; +x_2 = l_Lean_Parser_Command_declId_formatter___closed__4; +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; +} +} +static lean_object* _init_l_Lean_Parser_Command_noncomputableSection_formatter___closed__4() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_declId_formatter___closed__4; +x_1 = l_Lean_Parser_Command_noncomputableSection_formatter___closed__3; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_optional_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_noncomputableSection_formatter___closed__4() { +static lean_object* _init_l_Lean_Parser_Command_noncomputableSection_formatter___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_noncomputableSection_formatter___closed__2; -x_2 = l_Lean_Parser_Command_noncomputableSection_formatter___closed__3; +x_2 = l_Lean_Parser_Command_noncomputableSection_formatter___closed__4; 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; } } -static lean_object* _init_l_Lean_Parser_Command_noncomputableSection_formatter___closed__5() { +static lean_object* _init_l_Lean_Parser_Command_noncomputableSection_formatter___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_noncomputable_formatter___closed__2; -x_2 = l_Lean_Parser_Command_noncomputableSection_formatter___closed__4; +x_2 = l_Lean_Parser_Command_noncomputableSection_formatter___closed__5; 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; } } -static lean_object* _init_l_Lean_Parser_Command_noncomputableSection_formatter___closed__6() { +static lean_object* _init_l_Lean_Parser_Command_noncomputableSection_formatter___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_noncomputableSection___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_noncomputableSection_formatter___closed__5; +x_3 = l_Lean_Parser_Command_noncomputableSection_formatter___closed__6; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -25706,7 +25961,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_noncomputableSection_formatter(le { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Command_noncomputableSection_formatter___closed__1; -x_7 = l_Lean_Parser_Command_noncomputableSection_formatter___closed__6; +x_7 = l_Lean_Parser_Command_noncomputableSection_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; } @@ -25775,44 +26030,56 @@ return x_2; static lean_object* _init_l_Lean_Parser_Command_noncomputableSection_parenthesizer___closed__3() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Command_optDeclSig_parenthesizer___closed__5; +x_2 = l_Lean_Parser_Command_declId_parenthesizer___closed__4; +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; +} +} +static lean_object* _init_l_Lean_Parser_Command_noncomputableSection_parenthesizer___closed__4() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_declId_parenthesizer___closed__4; +x_1 = l_Lean_Parser_Command_noncomputableSection_parenthesizer___closed__3; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_optional_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_noncomputableSection_parenthesizer___closed__4() { +static lean_object* _init_l_Lean_Parser_Command_noncomputableSection_parenthesizer___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_noncomputableSection_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Command_noncomputableSection_parenthesizer___closed__3; +x_2 = l_Lean_Parser_Command_noncomputableSection_parenthesizer___closed__4; 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; } } -static lean_object* _init_l_Lean_Parser_Command_noncomputableSection_parenthesizer___closed__5() { +static lean_object* _init_l_Lean_Parser_Command_noncomputableSection_parenthesizer___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_noncomputable_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Command_noncomputableSection_parenthesizer___closed__4; +x_2 = l_Lean_Parser_Command_noncomputableSection_parenthesizer___closed__5; 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; } } -static lean_object* _init_l_Lean_Parser_Command_noncomputableSection_parenthesizer___closed__6() { +static lean_object* _init_l_Lean_Parser_Command_noncomputableSection_parenthesizer___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_noncomputableSection___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_noncomputableSection_parenthesizer___closed__5; +x_3 = l_Lean_Parser_Command_noncomputableSection_parenthesizer___closed__6; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -25825,7 +26092,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_noncomputableSection_parenthesize { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Command_noncomputableSection_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Command_noncomputableSection_parenthesizer___closed__6; +x_7 = l_Lean_Parser_Command_noncomputableSection_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; } @@ -25866,62 +26133,54 @@ return x_6; static lean_object* _init_l_Lean_Parser_Command_section___closed__1() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("section", 7); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Command_section___closed__2() { -_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_quot___closed__1; x_2 = l_Lean_Parser_Term_quot___closed__2; x_3 = l_Lean_Parser_Command_quot___closed__1; -x_4 = l_Lean_Parser_Command_section___closed__1; +x_4 = l_Lean_Parser_Command_noncomputableSection___closed__4; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Parser_Command_section___closed__3() { +static lean_object* _init_l_Lean_Parser_Command_section___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_Command_section___closed__1; -x_2 = l_Lean_Parser_Command_section___closed__2; +x_1 = l_Lean_Parser_Command_noncomputableSection___closed__4; +x_2 = l_Lean_Parser_Command_section___closed__1; 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_Command_section___closed__4() { +static lean_object* _init_l_Lean_Parser_Command_section___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_section___closed__2; +x_1 = l_Lean_Parser_Command_section___closed__1; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_noncomputableSection___closed__7; +x_3 = l_Lean_Parser_Command_noncomputableSection___closed__8; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Command_section___closed__5() { +static lean_object* _init_l_Lean_Parser_Command_section___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_section___closed__3; -x_2 = l_Lean_Parser_Command_section___closed__4; +x_1 = l_Lean_Parser_Command_section___closed__2; +x_2 = l_Lean_Parser_Command_section___closed__3; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_section___closed__6() { +static lean_object* _init_l_Lean_Parser_Command_section___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_section___closed__2; -x_2 = l_Lean_Parser_Command_section___closed__5; +x_1 = l_Lean_Parser_Command_section___closed__1; +x_2 = l_Lean_Parser_Command_section___closed__4; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -25930,7 +26189,7 @@ static lean_object* _init_l_Lean_Parser_Command_section() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Command_section___closed__6; +x_1 = l_Lean_Parser_Command_section___closed__5; return x_1; } } @@ -25939,7 +26198,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_section(lean_object* { 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_Lean_Parser_Command_quot___closed__5; -x_3 = l_Lean_Parser_Command_section___closed__2; +x_3 = l_Lean_Parser_Command_section___closed__1; x_4 = 1; x_5 = l_Lean_Parser_Command_section; x_6 = lean_unsigned_to_nat(1000u); @@ -25951,7 +26210,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(196u); +x_1 = lean_unsigned_to_nat(195u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -25963,8 +26222,8 @@ 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(197u); -x_2 = lean_unsigned_to_nat(30u); +x_1 = lean_unsigned_to_nat(196u); +x_2 = lean_unsigned_to_nat(42u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -25978,7 +26237,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Parser_Command_section_declRange___closed__1; x_2 = lean_unsigned_to_nat(26u); x_3 = l___regBuiltin_Lean_Parser_Command_section_declRange___closed__2; -x_4 = lean_unsigned_to_nat(30u); +x_4 = lean_unsigned_to_nat(42u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -25991,7 +26250,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(196u); +x_1 = lean_unsigned_to_nat(195u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -26003,7 +26262,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(196u); +x_1 = lean_unsigned_to_nat(195u); x_2 = lean_unsigned_to_nat(39u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -26043,7 +26302,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_section_declRange(le _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l_Lean_Parser_Command_section___closed__2; +x_2 = l_Lean_Parser_Command_section___closed__1; x_3 = l___regBuiltin_Lean_Parser_Command_section_declRange___closed__7; x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; @@ -26053,8 +26312,8 @@ static lean_object* _init_l_Lean_Parser_Command_section_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_Command_section___closed__1; -x_2 = l_Lean_Parser_Command_section___closed__2; +x_1 = l_Lean_Parser_Command_noncomputableSection___closed__4; +x_2 = l_Lean_Parser_Command_section___closed__1; x_3 = 1; x_4 = 0; x_5 = lean_box(x_3); @@ -26071,9 +26330,9 @@ static lean_object* _init_l_Lean_Parser_Command_section_formatter___closed__2() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Command_section___closed__2; +x_1 = l_Lean_Parser_Command_section___closed__1; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_noncomputableSection_formatter___closed__4; +x_3 = l_Lean_Parser_Command_noncomputableSection_formatter___closed__5; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -26098,7 +26357,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_quot___closed__1; x_2 = l_Lean_Parser_Term_quot___closed__2; x_3 = l_Lean_Parser_Command_quot___closed__1; -x_4 = l_Lean_Parser_Command_section___closed__1; +x_4 = l_Lean_Parser_Command_noncomputableSection___closed__4; x_5 = l___regBuiltin_Lean_Parser_Term_quot_formatter___closed__1; x_6 = l_Lean_Name_mkStr5(x_1, x_2, x_3, x_4, x_5); return x_6; @@ -26117,7 +26376,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_section_formatter(le { 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_Term_quot_formatter___closed__3; -x_3 = l_Lean_Parser_Command_section___closed__2; +x_3 = l_Lean_Parser_Command_section___closed__1; x_4 = l___regBuiltin_Lean_Parser_Command_section_formatter___closed__1; x_5 = l___regBuiltin_Lean_Parser_Command_section_formatter___closed__2; x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); @@ -26128,8 +26387,8 @@ static lean_object* _init_l_Lean_Parser_Command_section_parenthesizer___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_Command_section___closed__1; -x_2 = l_Lean_Parser_Command_section___closed__2; +x_1 = l_Lean_Parser_Command_noncomputableSection___closed__4; +x_2 = l_Lean_Parser_Command_section___closed__1; x_3 = 1; x_4 = 0; x_5 = lean_box(x_3); @@ -26146,9 +26405,9 @@ static lean_object* _init_l_Lean_Parser_Command_section_parenthesizer___closed__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Command_section___closed__2; +x_1 = l_Lean_Parser_Command_section___closed__1; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_noncomputableSection_parenthesizer___closed__4; +x_3 = l_Lean_Parser_Command_noncomputableSection_parenthesizer___closed__5; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -26173,7 +26432,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_quot___closed__1; x_2 = l_Lean_Parser_Term_quot___closed__2; x_3 = l_Lean_Parser_Command_quot___closed__1; -x_4 = l_Lean_Parser_Command_section___closed__1; +x_4 = l_Lean_Parser_Command_noncomputableSection___closed__4; x_5 = l___regBuiltin_Lean_Parser_Term_quot_parenthesizer___closed__1; x_6 = l_Lean_Name_mkStr5(x_1, x_2, x_3, x_4, x_5); return x_6; @@ -26192,7 +26451,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_section_parenthesize { 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_Term_quot_parenthesizer___closed__3; -x_3 = l_Lean_Parser_Command_section___closed__2; +x_3 = l_Lean_Parser_Command_section___closed__1; x_4 = l___regBuiltin_Lean_Parser_Command_section_parenthesizer___closed__1; x_5 = l___regBuiltin_Lean_Parser_Command_section_parenthesizer___closed__2; x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); @@ -26314,7 +26573,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(198u); +x_1 = lean_unsigned_to_nat(197u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -26326,7 +26585,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(199u); +x_1 = lean_unsigned_to_nat(198u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -26354,7 +26613,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(198u); +x_1 = lean_unsigned_to_nat(197u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -26366,7 +26625,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(198u); +x_1 = lean_unsigned_to_nat(197u); x_2 = lean_unsigned_to_nat(41u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -26641,57 +26900,49 @@ return x_5; static lean_object* _init_l_Lean_Parser_Command_end___closed__4() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("end ", 4); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Command_end___closed__5() { -_start: -{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_end___closed__4; +x_1 = l_Lean_Parser_Command_end___closed__1; x_2 = l_Lean_Parser_symbol(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_end___closed__6() { +static lean_object* _init_l_Lean_Parser_Command_end___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_end___closed__5; -x_2 = l_Lean_Parser_Command_noncomputableSection___closed__6; +x_1 = l_Lean_Parser_Command_end___closed__4; +x_2 = l_Lean_Parser_Command_noncomputableSection___closed__7; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_end___closed__7() { +static lean_object* _init_l_Lean_Parser_Command_end___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_end___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_end___closed__6; +x_3 = l_Lean_Parser_Command_end___closed__5; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Command_end___closed__8() { +static lean_object* _init_l_Lean_Parser_Command_end___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_end___closed__3; -x_2 = l_Lean_Parser_Command_end___closed__7; +x_2 = l_Lean_Parser_Command_end___closed__6; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_end___closed__9() { +static lean_object* _init_l_Lean_Parser_Command_end___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_end___closed__2; -x_2 = l_Lean_Parser_Command_end___closed__8; +x_2 = l_Lean_Parser_Command_end___closed__7; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -26700,7 +26951,7 @@ static lean_object* _init_l_Lean_Parser_Command_end() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Command_end___closed__9; +x_1 = l_Lean_Parser_Command_end___closed__8; return x_1; } } @@ -26721,7 +26972,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(200u); +x_1 = lean_unsigned_to_nat(199u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -26733,8 +26984,8 @@ 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(201u); -x_2 = lean_unsigned_to_nat(26u); +x_1 = lean_unsigned_to_nat(200u); +x_2 = lean_unsigned_to_nat(38u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -26744,23 +26995,24 @@ return x_3; static lean_object* _init_l___regBuiltin_Lean_Parser_Command_end_declRange___closed__3() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +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_Parser_Command_end_declRange___closed__1; x_2 = lean_unsigned_to_nat(26u); x_3 = l___regBuiltin_Lean_Parser_Command_end_declRange___closed__2; -x_4 = lean_alloc_ctor(0, 4, 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); -lean_ctor_set(x_4, 3, x_2); -return x_4; +x_4 = lean_unsigned_to_nat(38u); +x_5 = lean_alloc_ctor(0, 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; } } static lean_object* _init_l___regBuiltin_Lean_Parser_Command_end_declRange___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(200u); +x_1 = lean_unsigned_to_nat(199u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -26772,7 +27024,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(200u); +x_1 = lean_unsigned_to_nat(199u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -26840,7 +27092,7 @@ static lean_object* _init_l_Lean_Parser_Command_end_formatter___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_end___closed__4; +x_1 = l_Lean_Parser_Command_end___closed__1; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; @@ -26851,7 +27103,7 @@ static lean_object* _init_l_Lean_Parser_Command_end_formatter___closed__3() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_end_formatter___closed__2; -x_2 = l_Lean_Parser_Command_noncomputableSection_formatter___closed__3; +x_2 = l_Lean_Parser_Command_noncomputableSection_formatter___closed__4; 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); @@ -26937,7 +27189,7 @@ static lean_object* _init_l_Lean_Parser_Command_end_parenthesizer___closed__2() _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_end___closed__4; +x_1 = l_Lean_Parser_Command_end___closed__1; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; @@ -26948,7 +27200,7 @@ static lean_object* _init_l_Lean_Parser_Command_end_parenthesizer___closed__3() { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_end_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Command_noncomputableSection_parenthesizer___closed__3; +x_2 = l_Lean_Parser_Command_noncomputableSection_parenthesizer___closed__4; 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); @@ -27128,7 +27380,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(202u); +x_1 = lean_unsigned_to_nat(201u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -27140,7 +27392,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(203u); +x_1 = lean_unsigned_to_nat(202u); x_2 = lean_unsigned_to_nat(55u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -27168,7 +27420,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(202u); +x_1 = lean_unsigned_to_nat(201u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -27180,7 +27432,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(202u); +x_1 = lean_unsigned_to_nat(201u); x_2 = lean_unsigned_to_nat(40u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -27475,17 +27727,18 @@ return x_5; static lean_object* _init_l_Lean_Parser_Command_universe___closed__4() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("universe ", 9); -return x_1; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Command_universe___closed__1; +x_2 = l_Lean_Parser_symbol(x_1); +return x_2; } } static lean_object* _init_l_Lean_Parser_Command_universe___closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_universe___closed__4; -x_2 = l_Lean_Parser_symbol(x_1); +x_1 = l_Lean_Parser_Command_noncomputableSection___closed__6; +x_2 = l_Lean_Parser_many1(x_1); return x_2; } } @@ -27493,8 +27746,8 @@ static lean_object* _init_l_Lean_Parser_Command_universe___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_universe___closed__5; -x_2 = l_Lean_Parser_Command_structExplicitBinder___closed__6; +x_1 = l_Lean_Parser_Command_universe___closed__4; +x_2 = l_Lean_Parser_Command_universe___closed__5; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -27555,7 +27808,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(204u); +x_1 = lean_unsigned_to_nat(203u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -27567,8 +27820,8 @@ 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(205u); -x_2 = lean_unsigned_to_nat(28u); +x_1 = lean_unsigned_to_nat(204u); +x_2 = lean_unsigned_to_nat(40u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -27582,7 +27835,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Parser_Command_universe_declRange___closed__1; x_2 = lean_unsigned_to_nat(26u); x_3 = l___regBuiltin_Lean_Parser_Command_universe_declRange___closed__2; -x_4 = lean_unsigned_to_nat(28u); +x_4 = lean_unsigned_to_nat(40u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -27595,7 +27848,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(204u); +x_1 = lean_unsigned_to_nat(203u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -27607,7 +27860,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(204u); +x_1 = lean_unsigned_to_nat(203u); x_2 = lean_unsigned_to_nat(40u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -27675,7 +27928,7 @@ static lean_object* _init_l_Lean_Parser_Command_universe_formatter___closed__2() _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_universe___closed__4; +x_1 = l_Lean_Parser_Command_universe___closed__1; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; @@ -27684,22 +27937,32 @@ return x_2; static lean_object* _init_l_Lean_Parser_Command_universe_formatter___closed__3() { _start: { +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Command_noncomputableSection_formatter___closed__3; +x_2 = lean_alloc_closure((void*)(l_Lean_Parser_many1_formatter), 6, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Command_universe_formatter___closed__4() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_universe_formatter___closed__2; -x_2 = l_Lean_Parser_Command_structExplicitBinder_formatter___closed__4; +x_2 = l_Lean_Parser_Command_universe_formatter___closed__3; 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; } } -static lean_object* _init_l_Lean_Parser_Command_universe_formatter___closed__4() { +static lean_object* _init_l_Lean_Parser_Command_universe_formatter___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_universe___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_universe_formatter___closed__3; +x_3 = l_Lean_Parser_Command_universe_formatter___closed__4; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -27712,7 +27975,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_universe_formatter(lean_object* x { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Command_universe_formatter___closed__1; -x_7 = l_Lean_Parser_Command_universe_formatter___closed__4; +x_7 = l_Lean_Parser_Command_universe_formatter___closed__5; 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; } @@ -27772,7 +28035,7 @@ static lean_object* _init_l_Lean_Parser_Command_universe_parenthesizer___closed_ _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_universe___closed__4; +x_1 = l_Lean_Parser_Command_universe___closed__1; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; @@ -27781,22 +28044,32 @@ return x_2; static lean_object* _init_l_Lean_Parser_Command_universe_parenthesizer___closed__3() { _start: { +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Command_noncomputableSection_parenthesizer___closed__3; +x_2 = lean_alloc_closure((void*)(l_Lean_Parser_many1_parenthesizer), 6, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Command_universe_parenthesizer___closed__4() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_universe_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__4; +x_2 = l_Lean_Parser_Command_universe_parenthesizer___closed__3; 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; } } -static lean_object* _init_l_Lean_Parser_Command_universe_parenthesizer___closed__4() { +static lean_object* _init_l_Lean_Parser_Command_universe_parenthesizer___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_universe___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_universe_parenthesizer___closed__3; +x_3 = l_Lean_Parser_Command_universe_parenthesizer___closed__4; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -27809,7 +28082,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_universe_parenthesizer(lean_objec { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Command_universe_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Command_universe_parenthesizer___closed__4; +x_7 = l_Lean_Parser_Command_universe_parenthesizer___closed__5; 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; } @@ -27962,7 +28235,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(206u); +x_1 = lean_unsigned_to_nat(205u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -27974,7 +28247,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(207u); +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); @@ -28002,7 +28275,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(206u); +x_1 = lean_unsigned_to_nat(205u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -28014,7 +28287,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(206u); +x_1 = lean_unsigned_to_nat(205u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -28369,7 +28642,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(208u); +x_1 = lean_unsigned_to_nat(207u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -28381,7 +28654,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(209u); +x_1 = lean_unsigned_to_nat(208u); x_2 = lean_unsigned_to_nat(33u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -28409,7 +28682,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(208u); +x_1 = lean_unsigned_to_nat(207u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -28421,7 +28694,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(208u); +x_1 = lean_unsigned_to_nat(207u); x_2 = lean_unsigned_to_nat(43u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -28776,7 +29049,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(210u); +x_1 = lean_unsigned_to_nat(209u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -28788,7 +29061,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(211u); +x_1 = lean_unsigned_to_nat(210u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -28815,7 +29088,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(210u); +x_1 = lean_unsigned_to_nat(209u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -28827,7 +29100,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(210u); +x_1 = lean_unsigned_to_nat(209u); x_2 = lean_unsigned_to_nat(36u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29182,7 +29455,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(212u); +x_1 = lean_unsigned_to_nat(211u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29194,7 +29467,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(213u); +x_1 = lean_unsigned_to_nat(212u); x_2 = lean_unsigned_to_nat(24u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29222,7 +29495,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(212u); +x_1 = lean_unsigned_to_nat(211u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29234,7 +29507,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(212u); +x_1 = lean_unsigned_to_nat(211u); x_2 = lean_unsigned_to_nat(34u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29589,7 +29862,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(214u); +x_1 = lean_unsigned_to_nat(213u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29601,7 +29874,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(215u); +x_1 = lean_unsigned_to_nat(214u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29629,7 +29902,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(214u); +x_1 = lean_unsigned_to_nat(213u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29641,7 +29914,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(214u); +x_1 = lean_unsigned_to_nat(213u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29986,7 +30259,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(216u); +x_1 = lean_unsigned_to_nat(215u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29998,7 +30271,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(217u); +x_1 = lean_unsigned_to_nat(216u); x_2 = lean_unsigned_to_nat(9u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30026,7 +30299,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(216u); +x_1 = lean_unsigned_to_nat(215u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30038,7 +30311,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(216u); +x_1 = lean_unsigned_to_nat(215u); x_2 = lean_unsigned_to_nat(34u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30379,7 +30652,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(218u); +x_1 = lean_unsigned_to_nat(217u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30391,7 +30664,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(219u); +x_1 = lean_unsigned_to_nat(218u); x_2 = lean_unsigned_to_nat(33u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30419,7 +30692,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(218u); +x_1 = lean_unsigned_to_nat(217u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30431,7 +30704,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(218u); +x_1 = lean_unsigned_to_nat(217u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30837,7 +31110,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(220u); +x_1 = lean_unsigned_to_nat(219u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30849,7 +31122,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(221u); +x_1 = lean_unsigned_to_nat(220u); x_2 = lean_unsigned_to_nat(51u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30877,7 +31150,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(220u); +x_1 = lean_unsigned_to_nat(219u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30889,7 +31162,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(220u); +x_1 = lean_unsigned_to_nat(219u); x_2 = lean_unsigned_to_nat(41u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -31256,7 +31529,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(222u); +x_1 = lean_unsigned_to_nat(221u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -31268,7 +31541,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(223u); +x_1 = lean_unsigned_to_nat(222u); x_2 = lean_unsigned_to_nat(13u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -31296,7 +31569,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(222u); +x_1 = lean_unsigned_to_nat(221u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -31308,7 +31581,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(222u); +x_1 = lean_unsigned_to_nat(221u); x_2 = lean_unsigned_to_nat(41u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -31733,7 +32006,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(225u); +x_1 = lean_unsigned_to_nat(224u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -31745,7 +32018,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(226u); +x_1 = lean_unsigned_to_nat(225u); x_2 = lean_unsigned_to_nat(50u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -31773,7 +32046,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(225u); +x_1 = lean_unsigned_to_nat(224u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -31785,7 +32058,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(225u); +x_1 = lean_unsigned_to_nat(224u); x_2 = lean_unsigned_to_nat(42u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -32407,87 +32680,70 @@ return x_3; static lean_object* _init_l_Lean_Parser_Command_attribute___closed__9() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("] ", 2); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Command_attribute___closed__10() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_attribute___closed__9; -x_2 = l_Lean_Parser_symbol(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Parser_Command_attribute___closed__11() { -_start: -{ lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_attribute___closed__10; -x_2 = l_Lean_Parser_Command_structExplicitBinder___closed__6; +x_1 = l_Lean_Parser_Command_structInstBinder___closed__9; +x_2 = l_Lean_Parser_Command_universe___closed__5; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_attribute___closed__12() { +static lean_object* _init_l_Lean_Parser_Command_attribute___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_attribute___closed__8; -x_2 = l_Lean_Parser_Command_attribute___closed__11; +x_2 = l_Lean_Parser_Command_attribute___closed__9; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_attribute___closed__13() { +static lean_object* _init_l_Lean_Parser_Command_attribute___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_structInstBinder___closed__5; -x_2 = l_Lean_Parser_Command_attribute___closed__12; +x_2 = l_Lean_Parser_Command_attribute___closed__10; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_attribute___closed__14() { +static lean_object* _init_l_Lean_Parser_Command_attribute___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_attribute___closed__5; -x_2 = l_Lean_Parser_Command_attribute___closed__13; +x_2 = l_Lean_Parser_Command_attribute___closed__11; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_attribute___closed__15() { +static lean_object* _init_l_Lean_Parser_Command_attribute___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_attribute___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_attribute___closed__14; +x_3 = l_Lean_Parser_Command_attribute___closed__12; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Command_attribute___closed__16() { +static lean_object* _init_l_Lean_Parser_Command_attribute___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_attribute___closed__3; -x_2 = l_Lean_Parser_Command_attribute___closed__15; +x_2 = l_Lean_Parser_Command_attribute___closed__13; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_attribute___closed__17() { +static lean_object* _init_l_Lean_Parser_Command_attribute___closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_attribute___closed__2; -x_2 = l_Lean_Parser_Command_attribute___closed__16; +x_2 = l_Lean_Parser_Command_attribute___closed__14; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -32496,7 +32752,7 @@ static lean_object* _init_l_Lean_Parser_Command_attribute() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Command_attribute___closed__17; +x_1 = l_Lean_Parser_Command_attribute___closed__15; return x_1; } } @@ -32517,7 +32773,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(229u); +x_1 = lean_unsigned_to_nat(228u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -32529,8 +32785,8 @@ 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(232u); -x_2 = lean_unsigned_to_nat(21u); +x_1 = lean_unsigned_to_nat(231u); +x_2 = lean_unsigned_to_nat(33u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -32544,7 +32800,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Parser_Command_attribute_declRange___closed__1; x_2 = lean_unsigned_to_nat(26u); x_3 = l___regBuiltin_Lean_Parser_Command_attribute_declRange___closed__2; -x_4 = lean_unsigned_to_nat(21u); +x_4 = lean_unsigned_to_nat(33u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -32557,7 +32813,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(229u); +x_1 = lean_unsigned_to_nat(228u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -32569,7 +32825,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(229u); +x_1 = lean_unsigned_to_nat(228u); x_2 = lean_unsigned_to_nat(41u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -32790,68 +33046,58 @@ return x_2; static lean_object* _init_l_Lean_Parser_Command_attribute_formatter___closed__7() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_attribute___closed__9; -x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_formatter), 6, 1); -lean_closure_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Parser_Command_attribute_formatter___closed__8() { -_start: -{ lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_attribute_formatter___closed__7; -x_2 = l_Lean_Parser_Command_structExplicitBinder_formatter___closed__4; +x_1 = l_Lean_Parser_Command_structInstBinder_formatter___closed__5; +x_2 = l_Lean_Parser_Command_universe_formatter___closed__3; 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; } } -static lean_object* _init_l_Lean_Parser_Command_attribute_formatter___closed__9() { +static lean_object* _init_l_Lean_Parser_Command_attribute_formatter___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_attribute_formatter___closed__6; -x_2 = l_Lean_Parser_Command_attribute_formatter___closed__8; +x_2 = l_Lean_Parser_Command_attribute_formatter___closed__7; 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; } } -static lean_object* _init_l_Lean_Parser_Command_attribute_formatter___closed__10() { +static lean_object* _init_l_Lean_Parser_Command_attribute_formatter___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_structInstBinder_formatter___closed__2; -x_2 = l_Lean_Parser_Command_attribute_formatter___closed__9; +x_2 = l_Lean_Parser_Command_attribute_formatter___closed__8; 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; } } -static lean_object* _init_l_Lean_Parser_Command_attribute_formatter___closed__11() { +static lean_object* _init_l_Lean_Parser_Command_attribute_formatter___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_attribute_formatter___closed__2; -x_2 = l_Lean_Parser_Command_attribute_formatter___closed__10; +x_2 = l_Lean_Parser_Command_attribute_formatter___closed__9; 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; } } -static lean_object* _init_l_Lean_Parser_Command_attribute_formatter___closed__12() { +static lean_object* _init_l_Lean_Parser_Command_attribute_formatter___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_attribute___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_attribute_formatter___closed__11; +x_3 = l_Lean_Parser_Command_attribute_formatter___closed__10; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -32864,7 +33110,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_attribute_formatter(lean_object* { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Command_attribute_formatter___closed__1; -x_7 = l_Lean_Parser_Command_attribute_formatter___closed__12; +x_7 = l_Lean_Parser_Command_attribute_formatter___closed__11; 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; } @@ -33077,68 +33323,58 @@ return x_2; static lean_object* _init_l_Lean_Parser_Command_attribute_parenthesizer___closed__7() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_attribute___closed__9; -x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_parenthesizer), 6, 1); -lean_closure_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Parser_Command_attribute_parenthesizer___closed__8() { -_start: -{ lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_attribute_parenthesizer___closed__7; -x_2 = l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__4; +x_1 = l_Lean_Parser_Command_structInstBinder_parenthesizer___closed__5; +x_2 = l_Lean_Parser_Command_universe_parenthesizer___closed__3; 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; } } -static lean_object* _init_l_Lean_Parser_Command_attribute_parenthesizer___closed__9() { +static lean_object* _init_l_Lean_Parser_Command_attribute_parenthesizer___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_attribute_parenthesizer___closed__6; -x_2 = l_Lean_Parser_Command_attribute_parenthesizer___closed__8; +x_2 = l_Lean_Parser_Command_attribute_parenthesizer___closed__7; 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; } } -static lean_object* _init_l_Lean_Parser_Command_attribute_parenthesizer___closed__10() { +static lean_object* _init_l_Lean_Parser_Command_attribute_parenthesizer___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_structInstBinder_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Command_attribute_parenthesizer___closed__9; +x_2 = l_Lean_Parser_Command_attribute_parenthesizer___closed__8; 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; } } -static lean_object* _init_l_Lean_Parser_Command_attribute_parenthesizer___closed__11() { +static lean_object* _init_l_Lean_Parser_Command_attribute_parenthesizer___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_attribute_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Command_attribute_parenthesizer___closed__10; +x_2 = l_Lean_Parser_Command_attribute_parenthesizer___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); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_attribute_parenthesizer___closed__12() { +static lean_object* _init_l_Lean_Parser_Command_attribute_parenthesizer___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_attribute___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_attribute_parenthesizer___closed__11; +x_3 = l_Lean_Parser_Command_attribute_parenthesizer___closed__10; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -33151,7 +33387,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_attribute_parenthesizer(lean_obje { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Command_attribute_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Command_attribute_parenthesizer___closed__12; +x_7 = l_Lean_Parser_Command_attribute_parenthesizer___closed__11; 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; } @@ -33241,87 +33477,70 @@ return x_2; static lean_object* _init_l_Lean_Parser_Command_export___closed__6() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes(" (", 2); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Command_export___closed__7() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_export___closed__6; -x_2 = l_Lean_Parser_symbol(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Parser_Command_export___closed__8() { -_start: -{ lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_structExplicitBinder___closed__6; +x_1 = l_Lean_Parser_Command_structExplicitBinder___closed__8; x_2 = l_Lean_Parser_Term_quot___closed__14; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_export___closed__9() { +static lean_object* _init_l_Lean_Parser_Command_export___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_export___closed__7; -x_2 = l_Lean_Parser_Command_export___closed__8; +x_1 = l_Lean_Parser_Command_namedPrio___closed__5; +x_2 = l_Lean_Parser_Command_export___closed__6; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_export___closed__10() { +static lean_object* _init_l_Lean_Parser_Command_export___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_ident; -x_2 = l_Lean_Parser_Command_export___closed__9; +x_2 = l_Lean_Parser_Command_export___closed__7; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_export___closed__11() { +static lean_object* _init_l_Lean_Parser_Command_export___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_export___closed__5; -x_2 = l_Lean_Parser_Command_export___closed__10; +x_2 = l_Lean_Parser_Command_export___closed__8; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_export___closed__12() { +static lean_object* _init_l_Lean_Parser_Command_export___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_export___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_export___closed__11; +x_3 = l_Lean_Parser_Command_export___closed__9; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Command_export___closed__13() { +static lean_object* _init_l_Lean_Parser_Command_export___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_export___closed__3; -x_2 = l_Lean_Parser_Command_export___closed__12; +x_2 = l_Lean_Parser_Command_export___closed__10; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_export___closed__14() { +static lean_object* _init_l_Lean_Parser_Command_export___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_export___closed__2; -x_2 = l_Lean_Parser_Command_export___closed__13; +x_2 = l_Lean_Parser_Command_export___closed__11; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -33330,7 +33549,7 @@ static lean_object* _init_l_Lean_Parser_Command_export() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Command_export___closed__14; +x_1 = l_Lean_Parser_Command_export___closed__12; return x_1; } } @@ -33351,7 +33570,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(233u); +x_1 = lean_unsigned_to_nat(232u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -33363,7 +33582,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(234u); +x_1 = lean_unsigned_to_nat(233u); x_2 = lean_unsigned_to_nat(50u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -33391,7 +33610,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(233u); +x_1 = lean_unsigned_to_nat(232u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -33403,7 +33622,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(233u); +x_1 = lean_unsigned_to_nat(232u); x_2 = lean_unsigned_to_nat(38u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -33480,18 +33699,8 @@ return x_2; static lean_object* _init_l_Lean_Parser_Command_export_formatter___closed__3() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_export___closed__6; -x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_formatter), 6, 1); -lean_closure_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Parser_Command_export_formatter___closed__4() { -_start: -{ lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_structExplicitBinder_formatter___closed__4; +x_1 = l_Lean_Parser_Command_structExplicitBinder_formatter___closed__5; x_2 = l_Lean_Parser_Term_quot_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); @@ -33499,49 +33708,49 @@ lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_export_formatter___closed__5() { +static lean_object* _init_l_Lean_Parser_Command_export_formatter___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_export_formatter___closed__3; -x_2 = l_Lean_Parser_Command_export_formatter___closed__4; +x_1 = l_Lean_Parser_Command_namedPrio_formatter___closed__2; +x_2 = l_Lean_Parser_Command_export_formatter___closed__3; 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; } } -static lean_object* _init_l_Lean_Parser_Command_export_formatter___closed__6() { +static lean_object* _init_l_Lean_Parser_Command_export_formatter___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_declId_formatter___closed__4; -x_2 = l_Lean_Parser_Command_export_formatter___closed__5; +x_2 = l_Lean_Parser_Command_export_formatter___closed__4; 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; } } -static lean_object* _init_l_Lean_Parser_Command_export_formatter___closed__7() { +static lean_object* _init_l_Lean_Parser_Command_export_formatter___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_export_formatter___closed__2; -x_2 = l_Lean_Parser_Command_export_formatter___closed__6; +x_2 = l_Lean_Parser_Command_export_formatter___closed__5; 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; } } -static lean_object* _init_l_Lean_Parser_Command_export_formatter___closed__8() { +static lean_object* _init_l_Lean_Parser_Command_export_formatter___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_export___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_export_formatter___closed__7; +x_3 = l_Lean_Parser_Command_export_formatter___closed__6; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -33554,7 +33763,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_export_formatter(lean_object* x_1 { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Command_export_formatter___closed__1; -x_7 = l_Lean_Parser_Command_export_formatter___closed__8; +x_7 = l_Lean_Parser_Command_export_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; } @@ -33623,18 +33832,8 @@ return x_2; static lean_object* _init_l_Lean_Parser_Command_export_parenthesizer___closed__3() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_export___closed__6; -x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_parenthesizer), 6, 1); -lean_closure_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Parser_Command_export_parenthesizer___closed__4() { -_start: -{ lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__4; +x_1 = l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__5; x_2 = l_Lean_Parser_Term_quot_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); @@ -33642,49 +33841,49 @@ lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_export_parenthesizer___closed__5() { +static lean_object* _init_l_Lean_Parser_Command_export_parenthesizer___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_export_parenthesizer___closed__3; -x_2 = l_Lean_Parser_Command_export_parenthesizer___closed__4; +x_1 = l_Lean_Parser_Command_namedPrio_parenthesizer___closed__2; +x_2 = l_Lean_Parser_Command_export_parenthesizer___closed__3; 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; } } -static lean_object* _init_l_Lean_Parser_Command_export_parenthesizer___closed__6() { +static lean_object* _init_l_Lean_Parser_Command_export_parenthesizer___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_declId_parenthesizer___closed__4; -x_2 = l_Lean_Parser_Command_export_parenthesizer___closed__5; +x_2 = l_Lean_Parser_Command_export_parenthesizer___closed__4; 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; } } -static lean_object* _init_l_Lean_Parser_Command_export_parenthesizer___closed__7() { +static lean_object* _init_l_Lean_Parser_Command_export_parenthesizer___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_export_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Command_export_parenthesizer___closed__6; +x_2 = l_Lean_Parser_Command_export_parenthesizer___closed__5; 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; } } -static lean_object* _init_l_Lean_Parser_Command_export_parenthesizer___closed__8() { +static lean_object* _init_l_Lean_Parser_Command_export_parenthesizer___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_export___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_export_parenthesizer___closed__7; +x_3 = l_Lean_Parser_Command_export_parenthesizer___closed__6; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -33697,7 +33896,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_export_parenthesizer(lean_object* { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Command_export_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Command_export_parenthesizer___closed__8; +x_7 = l_Lean_Parser_Command_export_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; } @@ -33832,7 +34031,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_import_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(235u); +x_1 = lean_unsigned_to_nat(234u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -33844,7 +34043,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_import_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(235u); x_2 = lean_unsigned_to_nat(10u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -33872,7 +34071,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_import_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(235u); +x_1 = lean_unsigned_to_nat(234u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -33884,7 +34083,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_import_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(235u); +x_1 = lean_unsigned_to_nat(234u); x_2 = lean_unsigned_to_nat(38u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -34136,7 +34335,7 @@ static lean_object* _init_l_Lean_Parser_Command_openHiding___closed__4() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("hiding", 6); +x_1 = lean_mk_string_from_bytes(" hiding", 7); return x_1; } } @@ -34198,49 +34397,69 @@ return x_3; static lean_object* _init_l_Lean_Parser_Command_openHiding___closed__11() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_skip; +x_2 = l_Lean_Parser_Command_openHiding___closed__10; +x_3 = l_Lean_Parser_andthen(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Command_openHiding___closed__12() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_openHiding___closed__10; +x_1 = l_Lean_Parser_Command_openHiding___closed__11; x_2 = l_Lean_Parser_many1(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_openHiding___closed__12() { +static lean_object* _init_l_Lean_Parser_Command_openHiding___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_openHiding___closed__7; -x_2 = l_Lean_Parser_Command_openHiding___closed__11; +x_2 = l_Lean_Parser_Command_openHiding___closed__12; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_openHiding___closed__13() { +static lean_object* _init_l_Lean_Parser_Command_openHiding___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_skip; +x_2 = l_Lean_Parser_Command_openHiding___closed__13; +x_3 = l_Lean_Parser_andthen(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Command_openHiding___closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_openHiding___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_openHiding___closed__12; +x_3 = l_Lean_Parser_Command_openHiding___closed__14; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Command_openHiding___closed__14() { +static lean_object* _init_l_Lean_Parser_Command_openHiding___closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_openHiding___closed__3; -x_2 = l_Lean_Parser_Command_openHiding___closed__13; +x_2 = l_Lean_Parser_Command_openHiding___closed__15; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_openHiding___closed__15() { +static lean_object* _init_l_Lean_Parser_Command_openHiding___closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_openHiding___closed__2; -x_2 = l_Lean_Parser_Command_openHiding___closed__14; +x_2 = l_Lean_Parser_Command_openHiding___closed__16; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -34249,7 +34468,7 @@ static lean_object* _init_l_Lean_Parser_Command_openHiding() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Command_openHiding___closed__15; +x_1 = l_Lean_Parser_Command_openHiding___closed__17; return x_1; } } @@ -34406,7 +34625,7 @@ static lean_object* _init_l_Lean_Parser_Command_openRenaming___closed__4() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("renaming", 8); +x_1 = lean_mk_string_from_bytes(" renaming ", 10); return x_1; } } @@ -34463,30 +34682,40 @@ return x_3; static lean_object* _init_l_Lean_Parser_Command_openRenaming___closed__10() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_skip; +x_2 = l_Lean_Parser_Command_openRenaming___closed__9; +x_3 = l_Lean_Parser_andthen(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Command_openRenaming___closed__11() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_openRenaming___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_openRenaming___closed__9; +x_3 = l_Lean_Parser_Command_openRenaming___closed__10; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Command_openRenaming___closed__11() { +static lean_object* _init_l_Lean_Parser_Command_openRenaming___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_openRenaming___closed__3; -x_2 = l_Lean_Parser_Command_openRenaming___closed__10; +x_2 = l_Lean_Parser_Command_openRenaming___closed__11; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_openRenaming___closed__12() { +static lean_object* _init_l_Lean_Parser_Command_openRenaming___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_openRenaming___closed__2; -x_2 = l_Lean_Parser_Command_openRenaming___closed__11; +x_2 = l_Lean_Parser_Command_openRenaming___closed__12; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -34495,7 +34724,7 @@ static lean_object* _init_l_Lean_Parser_Command_openRenaming() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Command_openRenaming___closed__12; +x_1 = l_Lean_Parser_Command_openRenaming___closed__13; return x_1; } } @@ -34536,7 +34765,7 @@ static lean_object* _init_l_Lean_Parser_Command_openOnly___closed__4() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_ident; -x_2 = l_Lean_Parser_Command_export___closed__7; +x_2 = l_Lean_Parser_Command_namedPrio___closed__5; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -34555,7 +34784,7 @@ static lean_object* _init_l_Lean_Parser_Command_openOnly___closed__6() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_openOnly___closed__5; -x_2 = l_Lean_Parser_Command_export___closed__8; +x_2 = l_Lean_Parser_Command_export___closed__6; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -34563,30 +34792,40 @@ return x_3; static lean_object* _init_l_Lean_Parser_Command_openOnly___closed__7() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_skip; +x_2 = l_Lean_Parser_Command_openOnly___closed__6; +x_3 = l_Lean_Parser_andthen(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Command_openOnly___closed__8() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_openOnly___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_openOnly___closed__6; +x_3 = l_Lean_Parser_Command_openOnly___closed__7; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Command_openOnly___closed__8() { +static lean_object* _init_l_Lean_Parser_Command_openOnly___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_openOnly___closed__3; -x_2 = l_Lean_Parser_Command_openOnly___closed__7; +x_2 = l_Lean_Parser_Command_openOnly___closed__8; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_openOnly___closed__9() { +static lean_object* _init_l_Lean_Parser_Command_openOnly___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_openOnly___closed__2; -x_2 = l_Lean_Parser_Command_openOnly___closed__8; +x_2 = l_Lean_Parser_Command_openOnly___closed__9; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -34595,7 +34834,7 @@ static lean_object* _init_l_Lean_Parser_Command_openOnly() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Command_openOnly___closed__9; +x_1 = l_Lean_Parser_Command_openOnly___closed__10; return x_1; } } @@ -34637,7 +34876,7 @@ static lean_object* _init_l_Lean_Parser_Command_openSimple___closed__4() { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_openSimple___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_openHiding___closed__11; +x_3 = l_Lean_Parser_Command_openHiding___closed__12; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } @@ -34706,7 +34945,7 @@ static lean_object* _init_l_Lean_Parser_Command_openScoped___closed__4() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("scoped ", 7); +x_1 = lean_mk_string_from_bytes(" scoped", 7); return x_1; } } @@ -34724,7 +34963,7 @@ static lean_object* _init_l_Lean_Parser_Command_openScoped___closed__6() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_openScoped___closed__5; -x_2 = l_Lean_Parser_Command_openHiding___closed__11; +x_2 = l_Lean_Parser_Command_openHiding___closed__12; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -34892,66 +35131,58 @@ return x_5; static lean_object* _init_l_Lean_Parser_Command_open___closed__4() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("open ", 5); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Command_open___closed__5() { -_start: -{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_open___closed__4; +x_1 = l_Lean_Parser_Command_open___closed__1; x_2 = l_Lean_Parser_symbol(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_open___closed__6() { +static lean_object* _init_l_Lean_Parser_Command_open___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_open___closed__5; +x_1 = l_Lean_Parser_Command_open___closed__4; x_2 = l_Lean_Parser_Command_openDecl; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_open___closed__7() { +static lean_object* _init_l_Lean_Parser_Command_open___closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_open___closed__6; +x_1 = l_Lean_Parser_Command_open___closed__5; x_2 = l_Lean_Parser_withPosition(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_open___closed__8() { +static lean_object* _init_l_Lean_Parser_Command_open___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_open___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_open___closed__7; +x_3 = l_Lean_Parser_Command_open___closed__6; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Command_open___closed__9() { +static lean_object* _init_l_Lean_Parser_Command_open___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_open___closed__3; -x_2 = l_Lean_Parser_Command_open___closed__8; +x_2 = l_Lean_Parser_Command_open___closed__7; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_open___closed__10() { +static lean_object* _init_l_Lean_Parser_Command_open___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_open___closed__2; -x_2 = l_Lean_Parser_Command_open___closed__9; +x_2 = l_Lean_Parser_Command_open___closed__8; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -34960,7 +35191,7 @@ static lean_object* _init_l_Lean_Parser_Command_open() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Command_open___closed__10; +x_1 = l_Lean_Parser_Command_open___closed__9; return x_1; } } @@ -34981,7 +35212,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(252u); +x_1 = lean_unsigned_to_nat(251u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -34993,8 +35224,8 @@ 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(253u); -x_2 = lean_unsigned_to_nat(36u); +x_1 = lean_unsigned_to_nat(252u); +x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -35008,7 +35239,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Parser_Command_open_declRange___closed__1; x_2 = lean_unsigned_to_nat(26u); x_3 = l___regBuiltin_Lean_Parser_Command_open_declRange___closed__2; -x_4 = lean_unsigned_to_nat(36u); +x_4 = lean_unsigned_to_nat(35u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -35021,7 +35252,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(252u); +x_1 = lean_unsigned_to_nat(251u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -35033,7 +35264,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(252u); +x_1 = lean_unsigned_to_nat(251u); x_2 = lean_unsigned_to_nat(36u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -35152,32 +35383,56 @@ return x_3; static lean_object* _init_l_Lean_Parser_Command_openHiding_formatter___closed__7() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Command_optDeclSig_formatter___closed__5; +x_2 = l_Lean_Parser_Command_openHiding_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; +} +} +static lean_object* _init_l_Lean_Parser_Command_openHiding_formatter___closed__8() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_openHiding_formatter___closed__6; +x_1 = l_Lean_Parser_Command_openHiding_formatter___closed__7; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_many1_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_openHiding_formatter___closed__8() { +static lean_object* _init_l_Lean_Parser_Command_openHiding_formatter___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_openHiding_formatter___closed__4; -x_2 = l_Lean_Parser_Command_openHiding_formatter___closed__7; +x_2 = l_Lean_Parser_Command_openHiding_formatter___closed__8; 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; } } -static lean_object* _init_l_Lean_Parser_Command_openHiding_formatter___closed__9() { +static lean_object* _init_l_Lean_Parser_Command_openHiding_formatter___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Command_optDeclSig_formatter___closed__5; +x_2 = l_Lean_Parser_Command_openHiding_formatter___closed__9; +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; +} +} +static lean_object* _init_l_Lean_Parser_Command_openHiding_formatter___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_openHiding___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_openHiding_formatter___closed__8; +x_3 = l_Lean_Parser_Command_openHiding_formatter___closed__10; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -35190,7 +35445,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_openHiding_formatter(lean_object* { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Command_openHiding_formatter___closed__1; -x_7 = l_Lean_Parser_Command_openHiding_formatter___closed__9; +x_7 = l_Lean_Parser_Command_openHiding_formatter___closed__11; 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; } @@ -35421,10 +35676,22 @@ return x_3; static lean_object* _init_l_Lean_Parser_Command_openRenaming_formatter___closed__7() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Command_optDeclSig_formatter___closed__5; +x_2 = l_Lean_Parser_Command_openRenaming_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; +} +} +static lean_object* _init_l_Lean_Parser_Command_openRenaming_formatter___closed__8() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_openRenaming___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_openRenaming_formatter___closed__6; +x_3 = l_Lean_Parser_Command_openRenaming_formatter___closed__7; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -35437,7 +35704,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_openRenaming_formatter(lean_objec { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Command_openRenaming_formatter___closed__1; -x_7 = l_Lean_Parser_Command_openRenaming_formatter___closed__7; +x_7 = l_Lean_Parser_Command_openRenaming_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; } @@ -35498,7 +35765,7 @@ static lean_object* _init_l_Lean_Parser_Command_openOnly_formatter___closed__2() { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_declId_formatter___closed__4; -x_2 = l_Lean_Parser_Command_export_formatter___closed__3; +x_2 = l_Lean_Parser_Command_namedPrio_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); @@ -35520,7 +35787,7 @@ static lean_object* _init_l_Lean_Parser_Command_openOnly_formatter___closed__4() { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_openOnly_formatter___closed__3; -x_2 = l_Lean_Parser_Command_export_formatter___closed__4; +x_2 = l_Lean_Parser_Command_export_formatter___closed__3; 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); @@ -35530,10 +35797,22 @@ return x_3; static lean_object* _init_l_Lean_Parser_Command_openOnly_formatter___closed__5() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Command_optDeclSig_formatter___closed__5; +x_2 = l_Lean_Parser_Command_openOnly_formatter___closed__4; +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; +} +} +static lean_object* _init_l_Lean_Parser_Command_openOnly_formatter___closed__6() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_openOnly___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_openOnly_formatter___closed__4; +x_3 = l_Lean_Parser_Command_openOnly_formatter___closed__5; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -35546,7 +35825,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_openOnly_formatter(lean_object* x { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Command_openOnly_formatter___closed__1; -x_7 = l_Lean_Parser_Command_openOnly_formatter___closed__5; +x_7 = l_Lean_Parser_Command_openOnly_formatter___closed__6; 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; } @@ -35608,7 +35887,7 @@ static lean_object* _init_l_Lean_Parser_Command_openSimple_formatter___closed__2 lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_openSimple___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_openHiding_formatter___closed__7; +x_3 = l_Lean_Parser_Command_openHiding_formatter___closed__8; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -35692,7 +35971,7 @@ static lean_object* _init_l_Lean_Parser_Command_openScoped_formatter___closed__3 { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_openScoped_formatter___closed__2; -x_2 = l_Lean_Parser_Command_openHiding_formatter___closed__7; +x_2 = l_Lean_Parser_Command_openHiding_formatter___closed__8; 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); @@ -35853,7 +36132,7 @@ static lean_object* _init_l_Lean_Parser_Command_open_formatter___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_open___closed__4; +x_1 = l_Lean_Parser_Command_open___closed__1; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; @@ -36019,32 +36298,56 @@ return x_3; static lean_object* _init_l_Lean_Parser_Command_openHiding_parenthesizer___closed__7() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Command_optDeclSig_parenthesizer___closed__5; +x_2 = l_Lean_Parser_Command_openHiding_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; +} +} +static lean_object* _init_l_Lean_Parser_Command_openHiding_parenthesizer___closed__8() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_openHiding_parenthesizer___closed__6; +x_1 = l_Lean_Parser_Command_openHiding_parenthesizer___closed__7; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_many1_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_openHiding_parenthesizer___closed__8() { +static lean_object* _init_l_Lean_Parser_Command_openHiding_parenthesizer___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_openHiding_parenthesizer___closed__4; -x_2 = l_Lean_Parser_Command_openHiding_parenthesizer___closed__7; +x_2 = l_Lean_Parser_Command_openHiding_parenthesizer___closed__8; 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; } } -static lean_object* _init_l_Lean_Parser_Command_openHiding_parenthesizer___closed__9() { +static lean_object* _init_l_Lean_Parser_Command_openHiding_parenthesizer___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Command_optDeclSig_parenthesizer___closed__5; +x_2 = l_Lean_Parser_Command_openHiding_parenthesizer___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); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Command_openHiding_parenthesizer___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_openHiding___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_openHiding_parenthesizer___closed__8; +x_3 = l_Lean_Parser_Command_openHiding_parenthesizer___closed__10; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -36057,7 +36360,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_openHiding_parenthesizer(lean_obj { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Command_openHiding_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Command_openHiding_parenthesizer___closed__9; +x_7 = l_Lean_Parser_Command_openHiding_parenthesizer___closed__11; 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; } @@ -36288,10 +36591,22 @@ return x_3; static lean_object* _init_l_Lean_Parser_Command_openRenaming_parenthesizer___closed__7() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Command_optDeclSig_parenthesizer___closed__5; +x_2 = l_Lean_Parser_Command_openRenaming_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; +} +} +static lean_object* _init_l_Lean_Parser_Command_openRenaming_parenthesizer___closed__8() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_openRenaming___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_openRenaming_parenthesizer___closed__6; +x_3 = l_Lean_Parser_Command_openRenaming_parenthesizer___closed__7; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -36304,7 +36619,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_openRenaming_parenthesizer(lean_o { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Command_openRenaming_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Command_openRenaming_parenthesizer___closed__7; +x_7 = l_Lean_Parser_Command_openRenaming_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; } @@ -36365,7 +36680,7 @@ static lean_object* _init_l_Lean_Parser_Command_openOnly_parenthesizer___closed_ { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_declId_parenthesizer___closed__4; -x_2 = l_Lean_Parser_Command_export_parenthesizer___closed__3; +x_2 = l_Lean_Parser_Command_namedPrio_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); @@ -36387,7 +36702,7 @@ static lean_object* _init_l_Lean_Parser_Command_openOnly_parenthesizer___closed_ { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_openOnly_parenthesizer___closed__3; -x_2 = l_Lean_Parser_Command_export_parenthesizer___closed__4; +x_2 = l_Lean_Parser_Command_export_parenthesizer___closed__3; 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); @@ -36397,10 +36712,22 @@ return x_3; static lean_object* _init_l_Lean_Parser_Command_openOnly_parenthesizer___closed__5() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Command_optDeclSig_parenthesizer___closed__5; +x_2 = l_Lean_Parser_Command_openOnly_parenthesizer___closed__4; +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; +} +} +static lean_object* _init_l_Lean_Parser_Command_openOnly_parenthesizer___closed__6() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_openOnly___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_openOnly_parenthesizer___closed__4; +x_3 = l_Lean_Parser_Command_openOnly_parenthesizer___closed__5; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -36413,7 +36740,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_openOnly_parenthesizer(lean_objec { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Command_openOnly_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Command_openOnly_parenthesizer___closed__5; +x_7 = l_Lean_Parser_Command_openOnly_parenthesizer___closed__6; 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; } @@ -36475,7 +36802,7 @@ static lean_object* _init_l_Lean_Parser_Command_openSimple_parenthesizer___close lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_openSimple___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_openHiding_parenthesizer___closed__7; +x_3 = l_Lean_Parser_Command_openHiding_parenthesizer___closed__8; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -36559,7 +36886,7 @@ static lean_object* _init_l_Lean_Parser_Command_openScoped_parenthesizer___close { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_openScoped_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Command_openHiding_parenthesizer___closed__7; +x_2 = l_Lean_Parser_Command_openHiding_parenthesizer___closed__8; 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); @@ -36720,7 +37047,7 @@ static lean_object* _init_l_Lean_Parser_Command_open_parenthesizer___closed__2() _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_open___closed__4; +x_1 = l_Lean_Parser_Command_open___closed__1; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; @@ -36848,135 +37175,118 @@ return x_5; static lean_object* _init_l_Lean_Parser_Command_mutual___closed__4() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("mutual ", 7); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Command_mutual___closed__5() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_mutual___closed__4; -x_2 = l_Lean_Parser_symbol(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Parser_Command_mutual___closed__6() { -_start: -{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_end___closed__1; +x_1 = l_Lean_Parser_Command_mutual___closed__1; x_2 = l_Lean_Parser_symbol(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_mutual___closed__7() { +static lean_object* _init_l_Lean_Parser_Command_mutual___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_mutual___closed__6; +x_1 = l_Lean_Parser_Command_end___closed__4; x_2 = l_Lean_Parser_Command_end___closed__1; x_3 = l_Lean_Parser_notFollowedBy(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_mutual___closed__8() { +static lean_object* _init_l_Lean_Parser_Command_mutual___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_mutual___closed__7; +x_1 = l_Lean_Parser_Command_mutual___closed__5; x_2 = l_Lean_Parser_Command_quot___closed__6; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_mutual___closed__9() { +static lean_object* _init_l_Lean_Parser_Command_mutual___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_skip; -x_2 = l_Lean_Parser_Command_mutual___closed__8; +x_2 = l_Lean_Parser_Command_mutual___closed__6; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_mutual___closed__10() { +static lean_object* _init_l_Lean_Parser_Command_mutual___closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_mutual___closed__9; +x_1 = l_Lean_Parser_Command_mutual___closed__7; x_2 = l_Lean_Parser_many1(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_mutual___closed__11() { +static lean_object* _init_l_Lean_Parser_Command_mutual___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_skip; -x_2 = l_Lean_Parser_Command_mutual___closed__6; +x_2 = l_Lean_Parser_Command_end___closed__4; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_mutual___closed__12() { +static lean_object* _init_l_Lean_Parser_Command_mutual___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_mutual___closed__11; +x_1 = l_Lean_Parser_Command_mutual___closed__9; x_2 = l_Lean_Parser_Command_terminationSuffix; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_mutual___closed__13() { +static lean_object* _init_l_Lean_Parser_Command_mutual___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_mutual___closed__10; -x_2 = l_Lean_Parser_Command_mutual___closed__12; +x_1 = l_Lean_Parser_Command_mutual___closed__8; +x_2 = l_Lean_Parser_Command_mutual___closed__10; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_mutual___closed__14() { +static lean_object* _init_l_Lean_Parser_Command_mutual___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_mutual___closed__5; -x_2 = l_Lean_Parser_Command_mutual___closed__13; +x_1 = l_Lean_Parser_Command_mutual___closed__4; +x_2 = l_Lean_Parser_Command_mutual___closed__11; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_mutual___closed__15() { +static lean_object* _init_l_Lean_Parser_Command_mutual___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_mutual___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_mutual___closed__14; +x_3 = l_Lean_Parser_Command_mutual___closed__12; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Command_mutual___closed__16() { +static lean_object* _init_l_Lean_Parser_Command_mutual___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_mutual___closed__3; -x_2 = l_Lean_Parser_Command_mutual___closed__15; +x_2 = l_Lean_Parser_Command_mutual___closed__13; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_mutual___closed__17() { +static lean_object* _init_l_Lean_Parser_Command_mutual___closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_mutual___closed__2; -x_2 = l_Lean_Parser_Command_mutual___closed__16; +x_2 = l_Lean_Parser_Command_mutual___closed__14; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -36985,7 +37295,7 @@ static lean_object* _init_l_Lean_Parser_Command_mutual() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Command_mutual___closed__17; +x_1 = l_Lean_Parser_Command_mutual___closed__15; return x_1; } } @@ -37006,7 +37316,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(255u); +x_1 = lean_unsigned_to_nat(254u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -37018,7 +37328,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(257u); +x_1 = lean_unsigned_to_nat(256u); x_2 = lean_unsigned_to_nat(49u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -37046,7 +37356,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(255u); +x_1 = lean_unsigned_to_nat(254u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -37058,7 +37368,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(255u); +x_1 = lean_unsigned_to_nat(254u); x_2 = lean_unsigned_to_nat(38u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -37126,7 +37436,7 @@ static lean_object* _init_l_Lean_Parser_Command_mutual_formatter___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_mutual___closed__4; +x_1 = l_Lean_Parser_Command_mutual___closed__1; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; @@ -37179,40 +37489,30 @@ return x_2; static lean_object* _init_l_Lean_Parser_Command_mutual_formatter___closed__7() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_end___closed__1; -x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_formatter), 6, 1); -lean_closure_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Parser_Command_mutual_formatter___closed__8() { -_start: -{ lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_moduleDoc_formatter___closed__4; -x_2 = l_Lean_Parser_Command_mutual_formatter___closed__7; +x_2 = l_Lean_Parser_Command_end_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; } } -static lean_object* _init_l_Lean_Parser_Command_mutual_formatter___closed__9() { +static lean_object* _init_l_Lean_Parser_Command_mutual_formatter___closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_mutual_formatter___closed__8; +x_1 = l_Lean_Parser_Command_mutual_formatter___closed__7; x_2 = lean_alloc_closure((void*)(l_Lean_ppDedent_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_mutual_formatter___closed__10() { +static lean_object* _init_l_Lean_Parser_Command_mutual_formatter___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_mutual_formatter___closed__9; +x_1 = l_Lean_Parser_Command_mutual_formatter___closed__8; x_2 = l_Lean_Parser_Command_abbrev_formatter___closed__5; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -37220,37 +37520,37 @@ lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_mutual_formatter___closed__11() { +static lean_object* _init_l_Lean_Parser_Command_mutual_formatter___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_mutual_formatter___closed__6; -x_2 = l_Lean_Parser_Command_mutual_formatter___closed__10; +x_2 = l_Lean_Parser_Command_mutual_formatter___closed__9; 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; } } -static lean_object* _init_l_Lean_Parser_Command_mutual_formatter___closed__12() { +static lean_object* _init_l_Lean_Parser_Command_mutual_formatter___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_mutual_formatter___closed__2; -x_2 = l_Lean_Parser_Command_mutual_formatter___closed__11; +x_2 = l_Lean_Parser_Command_mutual_formatter___closed__10; 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; } } -static lean_object* _init_l_Lean_Parser_Command_mutual_formatter___closed__13() { +static lean_object* _init_l_Lean_Parser_Command_mutual_formatter___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_mutual___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_mutual_formatter___closed__12; +x_3 = l_Lean_Parser_Command_mutual_formatter___closed__11; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -37263,7 +37563,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_mutual_formatter(lean_object* x_1 { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Command_mutual_formatter___closed__1; -x_7 = l_Lean_Parser_Command_mutual_formatter___closed__13; +x_7 = l_Lean_Parser_Command_mutual_formatter___closed__12; 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; } @@ -37323,7 +37623,7 @@ static lean_object* _init_l_Lean_Parser_Command_mutual_parenthesizer___closed__2 _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_mutual___closed__4; +x_1 = l_Lean_Parser_Command_mutual___closed__1; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; @@ -37376,40 +37676,30 @@ return x_2; static lean_object* _init_l_Lean_Parser_Command_mutual_parenthesizer___closed__7() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_end___closed__1; -x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_parenthesizer), 6, 1); -lean_closure_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Parser_Command_mutual_parenthesizer___closed__8() { -_start: -{ lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_moduleDoc_parenthesizer___closed__4; -x_2 = l_Lean_Parser_Command_mutual_parenthesizer___closed__7; +x_2 = l_Lean_Parser_Command_end_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; } } -static lean_object* _init_l_Lean_Parser_Command_mutual_parenthesizer___closed__9() { +static lean_object* _init_l_Lean_Parser_Command_mutual_parenthesizer___closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_mutual_parenthesizer___closed__8; +x_1 = l_Lean_Parser_Command_mutual_parenthesizer___closed__7; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_ppDedent_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_mutual_parenthesizer___closed__10() { +static lean_object* _init_l_Lean_Parser_Command_mutual_parenthesizer___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_mutual_parenthesizer___closed__9; +x_1 = l_Lean_Parser_Command_mutual_parenthesizer___closed__8; x_2 = l_Lean_Parser_Command_abbrev_parenthesizer___closed__5; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -37417,37 +37707,37 @@ lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_mutual_parenthesizer___closed__11() { +static lean_object* _init_l_Lean_Parser_Command_mutual_parenthesizer___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_mutual_parenthesizer___closed__6; -x_2 = l_Lean_Parser_Command_mutual_parenthesizer___closed__10; +x_2 = l_Lean_Parser_Command_mutual_parenthesizer___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); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_mutual_parenthesizer___closed__12() { +static lean_object* _init_l_Lean_Parser_Command_mutual_parenthesizer___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_mutual_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Command_mutual_parenthesizer___closed__11; +x_2 = l_Lean_Parser_Command_mutual_parenthesizer___closed__10; 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; } } -static lean_object* _init_l_Lean_Parser_Command_mutual_parenthesizer___closed__13() { +static lean_object* _init_l_Lean_Parser_Command_mutual_parenthesizer___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_mutual___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_mutual_parenthesizer___closed__12; +x_3 = l_Lean_Parser_Command_mutual_parenthesizer___closed__11; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -37460,7 +37750,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_mutual_parenthesizer(lean_object* { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Command_mutual_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Command_mutual_parenthesizer___closed__13; +x_7 = l_Lean_Parser_Command_mutual_parenthesizer___closed__12; 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; } @@ -37649,7 +37939,7 @@ static lean_object* _init_l_Lean_Parser_Command_initialize___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_typeSpec; +x_1 = l_Lean_Parser_skip; x_2 = l_Lean_Parser_Term_leftArrow; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; @@ -37659,7 +37949,7 @@ static lean_object* _init_l_Lean_Parser_Command_initialize___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_ident; +x_1 = l_Lean_Parser_Term_typeSpec; x_2 = l_Lean_Parser_Command_initialize___closed__4; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; @@ -37668,78 +37958,88 @@ return x_3; static lean_object* _init_l_Lean_Parser_Command_initialize___closed__6() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_ident; +x_2 = l_Lean_Parser_Command_initialize___closed__5; +x_3 = l_Lean_Parser_andthen(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Command_initialize___closed__7() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_initialize___closed__5; +x_1 = l_Lean_Parser_Command_initialize___closed__6; x_2 = l_Lean_Parser_atomic(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initialize___closed__7() { +static lean_object* _init_l_Lean_Parser_Command_initialize___closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_initialize___closed__6; +x_1 = l_Lean_Parser_Command_initialize___closed__7; x_2 = l_Lean_Parser_optional(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initialize___closed__8() { +static lean_object* _init_l_Lean_Parser_Command_initialize___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_initialize___closed__7; +x_1 = l_Lean_Parser_Command_initialize___closed__8; x_2 = l_Lean_Parser_Term_doSeq; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_initialize___closed__9() { +static lean_object* _init_l_Lean_Parser_Command_initialize___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_initializeKeyword; -x_2 = l_Lean_Parser_Command_initialize___closed__8; +x_2 = l_Lean_Parser_Command_initialize___closed__9; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_initialize___closed__10() { +static lean_object* _init_l_Lean_Parser_Command_initialize___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_declaration___closed__4; -x_2 = l_Lean_Parser_Command_initialize___closed__9; +x_2 = l_Lean_Parser_Command_initialize___closed__10; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_initialize___closed__11() { +static lean_object* _init_l_Lean_Parser_Command_initialize___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_initialize___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_initialize___closed__10; +x_3 = l_Lean_Parser_Command_initialize___closed__11; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Command_initialize___closed__12() { +static lean_object* _init_l_Lean_Parser_Command_initialize___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_initialize___closed__3; -x_2 = l_Lean_Parser_Command_initialize___closed__11; +x_2 = l_Lean_Parser_Command_initialize___closed__12; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_initialize___closed__13() { +static lean_object* _init_l_Lean_Parser_Command_initialize___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_initialize___closed__2; -x_2 = l_Lean_Parser_Command_initialize___closed__12; +x_2 = l_Lean_Parser_Command_initialize___closed__13; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -37748,7 +38048,7 @@ static lean_object* _init_l_Lean_Parser_Command_initialize() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Command_initialize___closed__13; +x_1 = l_Lean_Parser_Command_initialize___closed__14; return x_1; } } @@ -37769,7 +38069,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(260u); +x_1 = lean_unsigned_to_nat(259u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -37781,8 +38081,8 @@ 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(262u); -x_2 = lean_unsigned_to_nat(76u); +x_1 = lean_unsigned_to_nat(261u); +x_2 = lean_unsigned_to_nat(87u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -37796,7 +38096,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Parser_Command_initialize_declRange___closed__1; x_2 = lean_unsigned_to_nat(26u); x_3 = l___regBuiltin_Lean_Parser_Command_initialize_declRange___closed__2; -x_4 = lean_unsigned_to_nat(76u); +x_4 = lean_unsigned_to_nat(87u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -37809,7 +38109,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(260u); +x_1 = lean_unsigned_to_nat(259u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -37821,7 +38121,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(260u); +x_1 = lean_unsigned_to_nat(259u); x_2 = lean_unsigned_to_nat(42u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -38004,7 +38304,7 @@ static lean_object* _init_l_Lean_Parser_Command_initialize_formatter___closed__3 _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_declSig_formatter___closed__2; +x_1 = l_Lean_Parser_Command_optDeclSig_formatter___closed__5; x_2 = l_Lean_Parser_Command_initialize_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); @@ -38016,7 +38316,7 @@ static lean_object* _init_l_Lean_Parser_Command_initialize_formatter___closed__4 _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_declId_formatter___closed__4; +x_1 = l_Lean_Parser_Command_declSig_formatter___closed__2; x_2 = l_Lean_Parser_Command_initialize_formatter___closed__3; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -38027,24 +38327,36 @@ return x_3; static lean_object* _init_l_Lean_Parser_Command_initialize_formatter___closed__5() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Command_declId_formatter___closed__4; +x_2 = l_Lean_Parser_Command_initialize_formatter___closed__4; +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; +} +} +static lean_object* _init_l_Lean_Parser_Command_initialize_formatter___closed__6() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_initialize_formatter___closed__4; +x_1 = l_Lean_Parser_Command_initialize_formatter___closed__5; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_atomic_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initialize_formatter___closed__6() { +static lean_object* _init_l_Lean_Parser_Command_initialize_formatter___closed__7() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_initialize_formatter___closed__5; +x_1 = l_Lean_Parser_Command_initialize_formatter___closed__6; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_optional_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initialize_formatter___closed__7() { +static lean_object* _init_l_Lean_Parser_Command_initialize_formatter___closed__8() { _start: { lean_object* x_1; @@ -38052,49 +38364,49 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Term_doSeq_formatter), 5, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_Command_initialize_formatter___closed__8() { +static lean_object* _init_l_Lean_Parser_Command_initialize_formatter___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_initialize_formatter___closed__6; -x_2 = l_Lean_Parser_Command_initialize_formatter___closed__7; +x_1 = l_Lean_Parser_Command_initialize_formatter___closed__7; +x_2 = l_Lean_Parser_Command_initialize_formatter___closed__8; 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; } } -static lean_object* _init_l_Lean_Parser_Command_initialize_formatter___closed__9() { +static lean_object* _init_l_Lean_Parser_Command_initialize_formatter___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_Lean_Parser_Command_initializeKeyword_formatter___closed__2; -x_2 = l_Lean_Parser_Command_initialize_formatter___closed__8; +x_2 = l_Lean_Parser_Command_initialize_formatter___closed__9; 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; } } -static lean_object* _init_l_Lean_Parser_Command_initialize_formatter___closed__10() { +static lean_object* _init_l_Lean_Parser_Command_initialize_formatter___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_declaration_formatter___closed__2; -x_2 = l_Lean_Parser_Command_initialize_formatter___closed__9; +x_2 = l_Lean_Parser_Command_initialize_formatter___closed__10; 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; } } -static lean_object* _init_l_Lean_Parser_Command_initialize_formatter___closed__11() { +static lean_object* _init_l_Lean_Parser_Command_initialize_formatter___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_initialize___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_initialize_formatter___closed__10; +x_3 = l_Lean_Parser_Command_initialize_formatter___closed__11; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -38107,7 +38419,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_initialize_formatter(lean_object* { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Command_initialize_formatter___closed__1; -x_7 = l_Lean_Parser_Command_initialize_formatter___closed__11; +x_7 = l_Lean_Parser_Command_initialize_formatter___closed__12; 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; } @@ -38282,7 +38594,7 @@ static lean_object* _init_l_Lean_Parser_Command_initialize_parenthesizer___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_declSig_parenthesizer___closed__2; +x_1 = l_Lean_Parser_Command_optDeclSig_parenthesizer___closed__5; x_2 = l_Lean_Parser_Command_initialize_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); @@ -38294,7 +38606,7 @@ static lean_object* _init_l_Lean_Parser_Command_initialize_parenthesizer___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_declId_parenthesizer___closed__4; +x_1 = l_Lean_Parser_Command_declSig_parenthesizer___closed__2; x_2 = l_Lean_Parser_Command_initialize_parenthesizer___closed__3; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -38305,24 +38617,36 @@ return x_3; static lean_object* _init_l_Lean_Parser_Command_initialize_parenthesizer___closed__5() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Command_declId_parenthesizer___closed__4; +x_2 = l_Lean_Parser_Command_initialize_parenthesizer___closed__4; +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; +} +} +static lean_object* _init_l_Lean_Parser_Command_initialize_parenthesizer___closed__6() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_initialize_parenthesizer___closed__4; +x_1 = l_Lean_Parser_Command_initialize_parenthesizer___closed__5; x_2 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_atomic_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initialize_parenthesizer___closed__6() { +static lean_object* _init_l_Lean_Parser_Command_initialize_parenthesizer___closed__7() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_initialize_parenthesizer___closed__5; +x_1 = l_Lean_Parser_Command_initialize_parenthesizer___closed__6; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_optional_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_initialize_parenthesizer___closed__7() { +static lean_object* _init_l_Lean_Parser_Command_initialize_parenthesizer___closed__8() { _start: { lean_object* x_1; @@ -38330,49 +38654,49 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Term_doSeq_parenthesizer), 5, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_Command_initialize_parenthesizer___closed__8() { +static lean_object* _init_l_Lean_Parser_Command_initialize_parenthesizer___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_initialize_parenthesizer___closed__6; -x_2 = l_Lean_Parser_Command_initialize_parenthesizer___closed__7; +x_1 = l_Lean_Parser_Command_initialize_parenthesizer___closed__7; +x_2 = l_Lean_Parser_Command_initialize_parenthesizer___closed__8; 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; } } -static lean_object* _init_l_Lean_Parser_Command_initialize_parenthesizer___closed__9() { +static lean_object* _init_l_Lean_Parser_Command_initialize_parenthesizer___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_Lean_Parser_Command_initializeKeyword_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Command_initialize_parenthesizer___closed__8; +x_2 = l_Lean_Parser_Command_initialize_parenthesizer___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); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_initialize_parenthesizer___closed__10() { +static lean_object* _init_l_Lean_Parser_Command_initialize_parenthesizer___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_declaration_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Command_initialize_parenthesizer___closed__9; +x_2 = l_Lean_Parser_Command_initialize_parenthesizer___closed__10; 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; } } -static lean_object* _init_l_Lean_Parser_Command_initialize_parenthesizer___closed__11() { +static lean_object* _init_l_Lean_Parser_Command_initialize_parenthesizer___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_initialize___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_initialize_parenthesizer___closed__10; +x_3 = l_Lean_Parser_Command_initialize_parenthesizer___closed__11; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -38385,7 +38709,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_initialize_parenthesizer(lean_obj { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Command_initialize_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Command_initialize_parenthesizer___closed__11; +x_7 = l_Lean_Parser_Command_initialize_parenthesizer___closed__12; 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; } @@ -38535,7 +38859,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(264u); +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); @@ -38548,7 +38872,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_in_declRange___clos { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_unsigned_to_nat(264u); -x_2 = lean_unsigned_to_nat(89u); +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); @@ -38562,7 +38886,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Parser_Command_in_declRange___closed__1; x_2 = lean_unsigned_to_nat(26u); x_3 = l___regBuiltin_Lean_Parser_Command_in_declRange___closed__2; -x_4 = lean_unsigned_to_nat(89u); +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); @@ -38575,7 +38899,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(264u); +x_1 = lean_unsigned_to_nat(263u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -38587,7 +38911,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(264u); +x_1 = lean_unsigned_to_nat(263u); x_2 = lean_unsigned_to_nat(34u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -38660,6 +38984,16 @@ static lean_object* _init_l_Lean_Parser_Command_in_formatter___closed__3() { { lean_object* x_1; lean_object* x_2; x_1 = l_Lean_Parser_Command_in_formatter___closed__2; +x_2 = lean_alloc_closure((void*)(l_Lean_ppDedent_formatter), 6, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Command_in_formatter___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Command_in_formatter___closed__3; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_withOpen_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; @@ -38672,7 +39006,7 @@ lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_obj x_6 = l_Lean_Parser_Command_in___closed__2; x_7 = lean_unsigned_to_nat(1024u); x_8 = lean_unsigned_to_nat(0u); -x_9 = l_Lean_Parser_Command_in_formatter___closed__3; +x_9 = l_Lean_Parser_Command_in_formatter___closed__4; x_10 = l_Lean_PrettyPrinter_Formatter_trailingNode_formatter(x_6, x_7, x_8, x_9, x_1, x_2, x_3, x_4, x_5); return x_10; } @@ -38737,6 +39071,16 @@ static lean_object* _init_l_Lean_Parser_Command_in_parenthesizer___closed__3() { { lean_object* x_1; lean_object* x_2; x_1 = l_Lean_Parser_Command_in_parenthesizer___closed__2; +x_2 = lean_alloc_closure((void*)(l_Lean_Parser_ppDedent_parenthesizer), 6, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Command_in_parenthesizer___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Command_in_parenthesizer___closed__3; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_withOpen_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; @@ -38749,7 +39093,7 @@ lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_obj x_6 = l_Lean_Parser_Command_in___closed__2; x_7 = lean_unsigned_to_nat(1024u); x_8 = lean_unsigned_to_nat(0u); -x_9 = l_Lean_Parser_Command_in_parenthesizer___closed__3; +x_9 = l_Lean_Parser_Command_in_parenthesizer___closed__4; x_10 = l_Lean_PrettyPrinter_Parenthesizer_trailingNode_parenthesizer(x_6, x_7, x_8, x_9, x_1, x_2, x_3, x_4, x_5); return x_10; } @@ -38823,7 +39167,7 @@ static lean_object* _init_l_Lean_Parser_Command_addDocString___closed__4() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("add_decl_doc", 12); +x_1 = lean_mk_string_from_bytes("add_decl_doc ", 13); return x_1; } } @@ -38925,7 +39269,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_addDocString_declRa { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_unsigned_to_nat(267u); -x_2 = lean_unsigned_to_nat(39u); +x_2 = lean_unsigned_to_nat(40u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -38939,7 +39283,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Parser_Command_addDocString_declRange___closed__1; x_2 = lean_unsigned_to_nat(26u); x_3 = l___regBuiltin_Lean_Parser_Command_addDocString_declRange___closed__2; -x_4 = lean_unsigned_to_nat(39u); +x_4 = lean_unsigned_to_nat(40u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -39911,7 +40255,7 @@ x_1 = l_Lean_Parser_Command_eoi___closed__5; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3040_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3138_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; @@ -39972,7 +40316,7 @@ x_1 = l_Lean_Parser_Command_ctor___closed__8; return x_1; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__1() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -39982,7 +40326,7 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__2() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__2() { _start: { lean_object* x_1; @@ -39990,19 +40334,19 @@ 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_3063____closed__3() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____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_quot___closed__1; x_2 = l_Lean_Parser_Term_quot___closed__2; x_3 = l_Lean_Parser_Command_quot___closed__1; -x_4 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__2; +x_4 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__2; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__4() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__4() { _start: { lean_object* x_1; lean_object* x_2; @@ -40012,7 +40356,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_3063____closed__5() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__5() { _start: { lean_object* x_1; lean_object* x_2; @@ -40022,7 +40366,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_3063____closed__6() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__6() { _start: { lean_object* x_1; lean_object* x_2; @@ -40032,12 +40376,12 @@ 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_3063____closed__7() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__7() { _start: { lean_object* x_1; lean_object* x_2; uint8_t x_3; lean_object* x_4; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__6; +x_2 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__6; x_3 = 1; x_4 = lean_alloc_ctor(0, 2, 1); lean_ctor_set(x_4, 0, x_1); @@ -40046,7 +40390,7 @@ lean_ctor_set_uint8(x_4, sizeof(void*)*2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__8() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__8() { _start: { lean_object* x_1; @@ -40054,17 +40398,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_3063____closed__9() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__9() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__8; +x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____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_3063____closed__10() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__10() { _start: { lean_object* x_1; @@ -40072,7 +40416,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_3063____closed__11() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__11() { _start: { lean_object* x_1; @@ -40080,17 +40424,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_3063____closed__12() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__12() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__11; +x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____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_3063____closed__13() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__13() { _start: { lean_object* x_1; @@ -40098,7 +40442,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_3063____closed__14() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__14() { _start: { lean_object* x_1; @@ -40106,17 +40450,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_3063____closed__15() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____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_3063____closed__14; +x_2 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__14; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__16() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__16() { _start: { lean_object* x_1; @@ -40124,19 +40468,19 @@ 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_3063____closed__17() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__17() { _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_quot___closed__1; x_2 = l_Lean_Parser_Term_quot___closed__2; x_3 = l_Lean_Parser_Command_quot___closed__1; -x_4 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__16; +x_4 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__16; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__18() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__18() { _start: { lean_object* x_1; lean_object* x_2; @@ -40146,7 +40490,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_3063____closed__19() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__19() { _start: { lean_object* x_1; @@ -40154,17 +40498,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_3063____closed__20() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__20() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__19; +x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____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_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__21() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__21() { _start: { lean_object* x_1; @@ -40172,17 +40516,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_3063____closed__22() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__22() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__21; +x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____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_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__23() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -40192,7 +40536,7 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__24() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__24() { _start: { lean_object* x_1; lean_object* x_2; @@ -40202,7 +40546,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_3063____closed__25() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__25() { _start: { lean_object* x_1; lean_object* x_2; @@ -40212,7 +40556,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_3063____closed__26() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__26() { _start: { lean_object* x_1; lean_object* x_2; @@ -40222,7 +40566,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_3063____closed__27() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__27() { _start: { lean_object* x_1; lean_object* x_2; @@ -40232,7 +40576,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_3063____closed__28() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__28() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -40242,7 +40586,7 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__29() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__29() { _start: { lean_object* x_1; lean_object* x_2; @@ -40252,7 +40596,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_3063____closed__30() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__30() { _start: { lean_object* x_1; lean_object* x_2; @@ -40262,7 +40606,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_3063____closed__31() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__31() { _start: { lean_object* x_1; lean_object* x_2; @@ -40272,7 +40616,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_3063____closed__32() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__32() { _start: { lean_object* x_1; lean_object* x_2; @@ -40282,7 +40626,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_3063____closed__33() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__33() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -40292,7 +40636,7 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__34() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__34() { _start: { lean_object* x_1; lean_object* x_2; @@ -40302,7 +40646,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_3063____closed__35() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__35() { _start: { lean_object* x_1; lean_object* x_2; @@ -40312,7 +40656,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_3063____closed__36() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__36() { _start: { lean_object* x_1; lean_object* x_2; @@ -40322,7 +40666,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_3063____closed__37() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__37() { _start: { lean_object* x_1; lean_object* x_2; @@ -40332,7 +40676,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_3063____closed__38() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__38() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -40342,7 +40686,7 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__39() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__39() { _start: { lean_object* x_1; lean_object* x_2; @@ -40352,7 +40696,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_3063____closed__40() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__40() { _start: { lean_object* x_1; lean_object* x_2; @@ -40362,7 +40706,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_3063____closed__41() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__41() { _start: { lean_object* x_1; lean_object* x_2; @@ -40372,7 +40716,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_3063____closed__42() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__42() { _start: { lean_object* x_1; lean_object* x_2; @@ -40382,7 +40726,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_3063____closed__43() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__43() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -40392,7 +40736,7 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__44() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__44() { _start: { lean_object* x_1; lean_object* x_2; @@ -40402,7 +40746,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_3063____closed__45() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__45() { _start: { lean_object* x_1; lean_object* x_2; @@ -40412,7 +40756,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_3063____closed__46() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__46() { _start: { lean_object* x_1; lean_object* x_2; @@ -40422,7 +40766,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_3063____closed__47() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__47() { _start: { lean_object* x_1; lean_object* x_2; @@ -40432,7 +40776,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_3063____closed__48() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__48() { _start: { lean_object* x_1; @@ -40440,29 +40784,29 @@ x_1 = lean_mk_string_from_bytes("docComment", 10); return x_1; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__49() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__49() { _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_3063____closed__48; +x_2 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__48; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__50() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__50() { _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_quot___closed__1; x_2 = l_Lean_Parser_Term_quot___closed__2; x_3 = l_Lean_Parser_Command_quot___closed__1; -x_4 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__48; +x_4 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__48; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__51() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__51() { _start: { lean_object* x_1; lean_object* x_2; @@ -40472,17 +40816,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_3063____closed__52() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__52() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__50; +x_1 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__50; 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_3063____closed__53() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__53() { _start: { lean_object* x_1; lean_object* x_2; @@ -40492,7 +40836,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_3063____closed__54() { +static lean_object* _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__54() { _start: { lean_object* x_1; lean_object* x_2; @@ -40502,15 +40846,15 @@ 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_3063_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161_(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; -x_2 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__1; -x_3 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__3; -x_4 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__4; -x_5 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__5; -x_6 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__7; +x_2 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__1; +x_3 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__3; +x_4 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__4; +x_5 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__5; +x_6 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__7; x_7 = l_Lean_Parser_registerAlias(x_2, x_3, x_4, x_5, x_6, x_1); if (lean_obj_tag(x_7) == 0) { @@ -40518,8 +40862,8 @@ 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_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__10; -x_10 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__9; +x_9 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__10; +x_10 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__9; x_11 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_2, x_10, x_8); if (lean_obj_tag(x_11) == 0) { @@ -40527,8 +40871,8 @@ 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_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__13; -x_14 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__12; +x_13 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__13; +x_14 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__12; x_15 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_2, x_14, x_12); if (lean_obj_tag(x_15) == 0) { @@ -40536,9 +40880,9 @@ lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean x_16 = lean_ctor_get(x_15, 1); lean_inc(x_16); lean_dec(x_15); -x_17 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__15; -x_18 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__17; -x_19 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__18; +x_17 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__15; +x_18 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__17; +x_19 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__18; x_20 = l_Lean_Parser_registerAlias(x_17, x_18, x_19, x_5, x_6, x_16); if (lean_obj_tag(x_20) == 0) { @@ -40546,7 +40890,7 @@ 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_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__20; +x_22 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__20; x_23 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_17, x_22, x_21); if (lean_obj_tag(x_23) == 0) { @@ -40554,7 +40898,7 @@ 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_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__22; +x_25 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__22; x_26 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_17, x_25, x_24); if (lean_obj_tag(x_26) == 0) { @@ -40562,10 +40906,10 @@ lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean x_27 = lean_ctor_get(x_26, 1); lean_inc(x_27); lean_dec(x_26); -x_28 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__23; +x_28 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__23; x_29 = l_Lean_Parser_Command_declId___closed__2; -x_30 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__24; -x_31 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__25; +x_30 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__24; +x_31 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__25; x_32 = l_Lean_Parser_registerAlias(x_28, x_29, x_30, x_31, x_6, x_27); if (lean_obj_tag(x_32) == 0) { @@ -40573,7 +40917,7 @@ 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_3063____closed__26; +x_34 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__26; x_35 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_28, x_34, x_33); if (lean_obj_tag(x_35) == 0) { @@ -40581,7 +40925,7 @@ 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 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__27; +x_37 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__27; x_38 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_28, x_37, x_36); if (lean_obj_tag(x_38) == 0) { @@ -40589,10 +40933,10 @@ lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean x_39 = lean_ctor_get(x_38, 1); lean_inc(x_39); lean_dec(x_38); -x_40 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__28; +x_40 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__28; x_41 = l_Lean_Parser_Command_declSig___closed__2; -x_42 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__29; -x_43 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__30; +x_42 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__29; +x_43 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__30; x_44 = l_Lean_Parser_registerAlias(x_40, x_41, x_42, x_43, x_6, x_39); if (lean_obj_tag(x_44) == 0) { @@ -40600,7 +40944,7 @@ 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_3063____closed__31; +x_46 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__31; x_47 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_40, x_46, x_45); if (lean_obj_tag(x_47) == 0) { @@ -40608,7 +40952,7 @@ lean_object* x_48; lean_object* x_49; lean_object* x_50; 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_3063____closed__32; +x_49 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__32; x_50 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_40, x_49, x_48); if (lean_obj_tag(x_50) == 0) { @@ -40616,10 +40960,10 @@ lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean x_51 = lean_ctor_get(x_50, 1); lean_inc(x_51); lean_dec(x_50); -x_52 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__33; +x_52 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__33; x_53 = l_Lean_Parser_Command_declVal___closed__2; -x_54 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__34; -x_55 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__35; +x_54 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__34; +x_55 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__35; x_56 = l_Lean_Parser_registerAlias(x_52, x_53, x_54, x_55, x_6, x_51); if (lean_obj_tag(x_56) == 0) { @@ -40627,7 +40971,7 @@ lean_object* x_57; lean_object* x_58; lean_object* x_59; x_57 = lean_ctor_get(x_56, 1); lean_inc(x_57); lean_dec(x_56); -x_58 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__36; +x_58 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__36; x_59 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_52, x_58, x_57); if (lean_obj_tag(x_59) == 0) { @@ -40635,7 +40979,7 @@ lean_object* x_60; lean_object* x_61; lean_object* x_62; x_60 = lean_ctor_get(x_59, 1); lean_inc(x_60); lean_dec(x_59); -x_61 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__37; +x_61 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__37; x_62 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_52, x_61, x_60); if (lean_obj_tag(x_62) == 0) { @@ -40643,10 +40987,10 @@ lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean 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_3063____closed__38; +x_64 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__38; x_65 = l_Lean_Parser_Command_optDeclSig___closed__2; -x_66 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__39; -x_67 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__40; +x_66 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__39; +x_67 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__40; x_68 = l_Lean_Parser_registerAlias(x_64, x_65, x_66, x_67, x_6, x_63); if (lean_obj_tag(x_68) == 0) { @@ -40654,7 +40998,7 @@ 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_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__41; +x_70 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__41; x_71 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_64, x_70, x_69); if (lean_obj_tag(x_71) == 0) { @@ -40662,7 +41006,7 @@ 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_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__42; +x_73 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__42; x_74 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_64, x_73, x_72); if (lean_obj_tag(x_74) == 0) { @@ -40670,10 +41014,10 @@ lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean 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_3063____closed__43; +x_76 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__43; x_77 = l_Lean_Parser_Command_openDecl___closed__2; -x_78 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__44; -x_79 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__45; +x_78 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__44; +x_79 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__45; x_80 = l_Lean_Parser_registerAlias(x_76, x_77, x_78, x_79, x_6, x_75); if (lean_obj_tag(x_80) == 0) { @@ -40681,7 +41025,7 @@ lean_object* x_81; lean_object* x_82; lean_object* x_83; x_81 = lean_ctor_get(x_80, 1); lean_inc(x_81); lean_dec(x_80); -x_82 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__46; +x_82 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__46; x_83 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_76, x_82, x_81); if (lean_obj_tag(x_83) == 0) { @@ -40689,7 +41033,7 @@ lean_object* x_84; lean_object* x_85; lean_object* x_86; x_84 = lean_ctor_get(x_83, 1); lean_inc(x_84); lean_dec(x_83); -x_85 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__47; +x_85 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__47; x_86 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_76, x_85, x_84); if (lean_obj_tag(x_86) == 0) { @@ -40697,10 +41041,10 @@ lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean x_87 = lean_ctor_get(x_86, 1); lean_inc(x_87); lean_dec(x_86); -x_88 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__49; -x_89 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__50; -x_90 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__51; -x_91 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__52; +x_88 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__49; +x_89 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__50; +x_90 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__51; +x_91 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__52; x_92 = l_Lean_Parser_registerAlias(x_88, x_89, x_90, x_91, x_6, x_87); if (lean_obj_tag(x_92) == 0) { @@ -40708,7 +41052,7 @@ lean_object* x_93; lean_object* x_94; lean_object* x_95; x_93 = lean_ctor_get(x_92, 1); lean_inc(x_93); lean_dec(x_92); -x_94 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__53; +x_94 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__53; x_95 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_88, x_94, x_93); if (lean_obj_tag(x_95) == 0) { @@ -40716,7 +41060,7 @@ 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_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__54; +x_97 = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__54; x_98 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_88, x_97, x_96); return x_98; } @@ -41299,7 +41643,7 @@ lean_closure_set(x_4, 0, x_3); lean_ctor_set(x_1, 1, x_4); x_5 = l_Lean_Parser_Command_openDecl; x_6 = l_Lean_Parser_andthen(x_5, x_1); -x_7 = l_Lean_Parser_Command_open___closed__5; +x_7 = l_Lean_Parser_Command_open___closed__4; x_8 = l_Lean_Parser_andthen(x_7, x_6); x_9 = l_Lean_Parser_Term_open___closed__1; x_10 = l_Lean_Parser_leadPrec; @@ -41324,7 +41668,7 @@ lean_ctor_set(x_18, 0, x_15); lean_ctor_set(x_18, 1, x_17); x_19 = l_Lean_Parser_Command_openDecl; x_20 = l_Lean_Parser_andthen(x_19, x_18); -x_21 = l_Lean_Parser_Command_open___closed__5; +x_21 = l_Lean_Parser_Command_open___closed__4; x_22 = l_Lean_Parser_andthen(x_21, x_20); x_23 = l_Lean_Parser_Term_open___closed__1; x_24 = l_Lean_Parser_leadPrec; @@ -41392,7 +41736,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_open_declRange___close { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_unsigned_to_nat(303u); -x_2 = lean_unsigned_to_nat(68u); +x_2 = lean_unsigned_to_nat(67u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -41406,7 +41750,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Parser_Term_open_declRange___closed__1; x_2 = lean_unsigned_to_nat(23u); x_3 = l___regBuiltin_Lean_Parser_Term_open_declRange___closed__2; -x_4 = lean_unsigned_to_nat(68u); +x_4 = lean_unsigned_to_nat(67u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -42252,6 +42596,23 @@ return x_5; static lean_object* _init_l_Lean_Parser_Tactic_open___closed__4() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("open ", 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_open___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Tactic_open___closed__4; +x_2 = l_Lean_Parser_symbol(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_open___closed__6() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_in___closed__4; x_2 = l_Lean_Parser_Tactic_tacticSeq; @@ -42259,11 +42620,11 @@ x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_open___closed__5() { +static lean_object* _init_l_Lean_Parser_Tactic_open___closed__7() { _start: { lean_object* x_1; uint8_t x_2; -x_1 = l_Lean_Parser_Tactic_open___closed__4; +x_1 = l_Lean_Parser_Tactic_open___closed__6; x_2 = !lean_is_exclusive(x_1); if (x_2 == 0) { @@ -42274,7 +42635,7 @@ lean_closure_set(x_4, 0, x_3); lean_ctor_set(x_1, 1, x_4); x_5 = l_Lean_Parser_Command_openDecl; x_6 = l_Lean_Parser_andthen(x_5, x_1); -x_7 = l_Lean_Parser_Command_open___closed__5; +x_7 = l_Lean_Parser_Tactic_open___closed__5; x_8 = l_Lean_Parser_andthen(x_7, x_6); x_9 = l_Lean_Parser_Tactic_open___closed__2; x_10 = l_Lean_Parser_leadPrec; @@ -42299,7 +42660,7 @@ lean_ctor_set(x_18, 0, x_15); lean_ctor_set(x_18, 1, x_17); x_19 = l_Lean_Parser_Command_openDecl; x_20 = l_Lean_Parser_andthen(x_19, x_18); -x_21 = l_Lean_Parser_Command_open___closed__5; +x_21 = l_Lean_Parser_Tactic_open___closed__5; x_22 = l_Lean_Parser_andthen(x_21, x_20); x_23 = l_Lean_Parser_Tactic_open___closed__2; x_24 = l_Lean_Parser_leadPrec; @@ -42315,7 +42676,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_open() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Tactic_open___closed__5; +x_1 = l_Lean_Parser_Tactic_open___closed__7; return x_1; } } @@ -42491,6 +42852,16 @@ return x_7; static lean_object* _init_l_Lean_Parser_Tactic_open_formatter___closed__2() { _start: { +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Tactic_open___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; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_open_formatter___closed__3() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_in_formatter___closed__1; x_2 = l_Lean_Parser_Command_decreasingBy_formatter___closed__3; @@ -42500,47 +42871,47 @@ lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_open_formatter___closed__3() { +static lean_object* _init_l_Lean_Parser_Tactic_open_formatter___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic_open_formatter___closed__2; +x_1 = l_Lean_Parser_Tactic_open_formatter___closed__3; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_withOpenDecl_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic_open_formatter___closed__4() { +static lean_object* _init_l_Lean_Parser_Tactic_open_formatter___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_open_formatter___closed__3; -x_2 = l_Lean_Parser_Tactic_open_formatter___closed__3; +x_2 = l_Lean_Parser_Tactic_open_formatter___closed__4; 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; } } -static lean_object* _init_l_Lean_Parser_Tactic_open_formatter___closed__5() { +static lean_object* _init_l_Lean_Parser_Tactic_open_formatter___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_open_formatter___closed__2; -x_2 = l_Lean_Parser_Tactic_open_formatter___closed__4; +x_1 = l_Lean_Parser_Tactic_open_formatter___closed__2; +x_2 = l_Lean_Parser_Tactic_open_formatter___closed__5; 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; } } -static lean_object* _init_l_Lean_Parser_Tactic_open_formatter___closed__6() { +static lean_object* _init_l_Lean_Parser_Tactic_open_formatter___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_open___closed__2; x_2 = l_Lean_Parser_leadPrec; -x_3 = l_Lean_Parser_Tactic_open_formatter___closed__5; +x_3 = l_Lean_Parser_Tactic_open_formatter___closed__6; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -42553,7 +42924,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_open_formatter(lean_object* x_1, l { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Tactic_open_formatter___closed__1; -x_7 = l_Lean_Parser_Tactic_open_formatter___closed__6; +x_7 = l_Lean_Parser_Tactic_open_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; } @@ -42612,6 +42983,16 @@ return x_7; static lean_object* _init_l_Lean_Parser_Tactic_open_parenthesizer___closed__2() { _start: { +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Tactic_open___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; +} +} +static lean_object* _init_l_Lean_Parser_Tactic_open_parenthesizer___closed__3() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_in_parenthesizer___closed__1; x_2 = l_Lean_Parser_Command_decreasingBy_parenthesizer___closed__3; @@ -42621,47 +43002,47 @@ lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_open_parenthesizer___closed__3() { +static lean_object* _init_l_Lean_Parser_Tactic_open_parenthesizer___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic_open_parenthesizer___closed__2; +x_1 = l_Lean_Parser_Tactic_open_parenthesizer___closed__3; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_withOpenDecl_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic_open_parenthesizer___closed__4() { +static lean_object* _init_l_Lean_Parser_Tactic_open_parenthesizer___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_open_parenthesizer___closed__3; -x_2 = l_Lean_Parser_Tactic_open_parenthesizer___closed__3; +x_2 = l_Lean_Parser_Tactic_open_parenthesizer___closed__4; 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; } } -static lean_object* _init_l_Lean_Parser_Tactic_open_parenthesizer___closed__5() { +static lean_object* _init_l_Lean_Parser_Tactic_open_parenthesizer___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_open_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Tactic_open_parenthesizer___closed__4; +x_1 = l_Lean_Parser_Tactic_open_parenthesizer___closed__2; +x_2 = l_Lean_Parser_Tactic_open_parenthesizer___closed__5; 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; } } -static lean_object* _init_l_Lean_Parser_Tactic_open_parenthesizer___closed__6() { +static lean_object* _init_l_Lean_Parser_Tactic_open_parenthesizer___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_open___closed__2; x_2 = l_Lean_Parser_leadPrec; -x_3 = l_Lean_Parser_Tactic_open_parenthesizer___closed__5; +x_3 = l_Lean_Parser_Tactic_open_parenthesizer___closed__6; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -42674,7 +43055,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_open_parenthesizer(lean_object* x_ { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Tactic_open_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Tactic_open_parenthesizer___closed__6; +x_7 = l_Lean_Parser_Tactic_open_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; } @@ -42741,7 +43122,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_set__option___closed__3() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_optionValue; -x_2 = l_Lean_Parser_Tactic_open___closed__4; +x_2 = l_Lean_Parser_Tactic_open___closed__6; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -42971,7 +43352,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_set__option_formatter___closed__2 { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_set__option_formatter___closed__3; -x_2 = l_Lean_Parser_Tactic_open_formatter___closed__2; +x_2 = l_Lean_Parser_Tactic_open_formatter___closed__3; 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); @@ -43094,7 +43475,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_set__option_parenthesizer___close { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_set__option_parenthesizer___closed__3; -x_2 = l_Lean_Parser_Tactic_open_parenthesizer___closed__2; +x_2 = l_Lean_Parser_Tactic_open_parenthesizer___closed__3; 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); @@ -43565,6 +43946,8 @@ l_Lean_Parser_Command_terminationByCore___closed__9 = _init_l_Lean_Parser_Comman lean_mark_persistent(l_Lean_Parser_Command_terminationByCore___closed__9); l_Lean_Parser_Command_terminationByCore___closed__10 = _init_l_Lean_Parser_Command_terminationByCore___closed__10(); lean_mark_persistent(l_Lean_Parser_Command_terminationByCore___closed__10); +l_Lean_Parser_Command_terminationByCore___closed__11 = _init_l_Lean_Parser_Command_terminationByCore___closed__11(); +lean_mark_persistent(l_Lean_Parser_Command_terminationByCore___closed__11); l_Lean_Parser_Command_terminationByCore = _init_l_Lean_Parser_Command_terminationByCore(); lean_mark_persistent(l_Lean_Parser_Command_terminationByCore); l_Lean_Parser_Command_decreasingBy___closed__1 = _init_l_Lean_Parser_Command_decreasingBy___closed__1(); @@ -43587,6 +43970,8 @@ l_Lean_Parser_Command_decreasingBy___closed__9 = _init_l_Lean_Parser_Command_dec lean_mark_persistent(l_Lean_Parser_Command_decreasingBy___closed__9); l_Lean_Parser_Command_decreasingBy___closed__10 = _init_l_Lean_Parser_Command_decreasingBy___closed__10(); lean_mark_persistent(l_Lean_Parser_Command_decreasingBy___closed__10); +l_Lean_Parser_Command_decreasingBy___closed__11 = _init_l_Lean_Parser_Command_decreasingBy___closed__11(); +lean_mark_persistent(l_Lean_Parser_Command_decreasingBy___closed__11); l_Lean_Parser_Command_decreasingBy = _init_l_Lean_Parser_Command_decreasingBy(); lean_mark_persistent(l_Lean_Parser_Command_decreasingBy); l_Lean_Parser_Command_terminationByElement___closed__1 = _init_l_Lean_Parser_Command_terminationByElement___closed__1(); @@ -43615,6 +44000,8 @@ l_Lean_Parser_Command_terminationByElement___closed__12 = _init_l_Lean_Parser_Co lean_mark_persistent(l_Lean_Parser_Command_terminationByElement___closed__12); l_Lean_Parser_Command_terminationByElement___closed__13 = _init_l_Lean_Parser_Command_terminationByElement___closed__13(); lean_mark_persistent(l_Lean_Parser_Command_terminationByElement___closed__13); +l_Lean_Parser_Command_terminationByElement___closed__14 = _init_l_Lean_Parser_Command_terminationByElement___closed__14(); +lean_mark_persistent(l_Lean_Parser_Command_terminationByElement___closed__14); l_Lean_Parser_Command_terminationByElement = _init_l_Lean_Parser_Command_terminationByElement(); lean_mark_persistent(l_Lean_Parser_Command_terminationByElement); l_Lean_Parser_Command_terminationBy___closed__1 = _init_l_Lean_Parser_Command_terminationBy___closed__1(); @@ -43789,8 +44176,6 @@ l_Lean_Parser_Command_namedPrio = _init_l_Lean_Parser_Command_namedPrio(); lean_mark_persistent(l_Lean_Parser_Command_namedPrio); l_Lean_Parser_Command_optNamedPrio___closed__1 = _init_l_Lean_Parser_Command_optNamedPrio___closed__1(); lean_mark_persistent(l_Lean_Parser_Command_optNamedPrio___closed__1); -l_Lean_Parser_Command_optNamedPrio___closed__2 = _init_l_Lean_Parser_Command_optNamedPrio___closed__2(); -lean_mark_persistent(l_Lean_Parser_Command_optNamedPrio___closed__2); l_Lean_Parser_Command_optNamedPrio = _init_l_Lean_Parser_Command_optNamedPrio(); lean_mark_persistent(l_Lean_Parser_Command_optNamedPrio); l_Lean_Parser_Command_private___closed__1 = _init_l_Lean_Parser_Command_private___closed__1(); @@ -44185,6 +44570,8 @@ l_Lean_Parser_Command_optDefDeriving___closed__8 = _init_l_Lean_Parser_Command_o lean_mark_persistent(l_Lean_Parser_Command_optDefDeriving___closed__8); l_Lean_Parser_Command_optDefDeriving___closed__9 = _init_l_Lean_Parser_Command_optDefDeriving___closed__9(); lean_mark_persistent(l_Lean_Parser_Command_optDefDeriving___closed__9); +l_Lean_Parser_Command_optDefDeriving___closed__10 = _init_l_Lean_Parser_Command_optDefDeriving___closed__10(); +lean_mark_persistent(l_Lean_Parser_Command_optDefDeriving___closed__10); l_Lean_Parser_Command_optDefDeriving = _init_l_Lean_Parser_Command_optDefDeriving(); lean_mark_persistent(l_Lean_Parser_Command_optDefDeriving); l_Lean_Parser_Command_def___closed__1 = _init_l_Lean_Parser_Command_def___closed__1(); @@ -44563,6 +44950,10 @@ l_Lean_Parser_Command_structExplicitBinder___closed__15 = _init_l_Lean_Parser_Co lean_mark_persistent(l_Lean_Parser_Command_structExplicitBinder___closed__15); l_Lean_Parser_Command_structExplicitBinder___closed__16 = _init_l_Lean_Parser_Command_structExplicitBinder___closed__16(); lean_mark_persistent(l_Lean_Parser_Command_structExplicitBinder___closed__16); +l_Lean_Parser_Command_structExplicitBinder___closed__17 = _init_l_Lean_Parser_Command_structExplicitBinder___closed__17(); +lean_mark_persistent(l_Lean_Parser_Command_structExplicitBinder___closed__17); +l_Lean_Parser_Command_structExplicitBinder___closed__18 = _init_l_Lean_Parser_Command_structExplicitBinder___closed__18(); +lean_mark_persistent(l_Lean_Parser_Command_structExplicitBinder___closed__18); l_Lean_Parser_Command_structExplicitBinder = _init_l_Lean_Parser_Command_structExplicitBinder(); lean_mark_persistent(l_Lean_Parser_Command_structExplicitBinder); l_Lean_Parser_Command_structImplicitBinder___closed__1 = _init_l_Lean_Parser_Command_structImplicitBinder___closed__1(); @@ -45155,6 +45546,8 @@ l_Lean_Parser_Command_terminationByElement_formatter___closed__12 = _init_l_Lean lean_mark_persistent(l_Lean_Parser_Command_terminationByElement_formatter___closed__12); l_Lean_Parser_Command_terminationByElement_formatter___closed__13 = _init_l_Lean_Parser_Command_terminationByElement_formatter___closed__13(); lean_mark_persistent(l_Lean_Parser_Command_terminationByElement_formatter___closed__13); +l_Lean_Parser_Command_terminationByElement_formatter___closed__14 = _init_l_Lean_Parser_Command_terminationByElement_formatter___closed__14(); +lean_mark_persistent(l_Lean_Parser_Command_terminationByElement_formatter___closed__14); l___regBuiltin_Lean_Parser_Command_terminationByElement_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_terminationByElement_formatter___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_terminationByElement_formatter___closed__1); l___regBuiltin_Lean_Parser_Command_terminationByElement_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_terminationByElement_formatter___closed__2(); @@ -45201,6 +45594,8 @@ l_Lean_Parser_Command_terminationByCore_formatter___closed__4 = _init_l_Lean_Par lean_mark_persistent(l_Lean_Parser_Command_terminationByCore_formatter___closed__4); l_Lean_Parser_Command_terminationByCore_formatter___closed__5 = _init_l_Lean_Parser_Command_terminationByCore_formatter___closed__5(); lean_mark_persistent(l_Lean_Parser_Command_terminationByCore_formatter___closed__5); +l_Lean_Parser_Command_terminationByCore_formatter___closed__6 = _init_l_Lean_Parser_Command_terminationByCore_formatter___closed__6(); +lean_mark_persistent(l_Lean_Parser_Command_terminationByCore_formatter___closed__6); l___regBuiltin_Lean_Parser_Command_terminationByCore_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_terminationByCore_formatter___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_terminationByCore_formatter___closed__1); l___regBuiltin_Lean_Parser_Command_terminationByCore_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_terminationByCore_formatter___closed__2(); @@ -45220,6 +45615,8 @@ l_Lean_Parser_Command_decreasingBy_formatter___closed__5 = _init_l_Lean_Parser_C lean_mark_persistent(l_Lean_Parser_Command_decreasingBy_formatter___closed__5); l_Lean_Parser_Command_decreasingBy_formatter___closed__6 = _init_l_Lean_Parser_Command_decreasingBy_formatter___closed__6(); lean_mark_persistent(l_Lean_Parser_Command_decreasingBy_formatter___closed__6); +l_Lean_Parser_Command_decreasingBy_formatter___closed__7 = _init_l_Lean_Parser_Command_decreasingBy_formatter___closed__7(); +lean_mark_persistent(l_Lean_Parser_Command_decreasingBy_formatter___closed__7); l___regBuiltin_Lean_Parser_Command_decreasingBy_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_decreasingBy_formatter___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_decreasingBy_formatter___closed__1); l___regBuiltin_Lean_Parser_Command_decreasingBy_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_decreasingBy_formatter___closed__2(); @@ -45270,6 +45667,8 @@ l_Lean_Parser_Command_optDefDeriving_formatter___closed__4 = _init_l_Lean_Parser lean_mark_persistent(l_Lean_Parser_Command_optDefDeriving_formatter___closed__4); l_Lean_Parser_Command_optDefDeriving_formatter___closed__5 = _init_l_Lean_Parser_Command_optDefDeriving_formatter___closed__5(); lean_mark_persistent(l_Lean_Parser_Command_optDefDeriving_formatter___closed__5); +l_Lean_Parser_Command_optDefDeriving_formatter___closed__6 = _init_l_Lean_Parser_Command_optDefDeriving_formatter___closed__6(); +lean_mark_persistent(l_Lean_Parser_Command_optDefDeriving_formatter___closed__6); l_Lean_Parser_Command_def_formatter___closed__1 = _init_l_Lean_Parser_Command_def_formatter___closed__1(); lean_mark_persistent(l_Lean_Parser_Command_def_formatter___closed__1); l_Lean_Parser_Command_def_formatter___closed__2 = _init_l_Lean_Parser_Command_def_formatter___closed__2(); @@ -45383,8 +45782,6 @@ lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_namedPrio_formatter___cl res = l___regBuiltin_Lean_Parser_Command_namedPrio_formatter(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l_Lean_Parser_Command_optNamedPrio_formatter___closed__1 = _init_l_Lean_Parser_Command_optNamedPrio_formatter___closed__1(); -lean_mark_persistent(l_Lean_Parser_Command_optNamedPrio_formatter___closed__1); l_Lean_Parser_Command_instance_formatter___closed__1 = _init_l_Lean_Parser_Command_instance_formatter___closed__1(); lean_mark_persistent(l_Lean_Parser_Command_instance_formatter___closed__1); l_Lean_Parser_Command_instance_formatter___closed__2 = _init_l_Lean_Parser_Command_instance_formatter___closed__2(); @@ -45533,6 +45930,8 @@ l_Lean_Parser_Command_derivingClasses_formatter___closed__5 = _init_l_Lean_Parse lean_mark_persistent(l_Lean_Parser_Command_derivingClasses_formatter___closed__5); l_Lean_Parser_Command_derivingClasses_formatter___closed__6 = _init_l_Lean_Parser_Command_derivingClasses_formatter___closed__6(); lean_mark_persistent(l_Lean_Parser_Command_derivingClasses_formatter___closed__6); +l_Lean_Parser_Command_derivingClasses_formatter___closed__7 = _init_l_Lean_Parser_Command_derivingClasses_formatter___closed__7(); +lean_mark_persistent(l_Lean_Parser_Command_derivingClasses_formatter___closed__7); l_Lean_Parser_Command_optDeriving_formatter___closed__1 = _init_l_Lean_Parser_Command_optDeriving_formatter___closed__1(); lean_mark_persistent(l_Lean_Parser_Command_optDeriving_formatter___closed__1); l_Lean_Parser_Command_optDeriving_formatter___closed__2 = _init_l_Lean_Parser_Command_optDeriving_formatter___closed__2(); @@ -45671,6 +46070,8 @@ l_Lean_Parser_Command_structCtor_formatter___closed__5 = _init_l_Lean_Parser_Com lean_mark_persistent(l_Lean_Parser_Command_structCtor_formatter___closed__5); l_Lean_Parser_Command_structCtor_formatter___closed__6 = _init_l_Lean_Parser_Command_structCtor_formatter___closed__6(); lean_mark_persistent(l_Lean_Parser_Command_structCtor_formatter___closed__6); +l_Lean_Parser_Command_structCtor_formatter___closed__7 = _init_l_Lean_Parser_Command_structCtor_formatter___closed__7(); +lean_mark_persistent(l_Lean_Parser_Command_structCtor_formatter___closed__7); l___regBuiltin_Lean_Parser_Command_structCtor_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_structCtor_formatter___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_structCtor_formatter___closed__1); l___regBuiltin_Lean_Parser_Command_structCtor_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_structCtor_formatter___closed__2(); @@ -45706,6 +46107,8 @@ l_Lean_Parser_Command_structExplicitBinder_formatter___closed__13 = _init_l_Lean lean_mark_persistent(l_Lean_Parser_Command_structExplicitBinder_formatter___closed__13); l_Lean_Parser_Command_structExplicitBinder_formatter___closed__14 = _init_l_Lean_Parser_Command_structExplicitBinder_formatter___closed__14(); lean_mark_persistent(l_Lean_Parser_Command_structExplicitBinder_formatter___closed__14); +l_Lean_Parser_Command_structExplicitBinder_formatter___closed__15 = _init_l_Lean_Parser_Command_structExplicitBinder_formatter___closed__15(); +lean_mark_persistent(l_Lean_Parser_Command_structExplicitBinder_formatter___closed__15); l___regBuiltin_Lean_Parser_Command_structExplicitBinder_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_structExplicitBinder_formatter___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_structExplicitBinder_formatter___closed__1); l___regBuiltin_Lean_Parser_Command_structExplicitBinder_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_structExplicitBinder_formatter___closed__2(); @@ -45843,6 +46246,8 @@ l_Lean_Parser_Command_structure_formatter___closed__17 = _init_l_Lean_Parser_Com lean_mark_persistent(l_Lean_Parser_Command_structure_formatter___closed__17); l_Lean_Parser_Command_structure_formatter___closed__18 = _init_l_Lean_Parser_Command_structure_formatter___closed__18(); lean_mark_persistent(l_Lean_Parser_Command_structure_formatter___closed__18); +l_Lean_Parser_Command_structure_formatter___closed__19 = _init_l_Lean_Parser_Command_structure_formatter___closed__19(); +lean_mark_persistent(l_Lean_Parser_Command_structure_formatter___closed__19); l___regBuiltin_Lean_Parser_Command_structure_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_structure_formatter___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_structure_formatter___closed__1); l___regBuiltin_Lean_Parser_Command_structure_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_structure_formatter___closed__2(); @@ -46177,6 +46582,8 @@ l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__12 = _init_l_ lean_mark_persistent(l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__12); l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__13 = _init_l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__13(); lean_mark_persistent(l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__13); +l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__14 = _init_l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__14(); +lean_mark_persistent(l_Lean_Parser_Command_terminationByElement_parenthesizer___closed__14); l___regBuiltin_Lean_Parser_Command_terminationByElement_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_terminationByElement_parenthesizer___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_terminationByElement_parenthesizer___closed__1); l___regBuiltin_Lean_Parser_Command_terminationByElement_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_terminationByElement_parenthesizer___closed__2(); @@ -46223,6 +46630,8 @@ l_Lean_Parser_Command_terminationByCore_parenthesizer___closed__4 = _init_l_Lean lean_mark_persistent(l_Lean_Parser_Command_terminationByCore_parenthesizer___closed__4); l_Lean_Parser_Command_terminationByCore_parenthesizer___closed__5 = _init_l_Lean_Parser_Command_terminationByCore_parenthesizer___closed__5(); lean_mark_persistent(l_Lean_Parser_Command_terminationByCore_parenthesizer___closed__5); +l_Lean_Parser_Command_terminationByCore_parenthesizer___closed__6 = _init_l_Lean_Parser_Command_terminationByCore_parenthesizer___closed__6(); +lean_mark_persistent(l_Lean_Parser_Command_terminationByCore_parenthesizer___closed__6); l___regBuiltin_Lean_Parser_Command_terminationByCore_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_terminationByCore_parenthesizer___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_terminationByCore_parenthesizer___closed__1); l___regBuiltin_Lean_Parser_Command_terminationByCore_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_terminationByCore_parenthesizer___closed__2(); @@ -46242,6 +46651,8 @@ l_Lean_Parser_Command_decreasingBy_parenthesizer___closed__5 = _init_l_Lean_Pars lean_mark_persistent(l_Lean_Parser_Command_decreasingBy_parenthesizer___closed__5); l_Lean_Parser_Command_decreasingBy_parenthesizer___closed__6 = _init_l_Lean_Parser_Command_decreasingBy_parenthesizer___closed__6(); lean_mark_persistent(l_Lean_Parser_Command_decreasingBy_parenthesizer___closed__6); +l_Lean_Parser_Command_decreasingBy_parenthesizer___closed__7 = _init_l_Lean_Parser_Command_decreasingBy_parenthesizer___closed__7(); +lean_mark_persistent(l_Lean_Parser_Command_decreasingBy_parenthesizer___closed__7); l___regBuiltin_Lean_Parser_Command_decreasingBy_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_decreasingBy_parenthesizer___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_decreasingBy_parenthesizer___closed__1); l___regBuiltin_Lean_Parser_Command_decreasingBy_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_decreasingBy_parenthesizer___closed__2(); @@ -46292,6 +46703,8 @@ l_Lean_Parser_Command_optDefDeriving_parenthesizer___closed__4 = _init_l_Lean_Pa lean_mark_persistent(l_Lean_Parser_Command_optDefDeriving_parenthesizer___closed__4); l_Lean_Parser_Command_optDefDeriving_parenthesizer___closed__5 = _init_l_Lean_Parser_Command_optDefDeriving_parenthesizer___closed__5(); lean_mark_persistent(l_Lean_Parser_Command_optDefDeriving_parenthesizer___closed__5); +l_Lean_Parser_Command_optDefDeriving_parenthesizer___closed__6 = _init_l_Lean_Parser_Command_optDefDeriving_parenthesizer___closed__6(); +lean_mark_persistent(l_Lean_Parser_Command_optDefDeriving_parenthesizer___closed__6); l_Lean_Parser_Command_def_parenthesizer___closed__1 = _init_l_Lean_Parser_Command_def_parenthesizer___closed__1(); lean_mark_persistent(l_Lean_Parser_Command_def_parenthesizer___closed__1); l_Lean_Parser_Command_def_parenthesizer___closed__2 = _init_l_Lean_Parser_Command_def_parenthesizer___closed__2(); @@ -46405,8 +46818,6 @@ lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_namedPrio_parenthesizer_ res = l___regBuiltin_Lean_Parser_Command_namedPrio_parenthesizer(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l_Lean_Parser_Command_optNamedPrio_parenthesizer___closed__1 = _init_l_Lean_Parser_Command_optNamedPrio_parenthesizer___closed__1(); -lean_mark_persistent(l_Lean_Parser_Command_optNamedPrio_parenthesizer___closed__1); l_Lean_Parser_Command_instance_parenthesizer___closed__1 = _init_l_Lean_Parser_Command_instance_parenthesizer___closed__1(); lean_mark_persistent(l_Lean_Parser_Command_instance_parenthesizer___closed__1); l_Lean_Parser_Command_instance_parenthesizer___closed__2 = _init_l_Lean_Parser_Command_instance_parenthesizer___closed__2(); @@ -46555,6 +46966,8 @@ l_Lean_Parser_Command_derivingClasses_parenthesizer___closed__5 = _init_l_Lean_P lean_mark_persistent(l_Lean_Parser_Command_derivingClasses_parenthesizer___closed__5); l_Lean_Parser_Command_derivingClasses_parenthesizer___closed__6 = _init_l_Lean_Parser_Command_derivingClasses_parenthesizer___closed__6(); lean_mark_persistent(l_Lean_Parser_Command_derivingClasses_parenthesizer___closed__6); +l_Lean_Parser_Command_derivingClasses_parenthesizer___closed__7 = _init_l_Lean_Parser_Command_derivingClasses_parenthesizer___closed__7(); +lean_mark_persistent(l_Lean_Parser_Command_derivingClasses_parenthesizer___closed__7); l_Lean_Parser_Command_optDeriving_parenthesizer___closed__1 = _init_l_Lean_Parser_Command_optDeriving_parenthesizer___closed__1(); lean_mark_persistent(l_Lean_Parser_Command_optDeriving_parenthesizer___closed__1); l_Lean_Parser_Command_optDeriving_parenthesizer___closed__2 = _init_l_Lean_Parser_Command_optDeriving_parenthesizer___closed__2(); @@ -46693,6 +47106,8 @@ l_Lean_Parser_Command_structCtor_parenthesizer___closed__5 = _init_l_Lean_Parser lean_mark_persistent(l_Lean_Parser_Command_structCtor_parenthesizer___closed__5); l_Lean_Parser_Command_structCtor_parenthesizer___closed__6 = _init_l_Lean_Parser_Command_structCtor_parenthesizer___closed__6(); lean_mark_persistent(l_Lean_Parser_Command_structCtor_parenthesizer___closed__6); +l_Lean_Parser_Command_structCtor_parenthesizer___closed__7 = _init_l_Lean_Parser_Command_structCtor_parenthesizer___closed__7(); +lean_mark_persistent(l_Lean_Parser_Command_structCtor_parenthesizer___closed__7); l___regBuiltin_Lean_Parser_Command_structCtor_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_structCtor_parenthesizer___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_structCtor_parenthesizer___closed__1); l___regBuiltin_Lean_Parser_Command_structCtor_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_structCtor_parenthesizer___closed__2(); @@ -46728,6 +47143,8 @@ l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__13 = _init_l_ lean_mark_persistent(l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__13); l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__14 = _init_l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__14(); lean_mark_persistent(l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__14); +l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__15 = _init_l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__15(); +lean_mark_persistent(l_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__15); l___regBuiltin_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__1); l___regBuiltin_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_structExplicitBinder_parenthesizer___closed__2(); @@ -46865,6 +47282,8 @@ l_Lean_Parser_Command_structure_parenthesizer___closed__17 = _init_l_Lean_Parser lean_mark_persistent(l_Lean_Parser_Command_structure_parenthesizer___closed__17); l_Lean_Parser_Command_structure_parenthesizer___closed__18 = _init_l_Lean_Parser_Command_structure_parenthesizer___closed__18(); lean_mark_persistent(l_Lean_Parser_Command_structure_parenthesizer___closed__18); +l_Lean_Parser_Command_structure_parenthesizer___closed__19 = _init_l_Lean_Parser_Command_structure_parenthesizer___closed__19(); +lean_mark_persistent(l_Lean_Parser_Command_structure_parenthesizer___closed__19); l___regBuiltin_Lean_Parser_Command_structure_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_structure_parenthesizer___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_structure_parenthesizer___closed__1); l___regBuiltin_Lean_Parser_Command_structure_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_structure_parenthesizer___closed__2(); @@ -47023,6 +47442,8 @@ l_Lean_Parser_Command_noncomputableSection___closed__10 = _init_l_Lean_Parser_Co lean_mark_persistent(l_Lean_Parser_Command_noncomputableSection___closed__10); l_Lean_Parser_Command_noncomputableSection___closed__11 = _init_l_Lean_Parser_Command_noncomputableSection___closed__11(); lean_mark_persistent(l_Lean_Parser_Command_noncomputableSection___closed__11); +l_Lean_Parser_Command_noncomputableSection___closed__12 = _init_l_Lean_Parser_Command_noncomputableSection___closed__12(); +lean_mark_persistent(l_Lean_Parser_Command_noncomputableSection___closed__12); l_Lean_Parser_Command_noncomputableSection = _init_l_Lean_Parser_Command_noncomputableSection(); lean_mark_persistent(l_Lean_Parser_Command_noncomputableSection); res = l___regBuiltin_Lean_Parser_Command_noncomputableSection(lean_io_mk_world()); @@ -47057,6 +47478,8 @@ l_Lean_Parser_Command_noncomputableSection_formatter___closed__5 = _init_l_Lean_ lean_mark_persistent(l_Lean_Parser_Command_noncomputableSection_formatter___closed__5); l_Lean_Parser_Command_noncomputableSection_formatter___closed__6 = _init_l_Lean_Parser_Command_noncomputableSection_formatter___closed__6(); lean_mark_persistent(l_Lean_Parser_Command_noncomputableSection_formatter___closed__6); +l_Lean_Parser_Command_noncomputableSection_formatter___closed__7 = _init_l_Lean_Parser_Command_noncomputableSection_formatter___closed__7(); +lean_mark_persistent(l_Lean_Parser_Command_noncomputableSection_formatter___closed__7); l___regBuiltin_Lean_Parser_Command_noncomputableSection_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_noncomputableSection_formatter___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_noncomputableSection_formatter___closed__1); l___regBuiltin_Lean_Parser_Command_noncomputableSection_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_noncomputableSection_formatter___closed__2(); @@ -47076,6 +47499,8 @@ l_Lean_Parser_Command_noncomputableSection_parenthesizer___closed__5 = _init_l_L lean_mark_persistent(l_Lean_Parser_Command_noncomputableSection_parenthesizer___closed__5); l_Lean_Parser_Command_noncomputableSection_parenthesizer___closed__6 = _init_l_Lean_Parser_Command_noncomputableSection_parenthesizer___closed__6(); lean_mark_persistent(l_Lean_Parser_Command_noncomputableSection_parenthesizer___closed__6); +l_Lean_Parser_Command_noncomputableSection_parenthesizer___closed__7 = _init_l_Lean_Parser_Command_noncomputableSection_parenthesizer___closed__7(); +lean_mark_persistent(l_Lean_Parser_Command_noncomputableSection_parenthesizer___closed__7); l___regBuiltin_Lean_Parser_Command_noncomputableSection_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_noncomputableSection_parenthesizer___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_noncomputableSection_parenthesizer___closed__1); l___regBuiltin_Lean_Parser_Command_noncomputableSection_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_noncomputableSection_parenthesizer___closed__2(); @@ -47093,8 +47518,6 @@ l_Lean_Parser_Command_section___closed__4 = _init_l_Lean_Parser_Command_section_ lean_mark_persistent(l_Lean_Parser_Command_section___closed__4); l_Lean_Parser_Command_section___closed__5 = _init_l_Lean_Parser_Command_section___closed__5(); lean_mark_persistent(l_Lean_Parser_Command_section___closed__5); -l_Lean_Parser_Command_section___closed__6 = _init_l_Lean_Parser_Command_section___closed__6(); -lean_mark_persistent(l_Lean_Parser_Command_section___closed__6); l_Lean_Parser_Command_section = _init_l_Lean_Parser_Command_section(); lean_mark_persistent(l_Lean_Parser_Command_section); res = l___regBuiltin_Lean_Parser_Command_section(lean_io_mk_world()); @@ -47225,8 +47648,6 @@ l_Lean_Parser_Command_end___closed__7 = _init_l_Lean_Parser_Command_end___closed lean_mark_persistent(l_Lean_Parser_Command_end___closed__7); l_Lean_Parser_Command_end___closed__8 = _init_l_Lean_Parser_Command_end___closed__8(); lean_mark_persistent(l_Lean_Parser_Command_end___closed__8); -l_Lean_Parser_Command_end___closed__9 = _init_l_Lean_Parser_Command_end___closed__9(); -lean_mark_persistent(l_Lean_Parser_Command_end___closed__9); l_Lean_Parser_Command_end = _init_l_Lean_Parser_Command_end(); lean_mark_persistent(l_Lean_Parser_Command_end); res = l___regBuiltin_Lean_Parser_Command_end(lean_io_mk_world()); @@ -47401,6 +47822,8 @@ l_Lean_Parser_Command_universe_formatter___closed__3 = _init_l_Lean_Parser_Comma lean_mark_persistent(l_Lean_Parser_Command_universe_formatter___closed__3); l_Lean_Parser_Command_universe_formatter___closed__4 = _init_l_Lean_Parser_Command_universe_formatter___closed__4(); lean_mark_persistent(l_Lean_Parser_Command_universe_formatter___closed__4); +l_Lean_Parser_Command_universe_formatter___closed__5 = _init_l_Lean_Parser_Command_universe_formatter___closed__5(); +lean_mark_persistent(l_Lean_Parser_Command_universe_formatter___closed__5); l___regBuiltin_Lean_Parser_Command_universe_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_universe_formatter___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_universe_formatter___closed__1); l___regBuiltin_Lean_Parser_Command_universe_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_universe_formatter___closed__2(); @@ -47416,6 +47839,8 @@ l_Lean_Parser_Command_universe_parenthesizer___closed__3 = _init_l_Lean_Parser_C lean_mark_persistent(l_Lean_Parser_Command_universe_parenthesizer___closed__3); l_Lean_Parser_Command_universe_parenthesizer___closed__4 = _init_l_Lean_Parser_Command_universe_parenthesizer___closed__4(); lean_mark_persistent(l_Lean_Parser_Command_universe_parenthesizer___closed__4); +l_Lean_Parser_Command_universe_parenthesizer___closed__5 = _init_l_Lean_Parser_Command_universe_parenthesizer___closed__5(); +lean_mark_persistent(l_Lean_Parser_Command_universe_parenthesizer___closed__5); l___regBuiltin_Lean_Parser_Command_universe_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_universe_parenthesizer___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_universe_parenthesizer___closed__1); l___regBuiltin_Lean_Parser_Command_universe_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_universe_parenthesizer___closed__2(); @@ -48227,10 +48652,6 @@ l_Lean_Parser_Command_attribute___closed__14 = _init_l_Lean_Parser_Command_attri lean_mark_persistent(l_Lean_Parser_Command_attribute___closed__14); l_Lean_Parser_Command_attribute___closed__15 = _init_l_Lean_Parser_Command_attribute___closed__15(); lean_mark_persistent(l_Lean_Parser_Command_attribute___closed__15); -l_Lean_Parser_Command_attribute___closed__16 = _init_l_Lean_Parser_Command_attribute___closed__16(); -lean_mark_persistent(l_Lean_Parser_Command_attribute___closed__16); -l_Lean_Parser_Command_attribute___closed__17 = _init_l_Lean_Parser_Command_attribute___closed__17(); -lean_mark_persistent(l_Lean_Parser_Command_attribute___closed__17); l_Lean_Parser_Command_attribute = _init_l_Lean_Parser_Command_attribute(); lean_mark_persistent(l_Lean_Parser_Command_attribute); res = l___regBuiltin_Lean_Parser_Command_attribute(lean_io_mk_world()); @@ -48290,8 +48711,6 @@ l_Lean_Parser_Command_attribute_formatter___closed__10 = _init_l_Lean_Parser_Com lean_mark_persistent(l_Lean_Parser_Command_attribute_formatter___closed__10); l_Lean_Parser_Command_attribute_formatter___closed__11 = _init_l_Lean_Parser_Command_attribute_formatter___closed__11(); lean_mark_persistent(l_Lean_Parser_Command_attribute_formatter___closed__11); -l_Lean_Parser_Command_attribute_formatter___closed__12 = _init_l_Lean_Parser_Command_attribute_formatter___closed__12(); -lean_mark_persistent(l_Lean_Parser_Command_attribute_formatter___closed__12); l___regBuiltin_Lean_Parser_Command_attribute_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_attribute_formatter___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_attribute_formatter___closed__1); l___regBuiltin_Lean_Parser_Command_attribute_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_attribute_formatter___closed__2(); @@ -48336,8 +48755,6 @@ l_Lean_Parser_Command_attribute_parenthesizer___closed__10 = _init_l_Lean_Parser lean_mark_persistent(l_Lean_Parser_Command_attribute_parenthesizer___closed__10); l_Lean_Parser_Command_attribute_parenthesizer___closed__11 = _init_l_Lean_Parser_Command_attribute_parenthesizer___closed__11(); lean_mark_persistent(l_Lean_Parser_Command_attribute_parenthesizer___closed__11); -l_Lean_Parser_Command_attribute_parenthesizer___closed__12 = _init_l_Lean_Parser_Command_attribute_parenthesizer___closed__12(); -lean_mark_persistent(l_Lean_Parser_Command_attribute_parenthesizer___closed__12); l___regBuiltin_Lean_Parser_Command_attribute_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_attribute_parenthesizer___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_attribute_parenthesizer___closed__1); l___regBuiltin_Lean_Parser_Command_attribute_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_attribute_parenthesizer___closed__2(); @@ -48369,10 +48786,6 @@ l_Lean_Parser_Command_export___closed__11 = _init_l_Lean_Parser_Command_export__ lean_mark_persistent(l_Lean_Parser_Command_export___closed__11); l_Lean_Parser_Command_export___closed__12 = _init_l_Lean_Parser_Command_export___closed__12(); lean_mark_persistent(l_Lean_Parser_Command_export___closed__12); -l_Lean_Parser_Command_export___closed__13 = _init_l_Lean_Parser_Command_export___closed__13(); -lean_mark_persistent(l_Lean_Parser_Command_export___closed__13); -l_Lean_Parser_Command_export___closed__14 = _init_l_Lean_Parser_Command_export___closed__14(); -lean_mark_persistent(l_Lean_Parser_Command_export___closed__14); l_Lean_Parser_Command_export = _init_l_Lean_Parser_Command_export(); lean_mark_persistent(l_Lean_Parser_Command_export); res = l___regBuiltin_Lean_Parser_Command_export(lean_io_mk_world()); @@ -48409,8 +48822,6 @@ l_Lean_Parser_Command_export_formatter___closed__6 = _init_l_Lean_Parser_Command lean_mark_persistent(l_Lean_Parser_Command_export_formatter___closed__6); l_Lean_Parser_Command_export_formatter___closed__7 = _init_l_Lean_Parser_Command_export_formatter___closed__7(); lean_mark_persistent(l_Lean_Parser_Command_export_formatter___closed__7); -l_Lean_Parser_Command_export_formatter___closed__8 = _init_l_Lean_Parser_Command_export_formatter___closed__8(); -lean_mark_persistent(l_Lean_Parser_Command_export_formatter___closed__8); l___regBuiltin_Lean_Parser_Command_export_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_export_formatter___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_export_formatter___closed__1); l___regBuiltin_Lean_Parser_Command_export_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_export_formatter___closed__2(); @@ -48432,8 +48843,6 @@ l_Lean_Parser_Command_export_parenthesizer___closed__6 = _init_l_Lean_Parser_Com lean_mark_persistent(l_Lean_Parser_Command_export_parenthesizer___closed__6); l_Lean_Parser_Command_export_parenthesizer___closed__7 = _init_l_Lean_Parser_Command_export_parenthesizer___closed__7(); lean_mark_persistent(l_Lean_Parser_Command_export_parenthesizer___closed__7); -l_Lean_Parser_Command_export_parenthesizer___closed__8 = _init_l_Lean_Parser_Command_export_parenthesizer___closed__8(); -lean_mark_persistent(l_Lean_Parser_Command_export_parenthesizer___closed__8); l___regBuiltin_Lean_Parser_Command_export_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_export_parenthesizer___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_export_parenthesizer___closed__1); l___regBuiltin_Lean_Parser_Command_export_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_export_parenthesizer___closed__2(); @@ -48533,6 +48942,10 @@ l_Lean_Parser_Command_openHiding___closed__14 = _init_l_Lean_Parser_Command_open lean_mark_persistent(l_Lean_Parser_Command_openHiding___closed__14); l_Lean_Parser_Command_openHiding___closed__15 = _init_l_Lean_Parser_Command_openHiding___closed__15(); lean_mark_persistent(l_Lean_Parser_Command_openHiding___closed__15); +l_Lean_Parser_Command_openHiding___closed__16 = _init_l_Lean_Parser_Command_openHiding___closed__16(); +lean_mark_persistent(l_Lean_Parser_Command_openHiding___closed__16); +l_Lean_Parser_Command_openHiding___closed__17 = _init_l_Lean_Parser_Command_openHiding___closed__17(); +lean_mark_persistent(l_Lean_Parser_Command_openHiding___closed__17); l_Lean_Parser_Command_openHiding = _init_l_Lean_Parser_Command_openHiding(); lean_mark_persistent(l_Lean_Parser_Command_openHiding); l_Lean_Parser_Command_openRenamingItem___closed__1 = _init_l_Lean_Parser_Command_openRenamingItem___closed__1(); @@ -48583,6 +48996,8 @@ l_Lean_Parser_Command_openRenaming___closed__11 = _init_l_Lean_Parser_Command_op lean_mark_persistent(l_Lean_Parser_Command_openRenaming___closed__11); l_Lean_Parser_Command_openRenaming___closed__12 = _init_l_Lean_Parser_Command_openRenaming___closed__12(); lean_mark_persistent(l_Lean_Parser_Command_openRenaming___closed__12); +l_Lean_Parser_Command_openRenaming___closed__13 = _init_l_Lean_Parser_Command_openRenaming___closed__13(); +lean_mark_persistent(l_Lean_Parser_Command_openRenaming___closed__13); l_Lean_Parser_Command_openRenaming = _init_l_Lean_Parser_Command_openRenaming(); lean_mark_persistent(l_Lean_Parser_Command_openRenaming); l_Lean_Parser_Command_openOnly___closed__1 = _init_l_Lean_Parser_Command_openOnly___closed__1(); @@ -48603,6 +49018,8 @@ l_Lean_Parser_Command_openOnly___closed__8 = _init_l_Lean_Parser_Command_openOnl lean_mark_persistent(l_Lean_Parser_Command_openOnly___closed__8); l_Lean_Parser_Command_openOnly___closed__9 = _init_l_Lean_Parser_Command_openOnly___closed__9(); lean_mark_persistent(l_Lean_Parser_Command_openOnly___closed__9); +l_Lean_Parser_Command_openOnly___closed__10 = _init_l_Lean_Parser_Command_openOnly___closed__10(); +lean_mark_persistent(l_Lean_Parser_Command_openOnly___closed__10); l_Lean_Parser_Command_openOnly = _init_l_Lean_Parser_Command_openOnly(); lean_mark_persistent(l_Lean_Parser_Command_openOnly); l_Lean_Parser_Command_openSimple___closed__1 = _init_l_Lean_Parser_Command_openSimple___closed__1(); @@ -48675,8 +49092,6 @@ l_Lean_Parser_Command_open___closed__8 = _init_l_Lean_Parser_Command_open___clos lean_mark_persistent(l_Lean_Parser_Command_open___closed__8); l_Lean_Parser_Command_open___closed__9 = _init_l_Lean_Parser_Command_open___closed__9(); lean_mark_persistent(l_Lean_Parser_Command_open___closed__9); -l_Lean_Parser_Command_open___closed__10 = _init_l_Lean_Parser_Command_open___closed__10(); -lean_mark_persistent(l_Lean_Parser_Command_open___closed__10); l_Lean_Parser_Command_open = _init_l_Lean_Parser_Command_open(); lean_mark_persistent(l_Lean_Parser_Command_open); res = l___regBuiltin_Lean_Parser_Command_open(lean_io_mk_world()); @@ -48717,6 +49132,10 @@ l_Lean_Parser_Command_openHiding_formatter___closed__8 = _init_l_Lean_Parser_Com lean_mark_persistent(l_Lean_Parser_Command_openHiding_formatter___closed__8); l_Lean_Parser_Command_openHiding_formatter___closed__9 = _init_l_Lean_Parser_Command_openHiding_formatter___closed__9(); lean_mark_persistent(l_Lean_Parser_Command_openHiding_formatter___closed__9); +l_Lean_Parser_Command_openHiding_formatter___closed__10 = _init_l_Lean_Parser_Command_openHiding_formatter___closed__10(); +lean_mark_persistent(l_Lean_Parser_Command_openHiding_formatter___closed__10); +l_Lean_Parser_Command_openHiding_formatter___closed__11 = _init_l_Lean_Parser_Command_openHiding_formatter___closed__11(); +lean_mark_persistent(l_Lean_Parser_Command_openHiding_formatter___closed__11); l___regBuiltin_Lean_Parser_Command_openHiding_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_openHiding_formatter___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_openHiding_formatter___closed__1); l___regBuiltin_Lean_Parser_Command_openHiding_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_openHiding_formatter___closed__2(); @@ -48755,6 +49174,8 @@ l_Lean_Parser_Command_openRenaming_formatter___closed__6 = _init_l_Lean_Parser_C lean_mark_persistent(l_Lean_Parser_Command_openRenaming_formatter___closed__6); l_Lean_Parser_Command_openRenaming_formatter___closed__7 = _init_l_Lean_Parser_Command_openRenaming_formatter___closed__7(); lean_mark_persistent(l_Lean_Parser_Command_openRenaming_formatter___closed__7); +l_Lean_Parser_Command_openRenaming_formatter___closed__8 = _init_l_Lean_Parser_Command_openRenaming_formatter___closed__8(); +lean_mark_persistent(l_Lean_Parser_Command_openRenaming_formatter___closed__8); l___regBuiltin_Lean_Parser_Command_openRenaming_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_openRenaming_formatter___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_openRenaming_formatter___closed__1); l___regBuiltin_Lean_Parser_Command_openRenaming_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_openRenaming_formatter___closed__2(); @@ -48772,6 +49193,8 @@ l_Lean_Parser_Command_openOnly_formatter___closed__4 = _init_l_Lean_Parser_Comma lean_mark_persistent(l_Lean_Parser_Command_openOnly_formatter___closed__4); l_Lean_Parser_Command_openOnly_formatter___closed__5 = _init_l_Lean_Parser_Command_openOnly_formatter___closed__5(); lean_mark_persistent(l_Lean_Parser_Command_openOnly_formatter___closed__5); +l_Lean_Parser_Command_openOnly_formatter___closed__6 = _init_l_Lean_Parser_Command_openOnly_formatter___closed__6(); +lean_mark_persistent(l_Lean_Parser_Command_openOnly_formatter___closed__6); l___regBuiltin_Lean_Parser_Command_openOnly_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_openOnly_formatter___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_openOnly_formatter___closed__1); l___regBuiltin_Lean_Parser_Command_openOnly_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_openOnly_formatter___closed__2(); @@ -48852,6 +49275,10 @@ l_Lean_Parser_Command_openHiding_parenthesizer___closed__8 = _init_l_Lean_Parser lean_mark_persistent(l_Lean_Parser_Command_openHiding_parenthesizer___closed__8); l_Lean_Parser_Command_openHiding_parenthesizer___closed__9 = _init_l_Lean_Parser_Command_openHiding_parenthesizer___closed__9(); lean_mark_persistent(l_Lean_Parser_Command_openHiding_parenthesizer___closed__9); +l_Lean_Parser_Command_openHiding_parenthesizer___closed__10 = _init_l_Lean_Parser_Command_openHiding_parenthesizer___closed__10(); +lean_mark_persistent(l_Lean_Parser_Command_openHiding_parenthesizer___closed__10); +l_Lean_Parser_Command_openHiding_parenthesizer___closed__11 = _init_l_Lean_Parser_Command_openHiding_parenthesizer___closed__11(); +lean_mark_persistent(l_Lean_Parser_Command_openHiding_parenthesizer___closed__11); l___regBuiltin_Lean_Parser_Command_openHiding_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_openHiding_parenthesizer___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_openHiding_parenthesizer___closed__1); l___regBuiltin_Lean_Parser_Command_openHiding_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_openHiding_parenthesizer___closed__2(); @@ -48890,6 +49317,8 @@ l_Lean_Parser_Command_openRenaming_parenthesizer___closed__6 = _init_l_Lean_Pars lean_mark_persistent(l_Lean_Parser_Command_openRenaming_parenthesizer___closed__6); l_Lean_Parser_Command_openRenaming_parenthesizer___closed__7 = _init_l_Lean_Parser_Command_openRenaming_parenthesizer___closed__7(); lean_mark_persistent(l_Lean_Parser_Command_openRenaming_parenthesizer___closed__7); +l_Lean_Parser_Command_openRenaming_parenthesizer___closed__8 = _init_l_Lean_Parser_Command_openRenaming_parenthesizer___closed__8(); +lean_mark_persistent(l_Lean_Parser_Command_openRenaming_parenthesizer___closed__8); l___regBuiltin_Lean_Parser_Command_openRenaming_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_openRenaming_parenthesizer___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_openRenaming_parenthesizer___closed__1); l___regBuiltin_Lean_Parser_Command_openRenaming_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_openRenaming_parenthesizer___closed__2(); @@ -48907,6 +49336,8 @@ l_Lean_Parser_Command_openOnly_parenthesizer___closed__4 = _init_l_Lean_Parser_C lean_mark_persistent(l_Lean_Parser_Command_openOnly_parenthesizer___closed__4); l_Lean_Parser_Command_openOnly_parenthesizer___closed__5 = _init_l_Lean_Parser_Command_openOnly_parenthesizer___closed__5(); lean_mark_persistent(l_Lean_Parser_Command_openOnly_parenthesizer___closed__5); +l_Lean_Parser_Command_openOnly_parenthesizer___closed__6 = _init_l_Lean_Parser_Command_openOnly_parenthesizer___closed__6(); +lean_mark_persistent(l_Lean_Parser_Command_openOnly_parenthesizer___closed__6); l___regBuiltin_Lean_Parser_Command_openOnly_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_openOnly_parenthesizer___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_openOnly_parenthesizer___closed__1); l___regBuiltin_Lean_Parser_Command_openOnly_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_openOnly_parenthesizer___closed__2(); @@ -48999,10 +49430,6 @@ l_Lean_Parser_Command_mutual___closed__14 = _init_l_Lean_Parser_Command_mutual__ lean_mark_persistent(l_Lean_Parser_Command_mutual___closed__14); l_Lean_Parser_Command_mutual___closed__15 = _init_l_Lean_Parser_Command_mutual___closed__15(); lean_mark_persistent(l_Lean_Parser_Command_mutual___closed__15); -l_Lean_Parser_Command_mutual___closed__16 = _init_l_Lean_Parser_Command_mutual___closed__16(); -lean_mark_persistent(l_Lean_Parser_Command_mutual___closed__16); -l_Lean_Parser_Command_mutual___closed__17 = _init_l_Lean_Parser_Command_mutual___closed__17(); -lean_mark_persistent(l_Lean_Parser_Command_mutual___closed__17); l_Lean_Parser_Command_mutual = _init_l_Lean_Parser_Command_mutual(); lean_mark_persistent(l_Lean_Parser_Command_mutual); res = l___regBuiltin_Lean_Parser_Command_mutual(lean_io_mk_world()); @@ -49049,8 +49476,6 @@ l_Lean_Parser_Command_mutual_formatter___closed__11 = _init_l_Lean_Parser_Comman lean_mark_persistent(l_Lean_Parser_Command_mutual_formatter___closed__11); l_Lean_Parser_Command_mutual_formatter___closed__12 = _init_l_Lean_Parser_Command_mutual_formatter___closed__12(); lean_mark_persistent(l_Lean_Parser_Command_mutual_formatter___closed__12); -l_Lean_Parser_Command_mutual_formatter___closed__13 = _init_l_Lean_Parser_Command_mutual_formatter___closed__13(); -lean_mark_persistent(l_Lean_Parser_Command_mutual_formatter___closed__13); l___regBuiltin_Lean_Parser_Command_mutual_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_mutual_formatter___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_mutual_formatter___closed__1); l___regBuiltin_Lean_Parser_Command_mutual_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_mutual_formatter___closed__2(); @@ -49082,8 +49507,6 @@ l_Lean_Parser_Command_mutual_parenthesizer___closed__11 = _init_l_Lean_Parser_Co lean_mark_persistent(l_Lean_Parser_Command_mutual_parenthesizer___closed__11); l_Lean_Parser_Command_mutual_parenthesizer___closed__12 = _init_l_Lean_Parser_Command_mutual_parenthesizer___closed__12(); lean_mark_persistent(l_Lean_Parser_Command_mutual_parenthesizer___closed__12); -l_Lean_Parser_Command_mutual_parenthesizer___closed__13 = _init_l_Lean_Parser_Command_mutual_parenthesizer___closed__13(); -lean_mark_persistent(l_Lean_Parser_Command_mutual_parenthesizer___closed__13); l___regBuiltin_Lean_Parser_Command_mutual_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_mutual_parenthesizer___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_mutual_parenthesizer___closed__1); l___regBuiltin_Lean_Parser_Command_mutual_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_mutual_parenthesizer___closed__2(); @@ -49141,6 +49564,8 @@ l_Lean_Parser_Command_initialize___closed__12 = _init_l_Lean_Parser_Command_init lean_mark_persistent(l_Lean_Parser_Command_initialize___closed__12); l_Lean_Parser_Command_initialize___closed__13 = _init_l_Lean_Parser_Command_initialize___closed__13(); lean_mark_persistent(l_Lean_Parser_Command_initialize___closed__13); +l_Lean_Parser_Command_initialize___closed__14 = _init_l_Lean_Parser_Command_initialize___closed__14(); +lean_mark_persistent(l_Lean_Parser_Command_initialize___closed__14); l_Lean_Parser_Command_initialize = _init_l_Lean_Parser_Command_initialize(); lean_mark_persistent(l_Lean_Parser_Command_initialize); res = l___regBuiltin_Lean_Parser_Command_initialize(lean_io_mk_world()); @@ -49202,6 +49627,8 @@ l_Lean_Parser_Command_initialize_formatter___closed__10 = _init_l_Lean_Parser_Co lean_mark_persistent(l_Lean_Parser_Command_initialize_formatter___closed__10); l_Lean_Parser_Command_initialize_formatter___closed__11 = _init_l_Lean_Parser_Command_initialize_formatter___closed__11(); lean_mark_persistent(l_Lean_Parser_Command_initialize_formatter___closed__11); +l_Lean_Parser_Command_initialize_formatter___closed__12 = _init_l_Lean_Parser_Command_initialize_formatter___closed__12(); +lean_mark_persistent(l_Lean_Parser_Command_initialize_formatter___closed__12); l___regBuiltin_Lean_Parser_Command_initialize_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_initialize_formatter___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_initialize_formatter___closed__1); l___regBuiltin_Lean_Parser_Command_initialize_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_initialize_formatter___closed__2(); @@ -49248,6 +49675,8 @@ l_Lean_Parser_Command_initialize_parenthesizer___closed__10 = _init_l_Lean_Parse lean_mark_persistent(l_Lean_Parser_Command_initialize_parenthesizer___closed__10); l_Lean_Parser_Command_initialize_parenthesizer___closed__11 = _init_l_Lean_Parser_Command_initialize_parenthesizer___closed__11(); lean_mark_persistent(l_Lean_Parser_Command_initialize_parenthesizer___closed__11); +l_Lean_Parser_Command_initialize_parenthesizer___closed__12 = _init_l_Lean_Parser_Command_initialize_parenthesizer___closed__12(); +lean_mark_persistent(l_Lean_Parser_Command_initialize_parenthesizer___closed__12); l___regBuiltin_Lean_Parser_Command_initialize_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_initialize_parenthesizer___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_initialize_parenthesizer___closed__1); l___regBuiltin_Lean_Parser_Command_initialize_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_initialize_parenthesizer___closed__2(); @@ -49295,6 +49724,8 @@ l_Lean_Parser_Command_in_formatter___closed__2 = _init_l_Lean_Parser_Command_in_ lean_mark_persistent(l_Lean_Parser_Command_in_formatter___closed__2); l_Lean_Parser_Command_in_formatter___closed__3 = _init_l_Lean_Parser_Command_in_formatter___closed__3(); lean_mark_persistent(l_Lean_Parser_Command_in_formatter___closed__3); +l_Lean_Parser_Command_in_formatter___closed__4 = _init_l_Lean_Parser_Command_in_formatter___closed__4(); +lean_mark_persistent(l_Lean_Parser_Command_in_formatter___closed__4); l___regBuiltin_Lean_Parser_Command_in_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_in_formatter___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_in_formatter___closed__1); l___regBuiltin_Lean_Parser_Command_in_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_in_formatter___closed__2(); @@ -49308,6 +49739,8 @@ l_Lean_Parser_Command_in_parenthesizer___closed__2 = _init_l_Lean_Parser_Command lean_mark_persistent(l_Lean_Parser_Command_in_parenthesizer___closed__2); l_Lean_Parser_Command_in_parenthesizer___closed__3 = _init_l_Lean_Parser_Command_in_parenthesizer___closed__3(); lean_mark_persistent(l_Lean_Parser_Command_in_parenthesizer___closed__3); +l_Lean_Parser_Command_in_parenthesizer___closed__4 = _init_l_Lean_Parser_Command_in_parenthesizer___closed__4(); +lean_mark_persistent(l_Lean_Parser_Command_in_parenthesizer___closed__4); l___regBuiltin_Lean_Parser_Command_in_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_in_parenthesizer___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_in_parenthesizer___closed__1); l___regBuiltin_Lean_Parser_Command_in_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_in_parenthesizer___closed__2(); @@ -49510,122 +49943,122 @@ l_Lean_Parser_Command_eoi___closed__5 = _init_l_Lean_Parser_Command_eoi___closed lean_mark_persistent(l_Lean_Parser_Command_eoi___closed__5); l_Lean_Parser_Command_eoi = _init_l_Lean_Parser_Command_eoi(); lean_mark_persistent(l_Lean_Parser_Command_eoi); -res = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3040_(lean_io_mk_world()); +res = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3138_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); l_Lean_Parser_Command_declModifiersF = _init_l_Lean_Parser_Command_declModifiersF(); 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_3063____closed__1 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__1(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__1); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__2 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__2(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__2); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__3 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__3(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__3); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__4 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__4(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__4); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__5 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__5(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__5); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__6 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__6(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__6); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__7 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__7(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__7); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__8 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__8(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__8); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__9 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__9(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__9); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__10 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__10(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__10); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__11 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__11(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__11); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__12 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__12(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__12); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__13 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__13(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__13); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__14 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__14(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__14); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__15 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__15(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__15); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__16 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__16(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__16); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__17 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__17(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__17); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__18 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__18(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__18); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__19 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__19(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__19); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__20 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__20(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__20); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__21 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__21(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__21); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__22 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__22(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__22); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__23 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__23(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__23); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__24 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__24(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__24); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__25 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__25(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__25); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__26 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__26(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__26); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__27 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__27(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__27); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__28 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__28(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__28); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__29 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__29(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__29); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__30 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__30(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__30); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__31 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__31(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__31); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__32 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__32(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__32); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__33 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__33(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__33); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__34 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__34(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__34); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__35 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__35(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__35); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__36 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__36(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__36); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__37 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__37(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__37); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__38 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__38(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__38); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__39 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__39(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__39); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__40 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__40(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__40); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__41 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__41(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__41); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__42 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__42(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__42); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__43 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__43(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__43); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__44 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__44(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__44); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__45 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__45(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__45); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__46 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__46(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__46); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__47 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__47(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__47); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__48 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__48(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__48); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__49 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__49(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__49); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__50 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__50(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__50); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__51 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__51(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__51); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__52 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__52(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__52); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__53 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__53(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__53); -l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__54 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__54(); -lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063____closed__54); -res = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3063_(lean_io_mk_world()); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__1 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__1(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__1); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__2 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__2(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__2); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__3 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__3(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__3); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__4 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__4(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__4); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__5 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__5(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__5); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__6 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__6(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__6); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__7 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__7(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__7); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__8 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__8(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__8); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__9 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__9(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__9); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__10 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__10(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__10); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__11 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__11(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__11); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__12 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__12(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__12); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__13 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__13(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__13); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__14 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__14(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__14); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__15 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__15(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__15); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__16 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__16(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__16); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__17 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__17(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__17); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__18 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__18(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__18); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__19 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__19(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__19); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__20 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__20(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__20); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__21 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__21(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__21); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__22 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__22(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__22); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__23 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__23(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__23); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__24 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__24(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__24); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__25 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__25(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__25); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__26 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__26(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__26); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__27 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__27(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__27); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__28 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__28(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__28); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__29 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__29(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__29); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__30 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__30(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__30); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__31 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__31(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__31); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__32 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__32(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__32); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__33 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__33(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__33); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__34 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__34(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__34); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__35 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__35(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__35); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__36 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__36(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__36); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__37 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__37(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__37); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__38 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__38(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__38); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__39 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__39(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__39); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__40 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__40(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__40); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__41 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__41(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__41); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__42 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__42(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__42); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__43 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__43(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__43); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__44 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__44(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__44); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__45 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__45(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__45); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__46 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__46(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__46); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__47 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__47(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__47); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__48 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__48(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__48); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__49 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__49(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__49); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__50 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__50(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__50); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__51 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__51(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__51); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__52 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__52(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__52); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__53 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__53(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__53); +l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__54 = _init_l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__54(); +lean_mark_persistent(l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161____closed__54); +res = l_Lean_Parser_Command_initFn____x40_Lean_Parser_Command___hyg_3161_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); l_Lean_Parser_Term_open___closed__1 = _init_l_Lean_Parser_Term_open___closed__1(); @@ -49798,6 +50231,10 @@ l_Lean_Parser_Tactic_open___closed__4 = _init_l_Lean_Parser_Tactic_open___closed lean_mark_persistent(l_Lean_Parser_Tactic_open___closed__4); l_Lean_Parser_Tactic_open___closed__5 = _init_l_Lean_Parser_Tactic_open___closed__5(); lean_mark_persistent(l_Lean_Parser_Tactic_open___closed__5); +l_Lean_Parser_Tactic_open___closed__6 = _init_l_Lean_Parser_Tactic_open___closed__6(); +lean_mark_persistent(l_Lean_Parser_Tactic_open___closed__6); +l_Lean_Parser_Tactic_open___closed__7 = _init_l_Lean_Parser_Tactic_open___closed__7(); +lean_mark_persistent(l_Lean_Parser_Tactic_open___closed__7); l_Lean_Parser_Tactic_open = _init_l_Lean_Parser_Tactic_open(); lean_mark_persistent(l_Lean_Parser_Tactic_open); l___regBuiltin_Lean_Parser_Tactic_open___closed__1 = _init_l___regBuiltin_Lean_Parser_Tactic_open___closed__1(); @@ -49841,6 +50278,8 @@ l_Lean_Parser_Tactic_open_formatter___closed__5 = _init_l_Lean_Parser_Tactic_ope lean_mark_persistent(l_Lean_Parser_Tactic_open_formatter___closed__5); l_Lean_Parser_Tactic_open_formatter___closed__6 = _init_l_Lean_Parser_Tactic_open_formatter___closed__6(); lean_mark_persistent(l_Lean_Parser_Tactic_open_formatter___closed__6); +l_Lean_Parser_Tactic_open_formatter___closed__7 = _init_l_Lean_Parser_Tactic_open_formatter___closed__7(); +lean_mark_persistent(l_Lean_Parser_Tactic_open_formatter___closed__7); l___regBuiltin_Lean_Parser_Tactic_open_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Tactic_open_formatter___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Tactic_open_formatter___closed__1); l___regBuiltin_Lean_Parser_Tactic_open_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Tactic_open_formatter___closed__2(); @@ -49860,6 +50299,8 @@ l_Lean_Parser_Tactic_open_parenthesizer___closed__5 = _init_l_Lean_Parser_Tactic lean_mark_persistent(l_Lean_Parser_Tactic_open_parenthesizer___closed__5); l_Lean_Parser_Tactic_open_parenthesizer___closed__6 = _init_l_Lean_Parser_Tactic_open_parenthesizer___closed__6(); lean_mark_persistent(l_Lean_Parser_Tactic_open_parenthesizer___closed__6); +l_Lean_Parser_Tactic_open_parenthesizer___closed__7 = _init_l_Lean_Parser_Tactic_open_parenthesizer___closed__7(); +lean_mark_persistent(l_Lean_Parser_Tactic_open_parenthesizer___closed__7); l___regBuiltin_Lean_Parser_Tactic_open_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Tactic_open_parenthesizer___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Tactic_open_parenthesizer___closed__1); l___regBuiltin_Lean_Parser_Tactic_open_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Tactic_open_parenthesizer___closed__2(); diff --git a/stage0/stdlib/Lean/Parser/Do.c b/stage0/stdlib/Lean/Parser/Do.c index 90d3b9f2c7e5..06c9cb4f581f 100644 --- a/stage0/stdlib/Lean/Parser/Do.c +++ b/stage0/stdlib/Lean/Parser/Do.c @@ -104,6 +104,7 @@ static lean_object* l_Lean_Parser_Term_doNested_parenthesizer___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_termTry_declRange___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Term_doIfLet; static lean_object* l_Lean_Parser_Term_liftMethod_formatter___closed__5; +static lean_object* l_Lean_Parser_Term_doCatch_parenthesizer___closed__13; lean_object* l_Lean_addBuiltinDocString(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_doLetArrow___closed__1; 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*); @@ -146,6 +147,7 @@ static lean_object* l_Lean_Parser_Term_doLetRec___closed__9; static lean_object* l___regBuiltin_Lean_Parser_Term_termReturn_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Term_termReturn_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_termFor___closed__3; +static lean_object* l_Lean_Parser_Term_doCatch_parenthesizer___closed__14; lean_object* l_Lean_PrettyPrinter_Formatter_orelse_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_notFollowedByRedefinedTermToken___closed__16; static lean_object* l_Lean_Parser_Term_doIfProp___closed__5; @@ -225,7 +227,6 @@ static lean_object* l_Lean_Parser_Term_doAssert___closed__8; static lean_object* l___regBuiltin_Lean_Parser_Term_liftMethod___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_doAssert_declRange___closed__1; lean_object* l_Lean_PrettyPrinter_Formatter_interpolatedStr_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Term_doHave___closed__9; static lean_object* l_Lean_Parser_Term_doLet___closed__9; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_termReturn_formatter(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_doIdDecl_parenthesizer___closed__2; @@ -297,6 +298,7 @@ static lean_object* l_Lean_Parser_Term_doIf_formatter___closed__13; static lean_object* l_Lean_Parser_Term_termTry___closed__6; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_doAssert_formatter(lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_doReassignArrow_declRange(lean_object*); +static lean_object* l_Lean_Parser_Term_doFor_parenthesizer___closed__9; static lean_object* l_Lean_Parser_Term_doSeqBracketed___closed__14; static lean_object* l_Lean_Parser_Term_doLetArrow_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Term_doIf_formatter___closed__31; @@ -309,6 +311,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_doBreak; static lean_object* l_Lean_Parser_Term_doExpr_parenthesizer___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Term_doReassign; static lean_object* l_Lean_Parser_Term_doLet___closed__7; +static lean_object* l_Lean_Parser_Term_doFor_formatter___closed__8; static lean_object* l___regBuiltin_Lean_Parser_Term_doBreak_declRange___closed__6; static lean_object* l_Lean_Parser_Term_doForDecl_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Term_doIf_formatter___closed__11; @@ -359,6 +362,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_termTry_parenthesizer(lean_object*, static lean_object* l_Lean_Parser_Term_doSeqIndent___closed__1; static lean_object* l_Lean_Parser_Term_doMatch___closed__6; static lean_object* l_Lean_Parser_Term_doForDecl_formatter___closed__2; +static lean_object* l_Lean_Parser_Term_doFor___closed__14; static lean_object* l___regBuiltin_Lean_Parser_Term_doAssert_declRange___closed__3; static lean_object* l_Lean_Parser_Term_doIf_parenthesizer___closed__27; LEAN_EXPORT lean_object* l_Lean_Parser_Term_doSeqIndent; @@ -450,6 +454,7 @@ static lean_object* l_Lean_Parser_Term_doLetArrow___closed__7; lean_object* l_Lean_Parser_Term_motive_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_doReassign___closed__6; lean_object* l_Lean_Parser_termParser_formatter___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Term_doCatch___closed__16; static lean_object* l___regBuiltin_Lean_Parser_Term_doLetRec_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_doMatch___closed__13; static lean_object* l_Lean_Parser_Term_doLet_parenthesizer___closed__5; @@ -486,6 +491,7 @@ static lean_object* l_Lean_Parser_Term_doExpr___closed__8; static lean_object* l___regBuiltin_Lean_Parser_Term_doReturn_formatter___closed__1; static lean_object* l_Lean_Parser_Term_doDbgTrace___closed__11; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_doNested_parenthesizer(lean_object*); +static lean_object* l_Lean_Parser_Term_doCatchMatch_parenthesizer___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Term_doNested; static lean_object* l_Lean_Parser_Term_doMatch_formatter___closed__4; static lean_object* l_Lean_Parser_Term_doReassignArrow___closed__2; @@ -801,6 +807,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_do_declRange___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Term_liftMethod_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_doIfLetBind_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_doIfLet_formatter___closed__1; +static lean_object* l_Lean_Parser_Term_doFor_parenthesizer___closed__8; static lean_object* l___regBuiltin_Lean_Parser_Term_doBreak_declRange___closed__2; static lean_object* l_Lean_Parser_Term_doTry_formatter___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_doIfLet_parenthesizer___closed__1; @@ -826,6 +833,7 @@ static lean_object* l_Lean_Parser_Term_doForDecl_parenthesizer___closed__9; LEAN_EXPORT lean_object* l_Lean_Parser_Term_notFollowedByRedefinedTermToken_parenthesizer___rarg(lean_object*); static lean_object* l_Lean_Parser_Term_doIfLetPure___closed__2; static lean_object* l_Lean_Parser_Term_doMatch_formatter___closed__2; +static lean_object* l_Lean_Parser_Term_doCatch_formatter___closed__14; static lean_object* l_Lean_Parser_Term_doIf_formatter___closed__12; static lean_object* l_Lean_Parser_Term_doIf___closed__32; static lean_object* l_Lean_Parser_Term_doTry___closed__9; @@ -866,6 +874,7 @@ static lean_object* l_Lean_Parser_Term_doExpr_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Term_doMatch_formatter___closed__13; lean_object* l_Lean_Parser_sepBy1(lean_object*, lean_object*, lean_object*, uint8_t); static lean_object* l_Lean_Parser_Term_doIdDecl_formatter___closed__5; +static lean_object* l_Lean_Parser_Term_doCatch_formatter___closed__13; lean_object* l_Lean_Parser_notFollowedBy(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_doFor___closed__7; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_doReassign(lean_object*); @@ -896,6 +905,7 @@ static lean_object* l_Lean_Parser_Term_doLetElse___closed__14; static lean_object* l_Lean_Parser_Term_doIf_formatter___closed__3; lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_doIfProp_formatter___closed__1; +static lean_object* l_Lean_Parser_Term_termFor_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Term_doLetElse___closed__13; static lean_object* l___regBuiltin_Lean_Parser_Term_do_formatter___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_termUnless_declRange___closed__4; @@ -953,6 +963,7 @@ static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Do___hyg_21____close static lean_object* l_Lean_Parser_Term_doUnless_formatter___closed__2; static lean_object* l_Lean_Parser_Term_doIdDecl_formatter___closed__6; static lean_object* l_Lean_Parser_Term_doFor_formatter___closed__6; +static lean_object* l_Lean_Parser_Term_doFor___closed__15; static lean_object* l___regBuiltin_Lean_Parser_Term_doHave_formatter___closed__2; static lean_object* l_Lean_Parser_Term_doIfLetPure___closed__4; static lean_object* l_Lean_Parser_Term_doIf_formatter___closed__33; @@ -1032,6 +1043,7 @@ static lean_object* l_Lean_Parser_Term_doIf_formatter___closed__24; static lean_object* l_Lean_Parser_Term_doIfLetBind___closed__2; static lean_object* l_Lean_Parser_Term_doIf_parenthesizer___closed__13; static lean_object* l___regBuiltin_Lean_Parser_Term_termReturn_declRange___closed__5; +static lean_object* l_Lean_Parser_Term_termFor___closed__8; static lean_object* l___regBuiltin_Lean_Parser_Term_doReassign_formatter___closed__2; static lean_object* l_Lean_Parser_Term_doTry_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Term_doAssert___closed__5; @@ -1060,6 +1072,7 @@ static lean_object* l_Lean_Parser_Term_doTry___closed__12; LEAN_EXPORT lean_object* l_Lean_Parser_doElemParser(lean_object*); lean_object* l_Lean_Parser_Term_ident_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_doCatchMatch_formatter___closed__1; +static lean_object* l_Lean_Parser_Term_doFinally___closed__10; static lean_object* l_Lean_Parser_Term_doSeqBracketed_parenthesizer___closed__8; static lean_object* l_Lean_Parser_Term_doFinally_parenthesizer___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Term_doTry_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1143,7 +1156,6 @@ static lean_object* l_Lean_Parser_Term_doReturn___closed__2; static lean_object* l_Lean_Parser_Term_doCatchMatch___closed__3; static lean_object* l_Lean_Parser_Term_doIf_parenthesizer___closed__31; 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_Term_doHave___closed__8; static lean_object* l_Lean_Parser_Term_doMatch_formatter___closed__7; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_doReturn_declRange(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_doLetRec_declRange___closed__2; @@ -1257,6 +1269,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_doIf_declRange___closed__5; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_letIdDeclNoBinders_formatter(lean_object*); lean_object* l_Lean_Parser_withPosition_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_doForDecl___closed__2; +static lean_object* l_Lean_Parser_Term_termFor_parenthesizer___closed__4; 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_Lean_Parser_Term_doExpr___closed__18; static lean_object* l_Lean_Parser_Term_doContinue_formatter___closed__2; @@ -1280,6 +1293,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_doForDecl_formatter(lea static lean_object* l_Lean_Parser_Term_doFinally___closed__4; static lean_object* l_Lean_Parser_Term_doIfLetBind___closed__3; static lean_object* l_Lean_Parser_Term_doLetRec___closed__7; +static lean_object* l_Lean_Parser_Term_doReturn_formatter___closed__8; static lean_object* l___regBuiltin_Lean_Parser_Term_doHave_declRange___closed__6; static lean_object* l_Lean_Parser_Term_doCatch_parenthesizer___closed__5; lean_object* l_Lean_Parser_optional_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1291,6 +1305,7 @@ static lean_object* l_Lean_Parser_Term_doIf_parenthesizer___closed__26; static lean_object* l___regBuiltin_Lean_Parser_Term_termFor_declRange___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Term_doMatch_declRange___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_doContinue_declRange(lean_object*); +static lean_object* l_Lean_Parser_Term_doFinally_parenthesizer___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Term_liftMethod_formatter___closed__1; static lean_object* l_Lean_Parser_Term_doCatch_parenthesizer___closed__8; static lean_object* l_Lean_Parser_Term_doForDecl___closed__13; @@ -1361,6 +1376,7 @@ static lean_object* l_Lean_Parser_Term_doFor___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Term_liftMethod_parenthesizer___closed__1; lean_object* l_Lean_Parser_withCache(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_notFollowedByRedefinedTermToken___closed__17; +static lean_object* l_Lean_Parser_Term_doReturn_parenthesizer___closed__8; static lean_object* l_Lean_Parser_Term_termFor_formatter___closed__2; static lean_object* l_Lean_Parser_Term_doBreak_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Term_doIfProp_formatter___closed__4; @@ -1375,6 +1391,7 @@ static lean_object* l_Lean_Parser_Term_doCatch_parenthesizer___closed__11; static lean_object* l___regBuiltin_Lean_Parser_Term_doReturn_docString___closed__1; lean_object* l_Lean_Parser_atomic_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Term_doAssert_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Term_doFor_formatter___closed__9; static lean_object* l_Lean_Parser_Term_doBreak___closed__6; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_doReassign_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_Term_doSeqBracketed___closed__13; @@ -1414,6 +1431,7 @@ static lean_object* l_Lean_Parser_Term_doBreak___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_doContinue_declRange___closed__2; 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_Lean_Parser_Term_doIfCond___closed__1; +static lean_object* l_Lean_Parser_Term_doFor___closed__16; static lean_object* l_Lean_Parser_Term_doReturn_formatter___closed__3; static lean_object* l_Lean_Parser_Term_letIdDeclNoBinders_formatter___closed__5; static lean_object* l_Lean_Parser_Term_doIf_parenthesizer___closed__10; @@ -1482,6 +1500,7 @@ static lean_object* l_Lean_Parser_Term_doFinally___closed__1; static lean_object* l_Lean_Parser_Term_doIf_parenthesizer___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Term_doIf_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_doHave_parenthesizer___closed__2; +static lean_object* l_Lean_Parser_Term_termFor___closed__7; static lean_object* l_Lean_Parser_Term_do___closed__2; static lean_object* l_Lean_Parser_Term_doIf___closed__30; static lean_object* l___regBuiltin_Lean_Parser_Term_doFor_formatter___closed__1; @@ -1503,6 +1522,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_doMatch_formatter___closed__ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_doForDecl_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_Term_doUnless_parenthesizer___closed__7; 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_Term_termFor_formatter___closed__3; static lean_object* l_Lean_Parser_Term_doIf_formatter___closed__30; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_doContinue(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_doLetArrow_declRange___closed__1; @@ -1649,6 +1669,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_doExpr_declRange___closed__6 extern lean_object* l_Lean_PrettyPrinter_formatterAttribute; lean_object* l_Lean_ppIndent_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Term_doFinally_formatter___closed__5; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Do___hyg_21____closed__10; LEAN_EXPORT lean_object* l_Lean_Parser_Term_doFinally; static lean_object* l_Lean_Parser_Term_doLet_formatter___closed__8; @@ -1727,6 +1748,7 @@ lean_object* l_Lean_PrettyPrinter_Formatter_categoryParser_formatter(lean_object static lean_object* l_Lean_Parser_Term_doSeqItem___closed__7; static lean_object* l_Lean_Parser_Term_doCatchMatch___closed__1; static lean_object* l_Lean_Parser_Term_doForDecl___closed__10; +static lean_object* l_Lean_Parser_Term_termFor_formatter___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Term_doReassign_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_letIdDeclNoBinders___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_doLetArrow_declRange(lean_object*); @@ -1773,7 +1795,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_doAssert_declRange___closed_ static lean_object* l_Lean_Parser_Term_doForDecl_formatter___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Term_notFollowedByRedefinedTermToken_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Parser_Term_letPatDecl; -static lean_object* l_Lean_Parser_Term_doReturn___closed__12; static lean_object* l_Lean_Parser_Term_doMatch___closed__17; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Do___hyg_21____closed__16; static lean_object* l_Lean_Parser_Term_doPatDecl_parenthesizer___closed__6; @@ -1790,6 +1811,7 @@ static lean_object* l_Lean_Parser_Term_elseIf_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Term_doIf_formatter___closed__16; lean_object* l_Lean_Parser_registerBuiltinDynamicParserAttribute(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_doFinally___closed__8; +static lean_object* l_Lean_Parser_Term_doCatchMatch___closed__8; static lean_object* l___regBuiltin_Lean_Parser_Term_doExpr_declRange___closed__4; static lean_object* l_Lean_Parser_Term_doIf_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Term_doIfLet_formatter___closed__3; @@ -1835,6 +1857,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_doElemParser_formatter___boxed(lean_objec static lean_object* l_Lean_Parser_Term_doIf_formatter___closed__19; static lean_object* l_Lean_Parser_Term_doDbgTrace_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Term_doCatch___closed__15; +static lean_object* l_Lean_Parser_Term_doCatchMatch_formatter___closed__4; static lean_object* l_Lean_Parser_Term_do_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Term_doPatDecl_parenthesizer___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_doFor_docString(lean_object*); @@ -8951,57 +8974,40 @@ return x_5; static lean_object* _init_l_Lean_Parser_Term_doHave___closed__4() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("have ", 5); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Term_doHave___closed__5() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_doHave___closed__4; -x_2 = l_Lean_Parser_symbol(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Parser_Term_doHave___closed__6() { -_start: -{ lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_doHave___closed__5; +x_1 = l_Lean_Parser_Term_notFollowedByRedefinedTermToken___closed__12; x_2 = l_Lean_Parser_Term_haveDecl; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_doHave___closed__7() { +static lean_object* _init_l_Lean_Parser_Term_doHave___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_doHave___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_doHave___closed__6; +x_3 = l_Lean_Parser_Term_doHave___closed__4; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Term_doHave___closed__8() { +static lean_object* _init_l_Lean_Parser_Term_doHave___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_doHave___closed__3; -x_2 = l_Lean_Parser_Term_doHave___closed__7; +x_2 = l_Lean_Parser_Term_doHave___closed__5; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_doHave___closed__9() { +static lean_object* _init_l_Lean_Parser_Term_doHave___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_doHave___closed__2; -x_2 = l_Lean_Parser_Term_doHave___closed__8; +x_2 = l_Lean_Parser_Term_doHave___closed__6; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -9010,7 +9016,7 @@ static lean_object* _init_l_Lean_Parser_Term_doHave() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Term_doHave___closed__9; +x_1 = l_Lean_Parser_Term_doHave___closed__7; return x_1; } } @@ -9044,7 +9050,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_doHave_declRange___clo { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_unsigned_to_nat(75u); -x_2 = lean_unsigned_to_nat(26u); +x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -9054,17 +9060,16 @@ return x_3; static lean_object* _init_l___regBuiltin_Lean_Parser_Term_doHave_declRange___closed__3() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___regBuiltin_Lean_Parser_Term_doHave_declRange___closed__1; x_2 = lean_unsigned_to_nat(25u); x_3 = l___regBuiltin_Lean_Parser_Term_doHave_declRange___closed__2; -x_4 = lean_unsigned_to_nat(26u); -x_5 = lean_alloc_ctor(0, 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; +x_4 = lean_alloc_ctor(0, 4, 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); +lean_ctor_set(x_4, 3, x_2); +return x_4; } } static lean_object* _init_l___regBuiltin_Lean_Parser_Term_doHave_declRange___closed__4() { @@ -9151,7 +9156,7 @@ static lean_object* _init_l_Lean_Parser_Term_doHave_formatter___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_doHave___closed__4; +x_1 = l_Lean_Parser_Term_notFollowedByRedefinedTermToken___closed__11; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; @@ -9256,7 +9261,7 @@ static lean_object* _init_l_Lean_Parser_Term_doHave_parenthesizer___closed__2() _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_doHave___closed__4; +x_1 = l_Lean_Parser_Term_notFollowedByRedefinedTermToken___closed__11; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; @@ -12187,7 +12192,7 @@ static lean_object* _init_l_Lean_Parser_Term_doUnless___closed__6() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("do ", 3); +x_1 = lean_mk_string_from_bytes(" do ", 4); return x_1; } } @@ -12299,7 +12304,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_doUnless_declRange___c { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_unsigned_to_nat(126u); -x_2 = lean_unsigned_to_nat(62u); +x_2 = lean_unsigned_to_nat(63u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -12313,7 +12318,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Parser_Term_doUnless_declRange___closed__1; x_2 = lean_unsigned_to_nat(25u); x_3 = l___regBuiltin_Lean_Parser_Term_doUnless_declRange___closed__2; -x_4 = lean_unsigned_to_nat(62u); +x_4 = lean_unsigned_to_nat(63u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -12910,50 +12915,77 @@ return x_5; static lean_object* _init_l_Lean_Parser_Term_doFor___closed__9() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("do ", 3); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Term_doFor___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Term_doFor___closed__9; +x_2 = l_Lean_Parser_symbol(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Term_doFor___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_doFor___closed__10; +x_2 = l_Lean_Parser_Term_doSeq; +x_3 = l_Lean_Parser_andthen(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Term_doFor___closed__12() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_doFor___closed__8; -x_2 = l_Lean_Parser_Term_doUnless___closed__8; +x_2 = l_Lean_Parser_Term_doFor___closed__11; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_doFor___closed__10() { +static lean_object* _init_l_Lean_Parser_Term_doFor___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_doFor___closed__5; -x_2 = l_Lean_Parser_Term_doFor___closed__9; +x_2 = l_Lean_Parser_Term_doFor___closed__12; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_doFor___closed__11() { +static lean_object* _init_l_Lean_Parser_Term_doFor___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_doFor___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_doFor___closed__10; +x_3 = l_Lean_Parser_Term_doFor___closed__13; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Term_doFor___closed__12() { +static lean_object* _init_l_Lean_Parser_Term_doFor___closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_doFor___closed__3; -x_2 = l_Lean_Parser_Term_doFor___closed__11; +x_2 = l_Lean_Parser_Term_doFor___closed__14; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_doFor___closed__13() { +static lean_object* _init_l_Lean_Parser_Term_doFor___closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_doFor___closed__2; -x_2 = l_Lean_Parser_Term_doFor___closed__12; +x_2 = l_Lean_Parser_Term_doFor___closed__15; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -12962,7 +12994,7 @@ static lean_object* _init_l_Lean_Parser_Term_doFor() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Term_doFor___closed__13; +x_1 = l_Lean_Parser_Term_doFor___closed__16; return x_1; } } @@ -13320,34 +13352,56 @@ return x_6; static lean_object* _init_l_Lean_Parser_Term_doFor_formatter___closed__5() { _start: { +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Term_doFor___closed__9; +x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_formatter), 6, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Term_doFor_formatter___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_doFor_formatter___closed__5; +x_2 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_210____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; +} +} +static lean_object* _init_l_Lean_Parser_Term_doFor_formatter___closed__7() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_doFor_formatter___closed__4; -x_2 = l_Lean_Parser_Term_doUnless_formatter___closed__5; +x_2 = l_Lean_Parser_Term_doFor_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; } } -static lean_object* _init_l_Lean_Parser_Term_doFor_formatter___closed__6() { +static lean_object* _init_l_Lean_Parser_Term_doFor_formatter___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_doFor_formatter___closed__2; -x_2 = l_Lean_Parser_Term_doFor_formatter___closed__5; +x_2 = l_Lean_Parser_Term_doFor_formatter___closed__7; 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; } } -static lean_object* _init_l_Lean_Parser_Term_doFor_formatter___closed__7() { +static lean_object* _init_l_Lean_Parser_Term_doFor_formatter___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_doFor___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_doFor_formatter___closed__6; +x_3 = l_Lean_Parser_Term_doFor_formatter___closed__8; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -13360,7 +13414,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_doFor_formatter(lean_object* x_1, le { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Term_doFor_formatter___closed__1; -x_7 = l_Lean_Parser_Term_doFor_formatter___closed__7; +x_7 = l_Lean_Parser_Term_doFor_formatter___closed__9; 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; } @@ -13619,34 +13673,56 @@ return x_6; static lean_object* _init_l_Lean_Parser_Term_doFor_parenthesizer___closed__5() { _start: { +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Term_doFor___closed__9; +x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_parenthesizer), 6, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Term_doFor_parenthesizer___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_doFor_parenthesizer___closed__5; +x_2 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_210____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); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Term_doFor_parenthesizer___closed__7() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_doFor_parenthesizer___closed__4; -x_2 = l_Lean_Parser_Term_doUnless_parenthesizer___closed__5; +x_2 = l_Lean_Parser_Term_doFor_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; } } -static lean_object* _init_l_Lean_Parser_Term_doFor_parenthesizer___closed__6() { +static lean_object* _init_l_Lean_Parser_Term_doFor_parenthesizer___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_doFor_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Term_doFor_parenthesizer___closed__5; +x_2 = l_Lean_Parser_Term_doFor_parenthesizer___closed__7; 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; } } -static lean_object* _init_l_Lean_Parser_Term_doFor_parenthesizer___closed__7() { +static lean_object* _init_l_Lean_Parser_Term_doFor_parenthesizer___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_doFor___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_doFor_parenthesizer___closed__6; +x_3 = l_Lean_Parser_Term_doFor_parenthesizer___closed__8; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -13659,7 +13735,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_doFor_parenthesizer(lean_object* x_1 { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Term_doFor_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Term_doFor_parenthesizer___closed__7; +x_7 = l_Lean_Parser_Term_doFor_parenthesizer___closed__9; 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; } @@ -13797,7 +13873,7 @@ static lean_object* _init_l_Lean_Parser_Term_doMatch___closed__9() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes(" with ", 6); +x_1 = lean_mk_string_from_bytes(" with", 5); return x_1; } } @@ -13929,7 +14005,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_doMatch_declRange___cl { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_unsigned_to_nat(142u); -x_2 = lean_unsigned_to_nat(51u); +x_2 = lean_unsigned_to_nat(50u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -13943,7 +14019,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Parser_Term_doMatch_declRange___closed__1; x_2 = lean_unsigned_to_nat(25u); x_3 = l___regBuiltin_Lean_Parser_Term_doMatch_declRange___closed__2; -x_4 = lean_unsigned_to_nat(51u); +x_4 = lean_unsigned_to_nat(50u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -14610,30 +14686,40 @@ return x_3; static lean_object* _init_l_Lean_Parser_Term_doCatch___closed__13() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_skip; +x_2 = l_Lean_Parser_Term_doCatch___closed__12; +x_3 = l_Lean_Parser_andthen(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Term_doCatch___closed__14() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_doCatch___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_doCatch___closed__12; +x_3 = l_Lean_Parser_Term_doCatch___closed__13; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Term_doCatch___closed__14() { +static lean_object* _init_l_Lean_Parser_Term_doCatch___closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_doCatch___closed__3; -x_2 = l_Lean_Parser_Term_doCatch___closed__13; +x_2 = l_Lean_Parser_Term_doCatch___closed__14; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_doCatch___closed__15() { +static lean_object* _init_l_Lean_Parser_Term_doCatch___closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_doCatch___closed__2; -x_2 = l_Lean_Parser_Term_doCatch___closed__14; +x_2 = l_Lean_Parser_Term_doCatch___closed__15; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -14642,7 +14728,7 @@ static lean_object* _init_l_Lean_Parser_Term_doCatch() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Term_doCatch___closed__15; +x_1 = l_Lean_Parser_Term_doCatch___closed__16; return x_1; } } @@ -14691,30 +14777,40 @@ return x_3; static lean_object* _init_l_Lean_Parser_Term_doCatchMatch___closed__5() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_skip; +x_2 = l_Lean_Parser_Term_doCatchMatch___closed__4; +x_3 = l_Lean_Parser_andthen(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Term_doCatchMatch___closed__6() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_doCatchMatch___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_doCatchMatch___closed__4; +x_3 = l_Lean_Parser_Term_doCatchMatch___closed__5; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Term_doCatchMatch___closed__6() { +static lean_object* _init_l_Lean_Parser_Term_doCatchMatch___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_doCatchMatch___closed__3; -x_2 = l_Lean_Parser_Term_doCatchMatch___closed__5; +x_2 = l_Lean_Parser_Term_doCatchMatch___closed__6; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_doCatchMatch___closed__7() { +static lean_object* _init_l_Lean_Parser_Term_doCatchMatch___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_doCatchMatch___closed__2; -x_2 = l_Lean_Parser_Term_doCatchMatch___closed__6; +x_2 = l_Lean_Parser_Term_doCatchMatch___closed__7; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -14723,7 +14819,7 @@ static lean_object* _init_l_Lean_Parser_Term_doCatchMatch() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Term_doCatchMatch___closed__7; +x_1 = l_Lean_Parser_Term_doCatchMatch___closed__8; return x_1; } } @@ -14789,30 +14885,40 @@ return x_3; static lean_object* _init_l_Lean_Parser_Term_doFinally___closed__7() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_skip; +x_2 = l_Lean_Parser_Term_doFinally___closed__6; +x_3 = l_Lean_Parser_andthen(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Term_doFinally___closed__8() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_doFinally___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_doFinally___closed__6; +x_3 = l_Lean_Parser_Term_doFinally___closed__7; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Term_doFinally___closed__8() { +static lean_object* _init_l_Lean_Parser_Term_doFinally___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_doFinally___closed__3; -x_2 = l_Lean_Parser_Term_doFinally___closed__7; +x_2 = l_Lean_Parser_Term_doFinally___closed__8; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_doFinally___closed__9() { +static lean_object* _init_l_Lean_Parser_Term_doFinally___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_doFinally___closed__2; -x_2 = l_Lean_Parser_Term_doFinally___closed__8; +x_2 = l_Lean_Parser_Term_doFinally___closed__9; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -14821,7 +14927,7 @@ static lean_object* _init_l_Lean_Parser_Term_doFinally() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Term_doFinally___closed__9; +x_1 = l_Lean_Parser_Term_doFinally___closed__10; return x_1; } } @@ -15108,13 +15214,23 @@ static lean_object* _init_l_Lean_Parser_Term_doCatch_formatter___closed__2() { _start: { lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Term_doSeqItem_formatter___closed__6; +x_2 = lean_alloc_closure((void*)(l_Lean_ppDedent_formatter), 6, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Term_doCatch_formatter___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; x_1 = l_Lean_Parser_Term_doCatch___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; } } -static lean_object* _init_l_Lean_Parser_Term_doCatch_formatter___closed__3() { +static lean_object* _init_l_Lean_Parser_Term_doCatch_formatter___closed__4() { _start: { lean_object* x_1; @@ -15122,29 +15238,29 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Term_binderIdent_formatter), 5, 0 return x_1; } } -static lean_object* _init_l_Lean_Parser_Term_doCatch_formatter___closed__4() { +static lean_object* _init_l_Lean_Parser_Term_doCatch_formatter___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_doCatch_formatter___closed__2; -x_2 = l_Lean_Parser_Term_doCatch_formatter___closed__3; +x_1 = l_Lean_Parser_Term_doCatch_formatter___closed__3; +x_2 = l_Lean_Parser_Term_doCatch_formatter___closed__4; 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; } } -static lean_object* _init_l_Lean_Parser_Term_doCatch_formatter___closed__5() { +static lean_object* _init_l_Lean_Parser_Term_doCatch_formatter___closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_doCatch_formatter___closed__4; +x_1 = l_Lean_Parser_Term_doCatch_formatter___closed__5; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_atomic_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_doCatch_formatter___closed__6() { +static lean_object* _init_l_Lean_Parser_Term_doCatch_formatter___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -15156,17 +15272,17 @@ lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_doCatch_formatter___closed__7() { +static lean_object* _init_l_Lean_Parser_Term_doCatch_formatter___closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_doCatch_formatter___closed__6; +x_1 = l_Lean_Parser_Term_doCatch_formatter___closed__7; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_optional_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_doCatch_formatter___closed__8() { +static lean_object* _init_l_Lean_Parser_Term_doCatch_formatter___closed__9() { _start: { lean_object* x_1; @@ -15174,11 +15290,11 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_darrow_formatter), 5, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_Term_doCatch_formatter___closed__9() { +static lean_object* _init_l_Lean_Parser_Term_doCatch_formatter___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_doCatch_formatter___closed__8; +x_1 = l_Lean_Parser_Term_doCatch_formatter___closed__9; x_2 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_210____closed__6; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -15186,37 +15302,49 @@ lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_doCatch_formatter___closed__10() { +static lean_object* _init_l_Lean_Parser_Term_doCatch_formatter___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_doCatch_formatter___closed__7; -x_2 = l_Lean_Parser_Term_doCatch_formatter___closed__9; +x_1 = l_Lean_Parser_Term_doCatch_formatter___closed__8; +x_2 = l_Lean_Parser_Term_doCatch_formatter___closed__10; 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; } } -static lean_object* _init_l_Lean_Parser_Term_doCatch_formatter___closed__11() { +static lean_object* _init_l_Lean_Parser_Term_doCatch_formatter___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_doCatch_formatter___closed__5; -x_2 = l_Lean_Parser_Term_doCatch_formatter___closed__10; +x_1 = l_Lean_Parser_Term_doCatch_formatter___closed__6; +x_2 = l_Lean_Parser_Term_doCatch_formatter___closed__11; 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; } } -static lean_object* _init_l_Lean_Parser_Term_doCatch_formatter___closed__12() { +static lean_object* _init_l_Lean_Parser_Term_doCatch_formatter___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_doCatch_formatter___closed__2; +x_2 = l_Lean_Parser_Term_doCatch_formatter___closed__12; +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; +} +} +static lean_object* _init_l_Lean_Parser_Term_doCatch_formatter___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_doCatch___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_doCatch_formatter___closed__11; +x_3 = l_Lean_Parser_Term_doCatch_formatter___closed__13; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -15229,7 +15357,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_doCatch_formatter(lean_object* x_1, { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Term_doCatch_formatter___closed__1; -x_7 = l_Lean_Parser_Term_doCatch_formatter___closed__12; +x_7 = l_Lean_Parser_Term_doCatch_formatter___closed__14; 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; } @@ -15289,7 +15417,7 @@ static lean_object* _init_l_Lean_Parser_Term_doCatchMatch_formatter___closed__2( _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_doCatch_formatter___closed__2; +x_1 = l_Lean_Parser_Term_doCatch_formatter___closed__3; x_2 = l_Lean_Parser_Term_doMatch_formatter___closed__10; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -15300,10 +15428,22 @@ return x_3; static lean_object* _init_l_Lean_Parser_Term_doCatchMatch_formatter___closed__3() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_doCatch_formatter___closed__2; +x_2 = l_Lean_Parser_Term_doCatchMatch_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; +} +} +static lean_object* _init_l_Lean_Parser_Term_doCatchMatch_formatter___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_doCatchMatch___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_doCatchMatch_formatter___closed__2; +x_3 = l_Lean_Parser_Term_doCatchMatch_formatter___closed__3; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -15316,7 +15456,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_doCatchMatch_formatter(lean_object* { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Term_doCatchMatch_formatter___closed__1; -x_7 = l_Lean_Parser_Term_doCatchMatch_formatter___closed__3; +x_7 = l_Lean_Parser_Term_doCatchMatch_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; } @@ -15397,10 +15537,22 @@ return x_3; static lean_object* _init_l_Lean_Parser_Term_doFinally_formatter___closed__4() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_doCatch_formatter___closed__2; +x_2 = l_Lean_Parser_Term_doFinally_formatter___closed__3; +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; +} +} +static lean_object* _init_l_Lean_Parser_Term_doFinally_formatter___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_doFinally___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_doFinally_formatter___closed__3; +x_3 = l_Lean_Parser_Term_doFinally_formatter___closed__4; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -15413,7 +15565,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_doFinally_formatter(lean_object* x_1 { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Term_doFinally_formatter___closed__1; -x_7 = l_Lean_Parser_Term_doFinally_formatter___closed__4; +x_7 = l_Lean_Parser_Term_doFinally_formatter___closed__5; 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; } @@ -15626,13 +15778,23 @@ static lean_object* _init_l_Lean_Parser_Term_doCatch_parenthesizer___closed__2() _start: { lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Term_doSeqItem_parenthesizer___closed__6; +x_2 = lean_alloc_closure((void*)(l_Lean_Parser_ppDedent_parenthesizer), 6, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Term_doCatch_parenthesizer___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; x_1 = l_Lean_Parser_Term_doCatch___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; } } -static lean_object* _init_l_Lean_Parser_Term_doCatch_parenthesizer___closed__3() { +static lean_object* _init_l_Lean_Parser_Term_doCatch_parenthesizer___closed__4() { _start: { lean_object* x_1; @@ -15640,29 +15802,29 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Term_binderIdent_parenthesizer), return x_1; } } -static lean_object* _init_l_Lean_Parser_Term_doCatch_parenthesizer___closed__4() { +static lean_object* _init_l_Lean_Parser_Term_doCatch_parenthesizer___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_doCatch_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Term_doCatch_parenthesizer___closed__3; +x_1 = l_Lean_Parser_Term_doCatch_parenthesizer___closed__3; +x_2 = l_Lean_Parser_Term_doCatch_parenthesizer___closed__4; 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; } } -static lean_object* _init_l_Lean_Parser_Term_doCatch_parenthesizer___closed__5() { +static lean_object* _init_l_Lean_Parser_Term_doCatch_parenthesizer___closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_doCatch_parenthesizer___closed__4; +x_1 = l_Lean_Parser_Term_doCatch_parenthesizer___closed__5; x_2 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_atomic_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_doCatch_parenthesizer___closed__6() { +static lean_object* _init_l_Lean_Parser_Term_doCatch_parenthesizer___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -15674,17 +15836,17 @@ lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_doCatch_parenthesizer___closed__7() { +static lean_object* _init_l_Lean_Parser_Term_doCatch_parenthesizer___closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_doCatch_parenthesizer___closed__6; +x_1 = l_Lean_Parser_Term_doCatch_parenthesizer___closed__7; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_optional_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_doCatch_parenthesizer___closed__8() { +static lean_object* _init_l_Lean_Parser_Term_doCatch_parenthesizer___closed__9() { _start: { lean_object* x_1; @@ -15692,11 +15854,11 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_darrow_parenthesizer), 5, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_Term_doCatch_parenthesizer___closed__9() { +static lean_object* _init_l_Lean_Parser_Term_doCatch_parenthesizer___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_doCatch_parenthesizer___closed__8; +x_1 = l_Lean_Parser_Term_doCatch_parenthesizer___closed__9; x_2 = l_Lean_Parser_Term_initFn____x40_Lean_Parser_Do___hyg_210____closed__9; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -15704,37 +15866,49 @@ lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_doCatch_parenthesizer___closed__10() { +static lean_object* _init_l_Lean_Parser_Term_doCatch_parenthesizer___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_doCatch_parenthesizer___closed__7; -x_2 = l_Lean_Parser_Term_doCatch_parenthesizer___closed__9; +x_1 = l_Lean_Parser_Term_doCatch_parenthesizer___closed__8; +x_2 = l_Lean_Parser_Term_doCatch_parenthesizer___closed__10; 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; } } -static lean_object* _init_l_Lean_Parser_Term_doCatch_parenthesizer___closed__11() { +static lean_object* _init_l_Lean_Parser_Term_doCatch_parenthesizer___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_doCatch_parenthesizer___closed__5; -x_2 = l_Lean_Parser_Term_doCatch_parenthesizer___closed__10; +x_1 = l_Lean_Parser_Term_doCatch_parenthesizer___closed__6; +x_2 = l_Lean_Parser_Term_doCatch_parenthesizer___closed__11; 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; } } -static lean_object* _init_l_Lean_Parser_Term_doCatch_parenthesizer___closed__12() { +static lean_object* _init_l_Lean_Parser_Term_doCatch_parenthesizer___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_doCatch_parenthesizer___closed__2; +x_2 = l_Lean_Parser_Term_doCatch_parenthesizer___closed__12; +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; +} +} +static lean_object* _init_l_Lean_Parser_Term_doCatch_parenthesizer___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_doCatch___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_doCatch_parenthesizer___closed__11; +x_3 = l_Lean_Parser_Term_doCatch_parenthesizer___closed__13; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -15747,7 +15921,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_doCatch_parenthesizer(lean_object* x { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Term_doCatch_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Term_doCatch_parenthesizer___closed__12; +x_7 = l_Lean_Parser_Term_doCatch_parenthesizer___closed__14; 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; } @@ -15807,7 +15981,7 @@ static lean_object* _init_l_Lean_Parser_Term_doCatchMatch_parenthesizer___closed _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_doCatch_parenthesizer___closed__2; +x_1 = l_Lean_Parser_Term_doCatch_parenthesizer___closed__3; x_2 = l_Lean_Parser_Term_doMatch_parenthesizer___closed__10; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -15818,10 +15992,22 @@ return x_3; static lean_object* _init_l_Lean_Parser_Term_doCatchMatch_parenthesizer___closed__3() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_doCatch_parenthesizer___closed__2; +x_2 = l_Lean_Parser_Term_doCatchMatch_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; +} +} +static lean_object* _init_l_Lean_Parser_Term_doCatchMatch_parenthesizer___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_doCatchMatch___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_doCatchMatch_parenthesizer___closed__2; +x_3 = l_Lean_Parser_Term_doCatchMatch_parenthesizer___closed__3; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -15834,7 +16020,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_doCatchMatch_parenthesizer(lean_obje { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Term_doCatchMatch_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Term_doCatchMatch_parenthesizer___closed__3; +x_7 = l_Lean_Parser_Term_doCatchMatch_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; } @@ -15915,10 +16101,22 @@ return x_3; static lean_object* _init_l_Lean_Parser_Term_doFinally_parenthesizer___closed__4() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_doCatch_parenthesizer___closed__2; +x_2 = l_Lean_Parser_Term_doFinally_parenthesizer___closed__3; +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; +} +} +static lean_object* _init_l_Lean_Parser_Term_doFinally_parenthesizer___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_doFinally___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_doFinally_parenthesizer___closed__3; +x_3 = l_Lean_Parser_Term_doFinally_parenthesizer___closed__4; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -15931,7 +16129,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_doFinally_parenthesizer(lean_object* { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Term_doFinally_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Term_doFinally_parenthesizer___closed__4; +x_7 = l_Lean_Parser_Term_doFinally_parenthesizer___closed__5; 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; } @@ -16939,85 +17137,78 @@ return x_5; static lean_object* _init_l_Lean_Parser_Term_doReturn___closed__4() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("return ", 7); -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_elseIf___closed__4; +x_2 = l_Lean_Parser_Term_liftMethod___closed__5; +x_3 = l_Lean_Parser_andthen(x_1, x_2); +return x_3; } } static lean_object* _init_l_Lean_Parser_Term_doReturn___closed__5() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_doReturn___closed__4; -x_2 = l_Lean_Parser_symbol(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Parser_Term_doReturn___closed__6() { -_start: -{ lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_elseIf___closed__4; -x_2 = l_Lean_Parser_Term_liftMethod___closed__5; +x_1 = l_Lean_Parser_skip; +x_2 = l_Lean_Parser_Term_doReturn___closed__4; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_doReturn___closed__7() { +static lean_object* _init_l_Lean_Parser_Term_doReturn___closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_doReturn___closed__6; +x_1 = l_Lean_Parser_Term_doReturn___closed__5; x_2 = l_Lean_Parser_optional(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_doReturn___closed__8() { +static lean_object* _init_l_Lean_Parser_Term_doReturn___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_doReturn___closed__5; -x_2 = l_Lean_Parser_Term_doReturn___closed__7; +x_1 = l_Lean_Parser_Term_notFollowedByRedefinedTermToken___closed__23; +x_2 = l_Lean_Parser_Term_doReturn___closed__6; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_doReturn___closed__9() { +static lean_object* _init_l_Lean_Parser_Term_doReturn___closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_doReturn___closed__8; +x_1 = l_Lean_Parser_Term_doReturn___closed__7; x_2 = l_Lean_Parser_withPosition(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_doReturn___closed__10() { +static lean_object* _init_l_Lean_Parser_Term_doReturn___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_doReturn___closed__2; x_2 = l_Lean_Parser_leadPrec; -x_3 = l_Lean_Parser_Term_doReturn___closed__9; +x_3 = l_Lean_Parser_Term_doReturn___closed__8; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Term_doReturn___closed__11() { +static lean_object* _init_l_Lean_Parser_Term_doReturn___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_doReturn___closed__3; -x_2 = l_Lean_Parser_Term_doReturn___closed__10; +x_2 = l_Lean_Parser_Term_doReturn___closed__9; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_doReturn___closed__12() { +static lean_object* _init_l_Lean_Parser_Term_doReturn___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_doReturn___closed__2; -x_2 = l_Lean_Parser_Term_doReturn___closed__11; +x_2 = l_Lean_Parser_Term_doReturn___closed__10; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -17026,7 +17217,7 @@ static lean_object* _init_l_Lean_Parser_Term_doReturn() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Term_doReturn___closed__12; +x_1 = l_Lean_Parser_Term_doReturn___closed__11; return x_1; } } @@ -17078,7 +17269,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_doReturn_declRange___c { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_unsigned_to_nat(168u); -x_2 = lean_unsigned_to_nat(66u); +x_2 = lean_unsigned_to_nat(76u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -17092,7 +17283,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Parser_Term_doReturn_declRange___closed__1; x_2 = lean_unsigned_to_nat(25u); x_3 = l___regBuiltin_Lean_Parser_Term_doReturn_declRange___closed__2; -x_4 = lean_unsigned_to_nat(66u); +x_4 = lean_unsigned_to_nat(76u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -17185,7 +17376,7 @@ static lean_object* _init_l_Lean_Parser_Term_doReturn_formatter___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_doReturn___closed__4; +x_1 = l_Lean_Parser_Term_notFollowedByRedefinedTermToken___closed__22; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; @@ -17206,42 +17397,54 @@ return x_3; static lean_object* _init_l_Lean_Parser_Term_doReturn_formatter___closed__4() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_doIdDecl_formatter___closed__2; +x_2 = l_Lean_Parser_Term_doReturn_formatter___closed__3; +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; +} +} +static lean_object* _init_l_Lean_Parser_Term_doReturn_formatter___closed__5() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_doReturn_formatter___closed__3; +x_1 = l_Lean_Parser_Term_doReturn_formatter___closed__4; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_optional_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_doReturn_formatter___closed__5() { +static lean_object* _init_l_Lean_Parser_Term_doReturn_formatter___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_doReturn_formatter___closed__2; -x_2 = l_Lean_Parser_Term_doReturn_formatter___closed__4; +x_2 = l_Lean_Parser_Term_doReturn_formatter___closed__5; 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; } } -static lean_object* _init_l_Lean_Parser_Term_doReturn_formatter___closed__6() { +static lean_object* _init_l_Lean_Parser_Term_doReturn_formatter___closed__7() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_doReturn_formatter___closed__5; +x_1 = l_Lean_Parser_Term_doReturn_formatter___closed__6; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_withPosition_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_doReturn_formatter___closed__7() { +static lean_object* _init_l_Lean_Parser_Term_doReturn_formatter___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_doReturn___closed__2; x_2 = l_Lean_Parser_leadPrec; -x_3 = l_Lean_Parser_Term_doReturn_formatter___closed__6; +x_3 = l_Lean_Parser_Term_doReturn_formatter___closed__7; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -17254,7 +17457,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_doReturn_formatter(lean_object* x_1, { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Term_doReturn_formatter___closed__1; -x_7 = l_Lean_Parser_Term_doReturn_formatter___closed__7; +x_7 = l_Lean_Parser_Term_doReturn_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; } @@ -17314,7 +17517,7 @@ static lean_object* _init_l_Lean_Parser_Term_doReturn_parenthesizer___closed__2( _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_doReturn___closed__4; +x_1 = l_Lean_Parser_Term_notFollowedByRedefinedTermToken___closed__22; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; @@ -17335,42 +17538,54 @@ return x_3; static lean_object* _init_l_Lean_Parser_Term_doReturn_parenthesizer___closed__4() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_doIdDecl_parenthesizer___closed__2; +x_2 = l_Lean_Parser_Term_doReturn_parenthesizer___closed__3; +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; +} +} +static lean_object* _init_l_Lean_Parser_Term_doReturn_parenthesizer___closed__5() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_doReturn_parenthesizer___closed__3; +x_1 = l_Lean_Parser_Term_doReturn_parenthesizer___closed__4; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_optional_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_doReturn_parenthesizer___closed__5() { +static lean_object* _init_l_Lean_Parser_Term_doReturn_parenthesizer___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_doReturn_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Term_doReturn_parenthesizer___closed__4; +x_2 = l_Lean_Parser_Term_doReturn_parenthesizer___closed__5; 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; } } -static lean_object* _init_l_Lean_Parser_Term_doReturn_parenthesizer___closed__6() { +static lean_object* _init_l_Lean_Parser_Term_doReturn_parenthesizer___closed__7() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_doReturn_parenthesizer___closed__5; +x_1 = l_Lean_Parser_Term_doReturn_parenthesizer___closed__6; x_2 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_withPosition_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_doReturn_parenthesizer___closed__7() { +static lean_object* _init_l_Lean_Parser_Term_doReturn_parenthesizer___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_doReturn___closed__2; x_2 = l_Lean_Parser_leadPrec; -x_3 = l_Lean_Parser_Term_doReturn_parenthesizer___closed__6; +x_3 = l_Lean_Parser_Term_doReturn_parenthesizer___closed__7; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -17383,7 +17598,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_doReturn_parenthesizer(lean_object* { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Term_doReturn_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Term_doReturn_parenthesizer___closed__7; +x_7 = l_Lean_Parser_Term_doReturn_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; } @@ -18957,7 +19172,7 @@ static lean_object* _init_l_Lean_Parser_Term_doNested___closed__4() { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_doNested___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_doUnless___closed__8; +x_3 = l_Lean_Parser_Term_doFor___closed__11; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } @@ -19129,7 +19344,7 @@ static lean_object* _init_l_Lean_Parser_Term_doNested_formatter___closed__2() { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_doNested___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_doUnless_formatter___closed__5; +x_3 = l_Lean_Parser_Term_doFor_formatter___closed__6; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -19204,7 +19419,7 @@ static lean_object* _init_l_Lean_Parser_Term_doNested_parenthesizer___closed__2( lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_doNested___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_doUnless_parenthesizer___closed__5; +x_3 = l_Lean_Parser_Term_doFor_parenthesizer___closed__6; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -19284,7 +19499,7 @@ static lean_object* _init_l_Lean_Parser_Term_do___closed__3() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_skip; -x_2 = l_Lean_Parser_Term_doUnless___closed__8; +x_2 = l_Lean_Parser_Term_doFor___closed__11; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -19474,7 +19689,7 @@ static lean_object* _init_l_Lean_Parser_Term_do_formatter___closed__3() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_do_formatter___closed__2; -x_2 = l_Lean_Parser_Term_doUnless_formatter___closed__5; +x_2 = l_Lean_Parser_Term_doFor_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); @@ -19569,7 +19784,7 @@ static lean_object* _init_l_Lean_Parser_Term_do_parenthesizer___closed__3() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_do_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Term_doUnless_parenthesizer___closed__5; +x_2 = l_Lean_Parser_Term_doFor_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); @@ -19752,7 +19967,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_termUnless_declRange__ { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_unsigned_to_nat(207u); -x_2 = lean_unsigned_to_nat(62u); +x_2 = lean_unsigned_to_nat(63u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -19766,7 +19981,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Parser_Term_termUnless_declRange___closed__1; x_2 = lean_unsigned_to_nat(23u); x_3 = l___regBuiltin_Lean_Parser_Term_termUnless_declRange___closed__2; -x_4 = lean_unsigned_to_nat(62u); +x_4 = lean_unsigned_to_nat(63u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -20022,30 +20237,50 @@ return x_5; static lean_object* _init_l_Lean_Parser_Term_termFor___closed__4() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_doFor___closed__8; +x_2 = l_Lean_Parser_Term_doUnless___closed__8; +x_3 = l_Lean_Parser_andthen(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Term_termFor___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_doFor___closed__5; +x_2 = l_Lean_Parser_Term_termFor___closed__4; +x_3 = l_Lean_Parser_andthen(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Term_termFor___closed__6() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_termFor___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_doFor___closed__10; +x_3 = l_Lean_Parser_Term_termFor___closed__5; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Term_termFor___closed__5() { +static lean_object* _init_l_Lean_Parser_Term_termFor___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_termFor___closed__3; -x_2 = l_Lean_Parser_Term_termFor___closed__4; +x_2 = l_Lean_Parser_Term_termFor___closed__6; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_termFor___closed__6() { +static lean_object* _init_l_Lean_Parser_Term_termFor___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_termFor___closed__2; -x_2 = l_Lean_Parser_Term_termFor___closed__5; +x_2 = l_Lean_Parser_Term_termFor___closed__7; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -20054,7 +20289,7 @@ static lean_object* _init_l_Lean_Parser_Term_termFor() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Term_termFor___closed__6; +x_1 = l_Lean_Parser_Term_termFor___closed__8; return x_1; } } @@ -20088,7 +20323,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_termFor_declRange___cl { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_unsigned_to_nat(209u); -x_2 = lean_unsigned_to_nat(51u); +x_2 = lean_unsigned_to_nat(52u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -20102,7 +20337,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Parser_Term_termFor_declRange___closed__1; x_2 = lean_unsigned_to_nat(23u); x_3 = l___regBuiltin_Lean_Parser_Term_termFor_declRange___closed__2; -x_4 = lean_unsigned_to_nat(51u); +x_4 = lean_unsigned_to_nat(52u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -20194,10 +20429,34 @@ return x_7; static lean_object* _init_l_Lean_Parser_Term_termFor_formatter___closed__2() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_doFor_formatter___closed__4; +x_2 = l_Lean_Parser_Term_doUnless_formatter___closed__5; +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; +} +} +static lean_object* _init_l_Lean_Parser_Term_termFor_formatter___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_doFor_formatter___closed__2; +x_2 = l_Lean_Parser_Term_termFor_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; +} +} +static lean_object* _init_l_Lean_Parser_Term_termFor_formatter___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_termFor___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_doFor_formatter___closed__6; +x_3 = l_Lean_Parser_Term_termFor_formatter___closed__3; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -20210,7 +20469,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_termFor_formatter(lean_object* x_1, { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Term_termFor_formatter___closed__1; -x_7 = l_Lean_Parser_Term_termFor_formatter___closed__2; +x_7 = l_Lean_Parser_Term_termFor_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; } @@ -20269,10 +20528,34 @@ return x_7; static lean_object* _init_l_Lean_Parser_Term_termFor_parenthesizer___closed__2() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_doFor_parenthesizer___closed__4; +x_2 = l_Lean_Parser_Term_doUnless_parenthesizer___closed__5; +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; +} +} +static lean_object* _init_l_Lean_Parser_Term_termFor_parenthesizer___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_doFor_parenthesizer___closed__2; +x_2 = l_Lean_Parser_Term_termFor_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; +} +} +static lean_object* _init_l_Lean_Parser_Term_termFor_parenthesizer___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_termFor___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_doFor_parenthesizer___closed__6; +x_3 = l_Lean_Parser_Term_termFor_parenthesizer___closed__3; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -20285,7 +20568,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_termFor_parenthesizer(lean_object* x { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Term_termFor_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Term_termFor_parenthesizer___closed__2; +x_7 = l_Lean_Parser_Term_termFor_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; } @@ -20697,7 +20980,7 @@ static lean_object* _init_l_Lean_Parser_Term_termReturn___closed__4() { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_termReturn___closed__2; x_2 = l_Lean_Parser_leadPrec; -x_3 = l_Lean_Parser_Term_doReturn___closed__9; +x_3 = l_Lean_Parser_Term_doReturn___closed__8; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } @@ -20778,7 +21061,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_termReturn_declRange__ { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_unsigned_to_nat(217u); -x_2 = lean_unsigned_to_nat(66u); +x_2 = lean_unsigned_to_nat(76u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -20792,7 +21075,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Parser_Term_termReturn_declRange___closed__1; x_2 = lean_unsigned_to_nat(23u); x_3 = l___regBuiltin_Lean_Parser_Term_termReturn_declRange___closed__2; -x_4 = lean_unsigned_to_nat(66u); +x_4 = lean_unsigned_to_nat(76u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -20887,7 +21170,7 @@ static lean_object* _init_l_Lean_Parser_Term_termReturn_formatter___closed__2() lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_termReturn___closed__2; x_2 = l_Lean_Parser_leadPrec; -x_3 = l_Lean_Parser_Term_doReturn_formatter___closed__6; +x_3 = l_Lean_Parser_Term_doReturn_formatter___closed__7; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -20962,7 +21245,7 @@ static lean_object* _init_l_Lean_Parser_Term_termReturn_parenthesizer___closed__ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_termReturn___closed__2; x_2 = l_Lean_Parser_leadPrec; -x_3 = l_Lean_Parser_Term_doReturn_parenthesizer___closed__6; +x_3 = l_Lean_Parser_Term_doReturn_parenthesizer___closed__7; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -22267,10 +22550,6 @@ l_Lean_Parser_Term_doHave___closed__6 = _init_l_Lean_Parser_Term_doHave___closed lean_mark_persistent(l_Lean_Parser_Term_doHave___closed__6); l_Lean_Parser_Term_doHave___closed__7 = _init_l_Lean_Parser_Term_doHave___closed__7(); lean_mark_persistent(l_Lean_Parser_Term_doHave___closed__7); -l_Lean_Parser_Term_doHave___closed__8 = _init_l_Lean_Parser_Term_doHave___closed__8(); -lean_mark_persistent(l_Lean_Parser_Term_doHave___closed__8); -l_Lean_Parser_Term_doHave___closed__9 = _init_l_Lean_Parser_Term_doHave___closed__9(); -lean_mark_persistent(l_Lean_Parser_Term_doHave___closed__9); l_Lean_Parser_Term_doHave = _init_l_Lean_Parser_Term_doHave(); lean_mark_persistent(l_Lean_Parser_Term_doHave); res = l___regBuiltin_Lean_Parser_Term_doHave(lean_io_mk_world()); @@ -22979,6 +23258,12 @@ l_Lean_Parser_Term_doFor___closed__12 = _init_l_Lean_Parser_Term_doFor___closed_ lean_mark_persistent(l_Lean_Parser_Term_doFor___closed__12); l_Lean_Parser_Term_doFor___closed__13 = _init_l_Lean_Parser_Term_doFor___closed__13(); lean_mark_persistent(l_Lean_Parser_Term_doFor___closed__13); +l_Lean_Parser_Term_doFor___closed__14 = _init_l_Lean_Parser_Term_doFor___closed__14(); +lean_mark_persistent(l_Lean_Parser_Term_doFor___closed__14); +l_Lean_Parser_Term_doFor___closed__15 = _init_l_Lean_Parser_Term_doFor___closed__15(); +lean_mark_persistent(l_Lean_Parser_Term_doFor___closed__15); +l_Lean_Parser_Term_doFor___closed__16 = _init_l_Lean_Parser_Term_doFor___closed__16(); +lean_mark_persistent(l_Lean_Parser_Term_doFor___closed__16); l_Lean_Parser_Term_doFor = _init_l_Lean_Parser_Term_doFor(); lean_mark_persistent(l_Lean_Parser_Term_doFor); res = l___regBuiltin_Lean_Parser_Term_doFor(lean_io_mk_world()); @@ -23047,6 +23332,10 @@ l_Lean_Parser_Term_doFor_formatter___closed__6 = _init_l_Lean_Parser_Term_doFor_ lean_mark_persistent(l_Lean_Parser_Term_doFor_formatter___closed__6); l_Lean_Parser_Term_doFor_formatter___closed__7 = _init_l_Lean_Parser_Term_doFor_formatter___closed__7(); lean_mark_persistent(l_Lean_Parser_Term_doFor_formatter___closed__7); +l_Lean_Parser_Term_doFor_formatter___closed__8 = _init_l_Lean_Parser_Term_doFor_formatter___closed__8(); +lean_mark_persistent(l_Lean_Parser_Term_doFor_formatter___closed__8); +l_Lean_Parser_Term_doFor_formatter___closed__9 = _init_l_Lean_Parser_Term_doFor_formatter___closed__9(); +lean_mark_persistent(l_Lean_Parser_Term_doFor_formatter___closed__9); l___regBuiltin_Lean_Parser_Term_doFor_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Term_doFor_formatter___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Term_doFor_formatter___closed__1); l___regBuiltin_Lean_Parser_Term_doFor_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Term_doFor_formatter___closed__2(); @@ -23095,6 +23384,10 @@ l_Lean_Parser_Term_doFor_parenthesizer___closed__6 = _init_l_Lean_Parser_Term_do lean_mark_persistent(l_Lean_Parser_Term_doFor_parenthesizer___closed__6); l_Lean_Parser_Term_doFor_parenthesizer___closed__7 = _init_l_Lean_Parser_Term_doFor_parenthesizer___closed__7(); lean_mark_persistent(l_Lean_Parser_Term_doFor_parenthesizer___closed__7); +l_Lean_Parser_Term_doFor_parenthesizer___closed__8 = _init_l_Lean_Parser_Term_doFor_parenthesizer___closed__8(); +lean_mark_persistent(l_Lean_Parser_Term_doFor_parenthesizer___closed__8); +l_Lean_Parser_Term_doFor_parenthesizer___closed__9 = _init_l_Lean_Parser_Term_doFor_parenthesizer___closed__9(); +lean_mark_persistent(l_Lean_Parser_Term_doFor_parenthesizer___closed__9); l___regBuiltin_Lean_Parser_Term_doFor_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Term_doFor_parenthesizer___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Term_doFor_parenthesizer___closed__1); l___regBuiltin_Lean_Parser_Term_doFor_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Term_doFor_parenthesizer___closed__2(); @@ -23274,6 +23567,8 @@ l_Lean_Parser_Term_doCatch___closed__14 = _init_l_Lean_Parser_Term_doCatch___clo lean_mark_persistent(l_Lean_Parser_Term_doCatch___closed__14); l_Lean_Parser_Term_doCatch___closed__15 = _init_l_Lean_Parser_Term_doCatch___closed__15(); lean_mark_persistent(l_Lean_Parser_Term_doCatch___closed__15); +l_Lean_Parser_Term_doCatch___closed__16 = _init_l_Lean_Parser_Term_doCatch___closed__16(); +lean_mark_persistent(l_Lean_Parser_Term_doCatch___closed__16); l_Lean_Parser_Term_doCatch = _init_l_Lean_Parser_Term_doCatch(); lean_mark_persistent(l_Lean_Parser_Term_doCatch); l_Lean_Parser_Term_doCatchMatch___closed__1 = _init_l_Lean_Parser_Term_doCatchMatch___closed__1(); @@ -23290,6 +23585,8 @@ l_Lean_Parser_Term_doCatchMatch___closed__6 = _init_l_Lean_Parser_Term_doCatchMa lean_mark_persistent(l_Lean_Parser_Term_doCatchMatch___closed__6); l_Lean_Parser_Term_doCatchMatch___closed__7 = _init_l_Lean_Parser_Term_doCatchMatch___closed__7(); lean_mark_persistent(l_Lean_Parser_Term_doCatchMatch___closed__7); +l_Lean_Parser_Term_doCatchMatch___closed__8 = _init_l_Lean_Parser_Term_doCatchMatch___closed__8(); +lean_mark_persistent(l_Lean_Parser_Term_doCatchMatch___closed__8); l_Lean_Parser_Term_doCatchMatch = _init_l_Lean_Parser_Term_doCatchMatch(); lean_mark_persistent(l_Lean_Parser_Term_doCatchMatch); l_Lean_Parser_Term_doFinally___closed__1 = _init_l_Lean_Parser_Term_doFinally___closed__1(); @@ -23310,6 +23607,8 @@ l_Lean_Parser_Term_doFinally___closed__8 = _init_l_Lean_Parser_Term_doFinally___ lean_mark_persistent(l_Lean_Parser_Term_doFinally___closed__8); l_Lean_Parser_Term_doFinally___closed__9 = _init_l_Lean_Parser_Term_doFinally___closed__9(); lean_mark_persistent(l_Lean_Parser_Term_doFinally___closed__9); +l_Lean_Parser_Term_doFinally___closed__10 = _init_l_Lean_Parser_Term_doFinally___closed__10(); +lean_mark_persistent(l_Lean_Parser_Term_doFinally___closed__10); l_Lean_Parser_Term_doFinally = _init_l_Lean_Parser_Term_doFinally(); lean_mark_persistent(l_Lean_Parser_Term_doFinally); l_Lean_Parser_Term_doTry___closed__1 = _init_l_Lean_Parser_Term_doTry___closed__1(); @@ -23386,6 +23685,10 @@ l_Lean_Parser_Term_doCatch_formatter___closed__11 = _init_l_Lean_Parser_Term_doC lean_mark_persistent(l_Lean_Parser_Term_doCatch_formatter___closed__11); l_Lean_Parser_Term_doCatch_formatter___closed__12 = _init_l_Lean_Parser_Term_doCatch_formatter___closed__12(); lean_mark_persistent(l_Lean_Parser_Term_doCatch_formatter___closed__12); +l_Lean_Parser_Term_doCatch_formatter___closed__13 = _init_l_Lean_Parser_Term_doCatch_formatter___closed__13(); +lean_mark_persistent(l_Lean_Parser_Term_doCatch_formatter___closed__13); +l_Lean_Parser_Term_doCatch_formatter___closed__14 = _init_l_Lean_Parser_Term_doCatch_formatter___closed__14(); +lean_mark_persistent(l_Lean_Parser_Term_doCatch_formatter___closed__14); l___regBuiltin_Lean_Parser_Term_doCatch_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Term_doCatch_formatter___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Term_doCatch_formatter___closed__1); l___regBuiltin_Lean_Parser_Term_doCatch_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Term_doCatch_formatter___closed__2(); @@ -23399,6 +23702,8 @@ l_Lean_Parser_Term_doCatchMatch_formatter___closed__2 = _init_l_Lean_Parser_Term lean_mark_persistent(l_Lean_Parser_Term_doCatchMatch_formatter___closed__2); l_Lean_Parser_Term_doCatchMatch_formatter___closed__3 = _init_l_Lean_Parser_Term_doCatchMatch_formatter___closed__3(); lean_mark_persistent(l_Lean_Parser_Term_doCatchMatch_formatter___closed__3); +l_Lean_Parser_Term_doCatchMatch_formatter___closed__4 = _init_l_Lean_Parser_Term_doCatchMatch_formatter___closed__4(); +lean_mark_persistent(l_Lean_Parser_Term_doCatchMatch_formatter___closed__4); l___regBuiltin_Lean_Parser_Term_doCatchMatch_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Term_doCatchMatch_formatter___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Term_doCatchMatch_formatter___closed__1); l___regBuiltin_Lean_Parser_Term_doCatchMatch_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Term_doCatchMatch_formatter___closed__2(); @@ -23414,6 +23719,8 @@ l_Lean_Parser_Term_doFinally_formatter___closed__3 = _init_l_Lean_Parser_Term_do lean_mark_persistent(l_Lean_Parser_Term_doFinally_formatter___closed__3); l_Lean_Parser_Term_doFinally_formatter___closed__4 = _init_l_Lean_Parser_Term_doFinally_formatter___closed__4(); lean_mark_persistent(l_Lean_Parser_Term_doFinally_formatter___closed__4); +l_Lean_Parser_Term_doFinally_formatter___closed__5 = _init_l_Lean_Parser_Term_doFinally_formatter___closed__5(); +lean_mark_persistent(l_Lean_Parser_Term_doFinally_formatter___closed__5); l___regBuiltin_Lean_Parser_Term_doFinally_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Term_doFinally_formatter___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Term_doFinally_formatter___closed__1); l___regBuiltin_Lean_Parser_Term_doFinally_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Term_doFinally_formatter___closed__2(); @@ -23470,6 +23777,10 @@ l_Lean_Parser_Term_doCatch_parenthesizer___closed__11 = _init_l_Lean_Parser_Term lean_mark_persistent(l_Lean_Parser_Term_doCatch_parenthesizer___closed__11); l_Lean_Parser_Term_doCatch_parenthesizer___closed__12 = _init_l_Lean_Parser_Term_doCatch_parenthesizer___closed__12(); lean_mark_persistent(l_Lean_Parser_Term_doCatch_parenthesizer___closed__12); +l_Lean_Parser_Term_doCatch_parenthesizer___closed__13 = _init_l_Lean_Parser_Term_doCatch_parenthesizer___closed__13(); +lean_mark_persistent(l_Lean_Parser_Term_doCatch_parenthesizer___closed__13); +l_Lean_Parser_Term_doCatch_parenthesizer___closed__14 = _init_l_Lean_Parser_Term_doCatch_parenthesizer___closed__14(); +lean_mark_persistent(l_Lean_Parser_Term_doCatch_parenthesizer___closed__14); l___regBuiltin_Lean_Parser_Term_doCatch_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Term_doCatch_parenthesizer___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Term_doCatch_parenthesizer___closed__1); l___regBuiltin_Lean_Parser_Term_doCatch_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Term_doCatch_parenthesizer___closed__2(); @@ -23483,6 +23794,8 @@ l_Lean_Parser_Term_doCatchMatch_parenthesizer___closed__2 = _init_l_Lean_Parser_ lean_mark_persistent(l_Lean_Parser_Term_doCatchMatch_parenthesizer___closed__2); l_Lean_Parser_Term_doCatchMatch_parenthesizer___closed__3 = _init_l_Lean_Parser_Term_doCatchMatch_parenthesizer___closed__3(); lean_mark_persistent(l_Lean_Parser_Term_doCatchMatch_parenthesizer___closed__3); +l_Lean_Parser_Term_doCatchMatch_parenthesizer___closed__4 = _init_l_Lean_Parser_Term_doCatchMatch_parenthesizer___closed__4(); +lean_mark_persistent(l_Lean_Parser_Term_doCatchMatch_parenthesizer___closed__4); l___regBuiltin_Lean_Parser_Term_doCatchMatch_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Term_doCatchMatch_parenthesizer___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Term_doCatchMatch_parenthesizer___closed__1); l___regBuiltin_Lean_Parser_Term_doCatchMatch_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Term_doCatchMatch_parenthesizer___closed__2(); @@ -23498,6 +23811,8 @@ l_Lean_Parser_Term_doFinally_parenthesizer___closed__3 = _init_l_Lean_Parser_Ter lean_mark_persistent(l_Lean_Parser_Term_doFinally_parenthesizer___closed__3); l_Lean_Parser_Term_doFinally_parenthesizer___closed__4 = _init_l_Lean_Parser_Term_doFinally_parenthesizer___closed__4(); lean_mark_persistent(l_Lean_Parser_Term_doFinally_parenthesizer___closed__4); +l_Lean_Parser_Term_doFinally_parenthesizer___closed__5 = _init_l_Lean_Parser_Term_doFinally_parenthesizer___closed__5(); +lean_mark_persistent(l_Lean_Parser_Term_doFinally_parenthesizer___closed__5); l___regBuiltin_Lean_Parser_Term_doFinally_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Term_doFinally_parenthesizer___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Term_doFinally_parenthesizer___closed__1); l___regBuiltin_Lean_Parser_Term_doFinally_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Term_doFinally_parenthesizer___closed__2(); @@ -23690,8 +24005,6 @@ l_Lean_Parser_Term_doReturn___closed__10 = _init_l_Lean_Parser_Term_doReturn___c lean_mark_persistent(l_Lean_Parser_Term_doReturn___closed__10); l_Lean_Parser_Term_doReturn___closed__11 = _init_l_Lean_Parser_Term_doReturn___closed__11(); lean_mark_persistent(l_Lean_Parser_Term_doReturn___closed__11); -l_Lean_Parser_Term_doReturn___closed__12 = _init_l_Lean_Parser_Term_doReturn___closed__12(); -lean_mark_persistent(l_Lean_Parser_Term_doReturn___closed__12); l_Lean_Parser_Term_doReturn = _init_l_Lean_Parser_Term_doReturn(); lean_mark_persistent(l_Lean_Parser_Term_doReturn); res = l___regBuiltin_Lean_Parser_Term_doReturn(lean_io_mk_world()); @@ -23733,6 +24046,8 @@ l_Lean_Parser_Term_doReturn_formatter___closed__6 = _init_l_Lean_Parser_Term_doR lean_mark_persistent(l_Lean_Parser_Term_doReturn_formatter___closed__6); l_Lean_Parser_Term_doReturn_formatter___closed__7 = _init_l_Lean_Parser_Term_doReturn_formatter___closed__7(); lean_mark_persistent(l_Lean_Parser_Term_doReturn_formatter___closed__7); +l_Lean_Parser_Term_doReturn_formatter___closed__8 = _init_l_Lean_Parser_Term_doReturn_formatter___closed__8(); +lean_mark_persistent(l_Lean_Parser_Term_doReturn_formatter___closed__8); l___regBuiltin_Lean_Parser_Term_doReturn_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Term_doReturn_formatter___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Term_doReturn_formatter___closed__1); l___regBuiltin_Lean_Parser_Term_doReturn_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Term_doReturn_formatter___closed__2(); @@ -23754,6 +24069,8 @@ l_Lean_Parser_Term_doReturn_parenthesizer___closed__6 = _init_l_Lean_Parser_Term lean_mark_persistent(l_Lean_Parser_Term_doReturn_parenthesizer___closed__6); l_Lean_Parser_Term_doReturn_parenthesizer___closed__7 = _init_l_Lean_Parser_Term_doReturn_parenthesizer___closed__7(); lean_mark_persistent(l_Lean_Parser_Term_doReturn_parenthesizer___closed__7); +l_Lean_Parser_Term_doReturn_parenthesizer___closed__8 = _init_l_Lean_Parser_Term_doReturn_parenthesizer___closed__8(); +lean_mark_persistent(l_Lean_Parser_Term_doReturn_parenthesizer___closed__8); l___regBuiltin_Lean_Parser_Term_doReturn_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Term_doReturn_parenthesizer___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Term_doReturn_parenthesizer___closed__1); l___regBuiltin_Lean_Parser_Term_doReturn_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Term_doReturn_parenthesizer___closed__2(); @@ -24218,6 +24535,10 @@ l_Lean_Parser_Term_termFor___closed__5 = _init_l_Lean_Parser_Term_termFor___clos lean_mark_persistent(l_Lean_Parser_Term_termFor___closed__5); l_Lean_Parser_Term_termFor___closed__6 = _init_l_Lean_Parser_Term_termFor___closed__6(); lean_mark_persistent(l_Lean_Parser_Term_termFor___closed__6); +l_Lean_Parser_Term_termFor___closed__7 = _init_l_Lean_Parser_Term_termFor___closed__7(); +lean_mark_persistent(l_Lean_Parser_Term_termFor___closed__7); +l_Lean_Parser_Term_termFor___closed__8 = _init_l_Lean_Parser_Term_termFor___closed__8(); +lean_mark_persistent(l_Lean_Parser_Term_termFor___closed__8); l_Lean_Parser_Term_termFor = _init_l_Lean_Parser_Term_termFor(); lean_mark_persistent(l_Lean_Parser_Term_termFor); res = l___regBuiltin_Lean_Parser_Term_termFor(lean_io_mk_world()); @@ -24244,6 +24565,10 @@ l_Lean_Parser_Term_termFor_formatter___closed__1 = _init_l_Lean_Parser_Term_term lean_mark_persistent(l_Lean_Parser_Term_termFor_formatter___closed__1); l_Lean_Parser_Term_termFor_formatter___closed__2 = _init_l_Lean_Parser_Term_termFor_formatter___closed__2(); lean_mark_persistent(l_Lean_Parser_Term_termFor_formatter___closed__2); +l_Lean_Parser_Term_termFor_formatter___closed__3 = _init_l_Lean_Parser_Term_termFor_formatter___closed__3(); +lean_mark_persistent(l_Lean_Parser_Term_termFor_formatter___closed__3); +l_Lean_Parser_Term_termFor_formatter___closed__4 = _init_l_Lean_Parser_Term_termFor_formatter___closed__4(); +lean_mark_persistent(l_Lean_Parser_Term_termFor_formatter___closed__4); l___regBuiltin_Lean_Parser_Term_termFor_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Term_termFor_formatter___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Term_termFor_formatter___closed__1); l___regBuiltin_Lean_Parser_Term_termFor_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Term_termFor_formatter___closed__2(); @@ -24255,6 +24580,10 @@ l_Lean_Parser_Term_termFor_parenthesizer___closed__1 = _init_l_Lean_Parser_Term_ lean_mark_persistent(l_Lean_Parser_Term_termFor_parenthesizer___closed__1); l_Lean_Parser_Term_termFor_parenthesizer___closed__2 = _init_l_Lean_Parser_Term_termFor_parenthesizer___closed__2(); lean_mark_persistent(l_Lean_Parser_Term_termFor_parenthesizer___closed__2); +l_Lean_Parser_Term_termFor_parenthesizer___closed__3 = _init_l_Lean_Parser_Term_termFor_parenthesizer___closed__3(); +lean_mark_persistent(l_Lean_Parser_Term_termFor_parenthesizer___closed__3); +l_Lean_Parser_Term_termFor_parenthesizer___closed__4 = _init_l_Lean_Parser_Term_termFor_parenthesizer___closed__4(); +lean_mark_persistent(l_Lean_Parser_Term_termFor_parenthesizer___closed__4); l___regBuiltin_Lean_Parser_Term_termFor_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Term_termFor_parenthesizer___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Term_termFor_parenthesizer___closed__1); l___regBuiltin_Lean_Parser_Term_termFor_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Term_termFor_parenthesizer___closed__2(); diff --git a/stage0/stdlib/Lean/Parser/Extension.c b/stage0/stdlib/Lean/Parser/Extension.c index a93612d5eee9..a885c0d83cdf 100644 --- a/stage0/stdlib/Lean/Parser/Extension.c +++ b/stage0/stdlib/Lean/Parser/Extension.c @@ -15,7 +15,6 @@ extern "C" { #endif static lean_object* l_Lean_Parser_mkParserOfConstantUnsafe___closed__5; lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); -uint8_t l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_addTrailingParser(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Parser_Extension___hyg_6874____closed__12; LEAN_EXPORT lean_object* l_Lean_Parser_getUnaryAlias___rarg___boxed(lean_object*, lean_object*, lean_object*); @@ -376,6 +375,7 @@ static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5192 lean_object* l_Lean_MapDeclarationExtension_find_x3f___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_runParserCategory___closed__2; static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_BuiltinParserAttribute_add___lambda__2___closed__10; +uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_ofExcept___at_Lean_Parser_mkParserOfConstantUnsafe___spec__1(lean_object*, lean_object*); lean_object* l_Lean_Parser_nodeWithAntiquot(lean_object*, lean_object*, lean_object*, uint8_t); static lean_object* l_Lean_Parser_parserOfStackFn___lambda__2___closed__5; @@ -434,7 +434,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_Parser static lean_object* l_Lean_Parser_registerAlias___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_addParser___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Environment_contains(lean_object*, lean_object*); -lean_object* l_Lean_Option_register___at_Std_Format_initFn____x40_Lean_Data_Format___hyg_95____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_resolveParserNameCore_isParser___closed__2; static lean_object* l_Lean_Parser_parserOfStack___closed__1; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_4551____closed__3; @@ -553,6 +552,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_getSyntaxNodeKinds(lean_object*); static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_4199____closed__17; lean_object* l_Lean_ResolveName_resolveGlobalName(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); +lean_object* l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_prattParser(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forM___at___private_Lean_Parser_Extension_0__Lean_Parser_ParserAttribute_add___spec__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at___private_Lean_Parser_Extension_0__Lean_Parser_addParserCategoryCore___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -8280,7 +8280,7 @@ 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_Extension___hyg_4848____closed__3; x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_4848____closed__5; x_4 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_4848____closed__6; -x_5 = l_Lean_Option_register___at_Std_Format_initFn____x40_Lean_Data_Format___hyg_95____spec__1(x_2, x_3, x_4, x_1); +x_5 = l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__1(x_2, x_3, x_4, x_1); return x_5; } } @@ -8324,7 +8324,7 @@ lean_inc(x_11); x_12 = lean_ctor_get(x_11, 1); lean_inc(x_12); x_13 = l_Lean_Parser_evalInsideQuot___elambda__1___closed__1; -x_14 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_12, x_13); +x_14 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_12, x_13); lean_dec(x_12); if (x_14 == 0) { @@ -15255,7 +15255,7 @@ x_8 = lean_ctor_get(x_2, 1); x_9 = lean_ctor_get(x_2, 2); x_10 = lean_ctor_get(x_2, 3); x_11 = l_Lean_Parser_evalInsideQuot___elambda__1___closed__1; -x_12 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_8, x_11); +x_12 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_8, x_11); if (x_12 == 0) { uint8_t x_13; @@ -15319,7 +15319,7 @@ lean_inc(x_26); lean_inc(x_25); lean_dec(x_2); x_29 = l_Lean_Parser_evalInsideQuot___elambda__1___closed__1; -x_30 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_26, x_29); +x_30 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_26, x_29); if (x_30 == 0) { lean_object* x_31; lean_object* x_32; uint8_t x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; diff --git a/stage0/stdlib/Lean/Parser/Extra.c b/stage0/stdlib/Lean/Parser/Extra.c index 5711560af60c..2e5fe257bb72 100644 --- a/stage0/stdlib/Lean/Parser/Extra.c +++ b/stage0/stdlib/Lean/Parser/Extra.c @@ -13,159 +13,149 @@ #ifdef __cplusplus extern "C" { #endif -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__89; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__55; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__103; 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_EXPORT lean_object* l_Lean_Parser_many1_formatter(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__1___closed__66; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__104; lean_object* l_Lean_PrettyPrinter_Formatter_setLhsPrec_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ppHardSpace_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__32; lean_object* l_Lean_KeyedDeclsAttribute_addBuiltin___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__23; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__122; lean_object* l_Lean_Parser_checkColEq(lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__42; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__51; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__121; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__41; LEAN_EXPORT lean_object* l_Lean_Parser_rawIdent_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__6; static lean_object* l_Lean_Parser_many_formatter___closed__3; lean_object* l_Lean_Syntax_mkNameLit(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__52; lean_object* l_Lean_PrettyPrinter_Formatter_pushAlign(uint8_t, lean_object*, lean_object*, lean_object*, 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_2047____closed__25; LEAN_EXPORT lean_object* l_Lean_Parser_charLit_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ppHardLineUnlessUngrouped; 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_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__28; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__37; static lean_object* l_Lean_Parser_sepByIndent_parenthesizer___closed__1; static lean_object* l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__3; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__53; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__121; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__57; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__19; LEAN_EXPORT lean_object* l_Lean_Parser_ppHardSpace_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*); -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_2065____closed__13; extern lean_object* l_Lean_Parser_pushNone; LEAN_EXPORT lean_object* l_Lean_Parser_rawIdent_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_antiquotNestedExpr_parenthesizer___closed__1; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__33; extern lean_object* l_Lean_Parser_nameLitNoAntiquot; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__101; 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*); -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_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__16; LEAN_EXPORT lean_object* l_Lean_Parser_withCache_parenthesizer(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__31; LEAN_EXPORT lean_object* l_Lean_Parser_ppRealGroup_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; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__48; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__26; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__37; LEAN_EXPORT lean_object* l_Lean_Parser_adaptCacheableContext_formatter(lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__100; LEAN_EXPORT lean_object* l_Lean_Parser_ppRealGroup___boxed(lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__113; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__13; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__98; -static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__13; 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_withForbidden_parenthesizer___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__36; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__30; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__47; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__91; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__35; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__13; LEAN_EXPORT lean_object* l_Lean_ppLine_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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_2047____closed__77; -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*); +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__74; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__87; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__110; lean_object* l_Lean_PrettyPrinter_Formatter_rawIdentNoAntiquot_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_forIn_loop___at_Lean_Parser_sepByIndent_formatter___spec__2___closed__4; lean_object* l_Lean_PrettyPrinter_Parenthesizer_strLitNoAntiquot_parenthesizer___boxed(lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__66; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__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__1___closed__38; LEAN_EXPORT lean_object* l_Lean_Parser_manyIndent(lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__122; -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; +static lean_object* l_Lean_Parser_hygieneInfo_formatter___closed__3; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__20; 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*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__117; -static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__29; -static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__14; -static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__32; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__79; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__17; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__50; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__73; LEAN_EXPORT lean_object* l_Lean_Parser_ppRealFill___boxed(lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__67; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__62; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__82; static lean_object* l_Lean_Parser_charLit___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_withCache_formatter___rarg(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__1___closed__39; lean_object* l_Lean_PrettyPrinter_Formatter_checkPrec_formatter___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__1___closed__49; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__15; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__12; static lean_object* l_Lean_ppDedent_formatter___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__114; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__13; LEAN_EXPORT lean_object* l_Lean_Parser_many_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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__1___closed__63; lean_object* l_Lean_Syntax_getId(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_dbgTraceState_parenthesizer___boxed(lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__16; static lean_object* l_List_forIn_loop___at_Lean_Parser_sepByIndent_formatter___spec__2___closed__1; 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_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__19; static lean_object* l_Lean_Parser_strLit_formatter___closed__2; lean_object* lean_array_push(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_group_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__26; lean_object* l_Lean_PrettyPrinter_Formatter_sepByNoAntiquot_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___lambda__1___closed__80; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__49; static lean_object* l_Lean_Parser_mkAntiquot_parenthesizer___closed__5; -static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__21; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__75; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__102; LEAN_EXPORT lean_object* l_Lean_Parser_notSymbol(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__28; static lean_object* l_Lean_Parser_optional___closed__1; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__28; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__17; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__68; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__78; static lean_object* l_Lean_Parser_rawIdent___closed__1; static lean_object* l_Lean_Parser_mkAntiquot_formatter___closed__8; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__44; +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; LEAN_EXPORT lean_object* l_Lean_Parser_ppSpace_parenthesizer___rarg(lean_object*); 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___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__12; uint8_t lean_usize_dec_eq(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Parser_notSymbol_parenthesizer___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_evalInsideQuot_parenthesizer___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ppHardLineUnlessUngrouped_formatter___boxed(lean_object*, lean_object*, lean_object*, 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*); lean_object* l_Lean_Syntax_getArgs(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__14; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__32; LEAN_EXPORT lean_object* l_Lean_Parser_sepByIndent_formatter___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_withCache_parenthesizer___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_dbgTraceState_formatter___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__67; LEAN_EXPORT lean_object* l_Lean_Parser_optional(lean_object*); lean_object* l_String_trim(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; LEAN_EXPORT lean_object* l_Lean_Parser_termParser_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__1___closed__61; LEAN_EXPORT lean_object* l_Lean_Parser_nameLit_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__8; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__25; static lean_object* l_Lean_Parser_strLit___closed__2; +LEAN_EXPORT lean_object* l_Lean_Parser_hygieneInfo_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); lean_object* l_ReaderT_bind___at_Lean_PrettyPrinter_Formatter_categoryFormatterCore___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__2; +LEAN_EXPORT lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29____________; LEAN_EXPORT lean_object* l_Lean_Parser_evalInsideQuot_parenthesizer___boxed(lean_object*); lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter(lean_object*); static lean_object* l_Lean_Parser_mkAntiquotSplice_formatter___closed__7; -static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_strLit; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__18; static lean_object* l_Lean_ppHardSpace_formatter___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_ppLine_parenthesizer(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_sepByElemParser_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_2065____closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_withOpen_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_checkLinebreakBefore(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_withOpenDecl_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___lambda__1___closed__47; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__30; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__76; lean_object* l_Lean_Macro_throwError___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_withPositionAfterLinebreak_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_2065____closed__44; LEAN_EXPORT lean_object* l_Lean_Parser_incQuotDepth_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_scientificLit_parenthesizer___closed__1; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__74; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__16; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__78; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__91; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__18; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__10; 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_mkAntiquot_parenthesizer___closed__8; lean_object* l_Lean_PrettyPrinter_Formatter_checkNoWsBefore_formatter___boxed(lean_object*); @@ -174,91 +164,95 @@ lean_object* l_Lean_mkIdentFrom(lean_object*, lean_object*, uint8_t); static lean_object* l_Lean_Parser_mkAntiquotSplice_parenthesizer___closed__1; static lean_object* l_Lean_Parser_ident_formatter___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_decQuotDepth_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_2065____closed__90; +LEAN_EXPORT lean_object* l_Lean_Parser_hygieneInfo; lean_object* l_Lean_PrettyPrinter_Parenthesizer_checkColEq_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__6; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__50; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__14; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__8; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__55; 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_EXPORT lean_object* l_Lean_Parser_adaptCacheableContext_parenthesizer___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_node5(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__4; static lean_object* l_Lean_Parser_antiquotExpr_formatter___closed__4; uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_withForbidden_formatter(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__1(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__58; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__35; static lean_object* l_Lean_Parser_antiquotExpr_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__1___closed__48; static lean_object* l_Lean_Parser_group_formatter___closed__1; static lean_object* l_Lean_Parser_rawIdent_parenthesizer___closed__1; 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_scientificLit_formatter___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_ppAllowUngrouped_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__2; static lean_object* l___regBuiltin_Lean_Parser_antiquotNestedExpr_parenthesizer___closed__2; 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___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__33; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__92; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__56; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__80; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__38; static lean_object* l_List_forIn_loop___at_Lean_Parser_sepByIndent_formatter___spec__2___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__73; static lean_object* l_Lean_Parser_ident_parenthesizer___closed__1; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__59; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__85; static lean_object* l_Lean_Parser_antiquotExpr_parenthesizer___closed__2; +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065_(lean_object*); lean_object* l_Lean_PrettyPrinter_Formatter_charLitNoAntiquot_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PrettyPrinter_Parenthesizer_identNoAntiquot_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__112; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__65; static lean_object* l_Lean_Parser_antiquotNestedExpr_formatter___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_setExpected_formatter___rarg(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__1___closed__36; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__15; static lean_object* l_Lean_Parser_mkAntiquot_parenthesizer___closed__2; static lean_object* l_Lean_Parser_leadingNode_formatter___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__110; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__72; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__65; LEAN_EXPORT lean_object* l_Lean_Parser_withCache_formatter(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ppLine_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PrettyPrinter_Parenthesizer_categoryParser_parenthesizer(lean_object*, lean_object*, 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*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__1; 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___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__43; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__103; static lean_object* l_Lean_Parser_strLit_formatter___closed__3; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__22; 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_mkAntiquot_parenthesizer___closed__10; 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_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__80; 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*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__32; static lean_object* l_Lean_Parser_mkAntiquot_parenthesizer___closed__4; -static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__27; 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*); +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_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_addQuotDepth_parenthesizer(lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__71; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__29; 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_mkAntiquot_formatter___closed__6; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__11; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__90; LEAN_EXPORT lean_object* l_Lean_Parser_setExpected_formatter(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__48; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__4; static lean_object* l_Lean_Parser_strLit_parenthesizer___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_withCache_parenthesizer___boxed(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___lambda__1___closed__18; LEAN_EXPORT lean_object* l_Lean_Parser_group(lean_object*); 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_2065____closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_setExpected_parenthesizer___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr3(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__57; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__45; static lean_object* l_Lean_Parser_sepByElemParser_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__1___closed__52; 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_ppAllowUngrouped_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_notSymbol_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__28; -static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________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__77; LEAN_EXPORT lean_object* l_Lean_Parser_ppGroup(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__79; size_t lean_usize_of_nat(lean_object*); lean_object* l_Lean_PrettyPrinter_Parenthesizer_checkNoWsBefore_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__21; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__55; 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_charLit_formatter___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__20; -static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__22; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__68; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__85; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__28; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__15; LEAN_EXPORT lean_object* l_Lean_Parser_notSymbol_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__53; lean_object* l_Lean_Macro_resolveGlobalName(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_antiquotExpr_parenthesizer___closed__3; static lean_object* l_Lean_Parser_many1Indent_formatter___closed__1; @@ -266,169 +260,185 @@ static lean_object* l_Lean_Parser_antiquotNestedExpr_formatter___closed__6; lean_object* l_Lean_Parser_withPosition(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*); LEAN_EXPORT lean_object* l_Lean_Parser_nodeWithAntiquot_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__1___closed__5; +extern lean_object* l_Lean_Parser_hygieneInfoNoAntiquot; static lean_object* l_Lean_Parser_mkAntiquot_formatter___closed__14; static lean_object* l_Lean_Parser_scientificLit_formatter___closed__4; lean_object* l_Lean_Parser_optionalNoAntiquot(lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__5; -static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__10; static lean_object* l_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__1___closed__67; lean_object* lean_st_ref_take(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__41; 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_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__43; +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; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l_Lean_ppHardSpace_formatter___closed__2; 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___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_mkAntiquot_formatter___closed__18; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__78; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__95; LEAN_EXPORT lean_object* l_Lean_Parser_setExpected_parenthesizer(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_notSymbol_formatter___rarg(lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__18; LEAN_EXPORT lean_object* l_Lean_Parser_many1_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__37; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__105; LEAN_EXPORT lean_object* l_Lean_Parser_commandParser_formatter(lean_object*); static lean_object* l_Lean_Parser_patternIgnore_formatter___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_withoutPosition_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_strLit___closed__1; +static lean_object* l_Lean_Parser_hygieneInfo_parenthesizer___closed__2; 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_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__93; extern lean_object* l_Lean_Parser_charLitNoAntiquot; lean_object* l_Lean_SourceInfo_fromRef(lean_object*, uint8_t); lean_object* lean_nat_to_int(lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__99; LEAN_EXPORT lean_object* l_Lean_Parser_many1Indent_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_2065____closed__115; lean_object* l_Lean_Syntax_node6(lean_object*, lean_object*, 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___lambda__1___closed__10; LEAN_EXPORT lean_object* l_Lean_ppAllowUngrouped_formatter___boxed(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_ppRealGroup_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__123; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__61; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__117; LEAN_EXPORT lean_object* l_Lean_Parser_withOpen_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_2047____closed__95; lean_object* l_Lean_PrettyPrinter_Parenthesizer_checkColGe_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_mkStrLit(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__68; lean_object* l_List_range(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_nonReservedSymbol_parenthesizer___boxed(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__1___closed__6; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__20; lean_object* l_Lean_PrettyPrinter_Parenthesizer_sepByNoAntiquot_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PrettyPrinter_Parenthesizer_atomic_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_antiquotNestedExpr_formatter___closed__4; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__72; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__39; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__23; static lean_object* l_Lean_Parser_charLit_formatter___closed__4; -static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__30; extern lean_object* l_Lean_PrettyPrinter_parenthesizerAttribute; -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_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__93; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__43; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__52; 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_mkAntiquot_parenthesizer___closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_ppHardLineUnlessUngrouped_parenthesizer___rarg(lean_object*); extern lean_object* l_Lean_Parser_rawIdentNoAntiquot; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__51; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__10; static lean_object* l_Lean_Parser_sepByIndent_parenthesizer___closed__3; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__12; LEAN_EXPORT lean_object* l_Lean_Parser_adaptCacheableContext_formatter___rarg(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__18; static lean_object* l_Lean_Parser_sepByElemParser_formatter___closed__3; LEAN_EXPORT lean_object* l_Lean_ppHardSpace_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__68; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__77; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__37; lean_object* l_Lean_quoteNameMk(lean_object*); static lean_object* l_Lean_Parser_antiquotNestedExpr_parenthesizer___closed__3; static lean_object* l_Lean_Parser_nameLit___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__63; static lean_object* l_Lean_Parser_antiquotNestedExpr_parenthesizer___closed__4; -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__41; +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__9; lean_object* lean_st_ref_get(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__54; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__71; LEAN_EXPORT lean_object* l_Lean_Parser_manyIndent_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___lambda__1___closed__42; +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; LEAN_EXPORT lean_object* l_Lean_ppLine_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___lambda__1___closed__46; lean_object* l_Lean_PrettyPrinter_Formatter_concat___lambda__1___boxed(lean_object*); static lean_object* l_Lean_Parser_nameLit_parenthesizer___closed__1; lean_object* l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer(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_2047____closed__10; lean_object* l_Lean_PrettyPrinter_Formatter_nameLitNoAntiquot_formatter(lean_object*, lean_object*, lean_object*, 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_2047____closed__76; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__14; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__36; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__59; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__38; +lean_object* l_Lean_PrettyPrinter_Formatter_hygieneInfoNoAntiquot_formatter___boxed(lean_object*); lean_object* l_Lean_Syntax_getOptional_x3f(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ppHardLineUnlessUngrouped_parenthesizer___boxed(lean_object*, 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_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__86; static lean_object* l_Lean_Parser_termParser_formatter___rarg___closed__2; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__22; lean_object* l_Lean_Parser_notFollowedBy(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__34; lean_object* l_Lean_PrettyPrinter_Formatter_pushWhitespace(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_symbol(lean_object*); static lean_object* l_Lean_Parser_sepByIndent___closed__4; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__40; 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_withForbidden_parenthesizer___boxed(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__5; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__31; LEAN_EXPORT lean_object* l_Lean_Parser_ppHardLineUnlessUngrouped_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Parser_sepByIndent_formatter___spec__3___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ppDedentIfGrouped_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_node3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ppGroup___boxed(lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__57; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__36; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__22; static lean_object* l_Lean_Parser_commandParser_formatter___rarg___closed__2; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__37; static lean_object* l_Lean_Parser_mkAntiquotSplice_formatter___closed__2; static lean_object* l_Lean_Parser_nameLit___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_symbol_formatter(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__1___closed__17; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__40; LEAN_EXPORT lean_object* l_Lean_Parser_withResetCache_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PrettyPrinter_Parenthesizer_checkLinebreakBefore_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ppSpace; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__66; static lean_object* l_Lean_Parser_optional_parenthesizer___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_nameLit_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_nameLit; -static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__24; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__45; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__116; +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_2065____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__8; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__116; static lean_object* l_Lean_Parser_mkAntiquot_formatter___closed__10; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__27; LEAN_EXPORT lean_object* l_Lean_Parser_suppressInsideQuot_formatter(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__1___closed__33; 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*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__26; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__5; static lean_object* l_Lean_Parser_sepByIndent___closed__2; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__64; +static lean_object* l_Lean_Parser_hygieneInfo_formatter___closed__4; static lean_object* l_Lean_Parser_scientificLit___closed__1; static lean_object* l_Lean_Parser_mkAntiquot_parenthesizer___closed__9; static lean_object* l_Lean_Parser_strLit_formatter___closed__4; 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*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__97; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__11; static lean_object* l_Lean_Parser_scientificLit_parenthesizer___closed__2; lean_object* l_Lean_Parser_withAntiquot(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__23; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__6; static lean_object* l_Lean_Parser_sepByIndent___closed__3; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__94; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__26; lean_object* l_Lean_addMacroScope(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___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__16; static lean_object* l_Lean_Parser_ident___closed__1; static lean_object* l___regBuiltin_Lean_Parser_antiquotNestedExpr_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__1___closed__35; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__107; 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_2047____closed__19; LEAN_EXPORT lean_object* l_Lean_Parser_antiquotExpr_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PrettyPrinter_Parenthesizer_checkNoImmediateColon_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PrettyPrinter_Formatter_identNoAntiquot_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__96; LEAN_EXPORT lean_object* l_Lean_Parser_ppAllowUngrouped; +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_2065____closed__29; static lean_object* l_Lean_Parser_mkAntiquot_formatter___closed__4; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__41; +static lean_object* l_Lean_Parser_hygieneInfo_formatter___closed__1; lean_object* l_Lean_PrettyPrinter_Parenthesizer_nonReservedSymbolNoAntiquot_parenthesizer___boxed(lean_object*, lean_object*, 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___lambda__1___closed__76; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__77; LEAN_EXPORT lean_object* l_Lean_Parser_withForbidden_formatter___boxed(lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__40; lean_object* l_Lean_Syntax_node2(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__60; 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___lambda__1___closed__64; static lean_object* l_Lean_Parser_many___closed__1; static lean_object* l_Lean_Parser_mkAntiquot_formatter___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__115; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__22; LEAN_EXPORT lean_object* l_Lean_Parser_suppressInsideQuot_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_nameLit_formatter___closed__2; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__40; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__54; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_many_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_withResetCache_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_sepByElemParser_formatter___closed__2; @@ -436,29 +446,32 @@ LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_addQuotDep LEAN_EXPORT lean_object* l_Lean_Parser_nonReservedSymbol_parenthesizer(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); 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*); -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__76; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__108; uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); static lean_object* l_Lean_Parser_mkAntiquot_formatter___closed__11; static lean_object* l_Lean_Parser_antiquotNestedExpr_parenthesizer___closed__2; static lean_object* l_Lean_Parser_nameLit_formatter___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_adaptCacheableContext_parenthesizer(lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__54; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__52; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__98; static lean_object* l_Lean_Parser_ident___closed__2; lean_object* l_Lean_Parser_many1NoAntiquot(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_charLit_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__69; lean_object* l_Lean_Parser_node(lean_object*, lean_object*); lean_object* l_Lean_PrettyPrinter_Parenthesizer_symbolNoAntiquot_parenthesizer___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Parser_numLit___closed__2; static lean_object* l_Lean_Parser_mkAntiquot_formatter___closed__12; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__54; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__42; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__75; extern lean_object* l_Lean_Parser_numLitNoAntiquot; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__38; static lean_object* l_Lean_Parser_mkAntiquotSplice_formatter___closed__8; static lean_object* l_Lean_Parser_patternIgnore_formatter___closed__1; static lean_object* l_Lean_Parser_many1Indent___closed__2; -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_2065____closed__74; static lean_object* l_Lean_Parser_sepByIndent_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_2065____closed__23; +LEAN_EXPORT lean_object* l_Lean_Parser_hygieneInfo_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_sepBy1Indent___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_charLit_parenthesizer___closed__2; static lean_object* l_List_forIn_loop___at_Lean_Parser_sepByIndent_formatter___spec__2___closed__2; @@ -467,83 +480,87 @@ LEAN_EXPORT lean_object* l_Array_mapIdxM_map___at_Lean_Parser_sepByIndent_format static lean_object* l_Lean_Parser_antiquotNestedExpr_formatter___closed__7; LEAN_EXPORT lean_object* l_Lean_ppHardSpace_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_rawIdent_formatter___closed__1; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__63; 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___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__24; LEAN_EXPORT lean_object* l_Lean_Parser_withForbidden_formatter___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_andthen(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__70; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_ppHardSpace; lean_object* l_Lean_PrettyPrinter_Parenthesizer_manyNoAntiquot_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__58; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__62; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__79; LEAN_EXPORT lean_object* l_Lean_Parser_withForbidden_parenthesizer(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__62; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__58; LEAN_EXPORT lean_object* l_Lean_Parser_setExpected_parenthesizer___boxed(lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_antiquotNestedExpr_parenthesizer(lean_object*); lean_object* l_Lean_Parser_checkColGe(lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__107; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__27; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__44; static lean_object* l_Lean_Parser_charLit_formatter___closed__3; static lean_object* l_Lean_Parser_optional_formatter___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_evalInsideQuot_parenthesizer(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ppSpace_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_int_sub(lean_object*, lean_object*); static lean_object* l_Lean_Parser_numLit_formatter___closed__1; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__46; LEAN_EXPORT lean_object* l_Lean_Parser_ident_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_charLit; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__119; 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*); LEAN_EXPORT lean_object* l_Lean_Parser_ppRealFill(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_charLit_formatter___closed__2; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__118; static lean_object* l_Lean_Parser_sepByIndent_parenthesizer___closed__4; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__39; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__8; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__71; static lean_object* l_Lean_Parser_termParser_formatter___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_rawIdent; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__38; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__32; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__11; +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; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__8; LEAN_EXPORT lean_object* l_Lean_Parser_evalInsideQuot_formatter(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__27; extern lean_object* l_Lean_PrettyPrinter_Formatter_formatterAliasesRef; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__66; LEAN_EXPORT lean_object* l_Lean_Parser_ppIndent(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_commandParser_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__1___closed__19; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__51; static lean_object* l_Lean_Parser_antiquotNestedExpr_formatter___closed__9; 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*); +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__40; LEAN_EXPORT lean_object* l_Lean_Parser_scientificLit; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__29; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__26; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__57; LEAN_EXPORT lean_object* l_Lean_Parser_many(lean_object*); LEAN_EXPORT lean_object* l_Lean_ppDedent_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_numLit_parenthesizer___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__12; 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_2047____closed__48; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__73; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__45; static lean_object* l_Lean_Parser_many1Indent___closed__1; +static lean_object* l_Lean_Parser_hygieneInfo___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__24; lean_object* l_Lean_PrettyPrinter_Formatter_pushNone_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__1___closed__51; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__16; LEAN_EXPORT lean_object* l_Lean_ppAllowUngrouped_formatter___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_many_formatter___closed__2; -static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__34; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__86; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__18; static lean_object* l_Lean_Parser_antiquotNestedExpr_formatter___closed__2; static lean_object* l_Lean_Parser_scientificLit___closed__2; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Parser_sepByIndent_formatter___spec__3(lean_object*, size_t, size_t); 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_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__109; static lean_object* l_Lean_Parser_mkAntiquot_parenthesizer___closed__1; static lean_object* l_Lean_Parser_mkAntiquotSplice_formatter___closed__4; -static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__33; LEAN_EXPORT lean_object* l_Lean_Parser_ppLine; LEAN_EXPORT lean_object* l_Lean_ppSpace_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__27; 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___lambda__1___closed__27; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__17; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__42; static lean_object* l_Lean_Parser_mkAntiquot_formatter___closed__2; -LEAN_EXPORT lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29________; LEAN_EXPORT lean_object* l_Lean_Parser_withPosition_formatter(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_Lean_Parser_hygieneInfo_parenthesizer___closed__1; lean_object* l_Lean_PrettyPrinter_Parenthesizer_withoutInfo_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PrettyPrinter_Parenthesizer_withAntiquot_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_adaptCacheableContext_formatter___boxed(lean_object*); @@ -551,154 +568,163 @@ static lean_object* l_Lean_Parser_optional_formatter___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_addQuotDepth_formatter___boxed(lean_object*); lean_object* l_Lean_Parser_registerAliasCore___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_patternIgnore_formatter(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__1___closed__7; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__56; LEAN_EXPORT lean_object* l_Lean_Parser_termParser_formatter(lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__69; LEAN_EXPORT lean_object* l_Lean_Parser_ident; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__23; LEAN_EXPORT lean_object* l_Lean_Parser_optional_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_unicodeSymbol_formatter(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*); lean_object* l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_many1Indent(lean_object*); static lean_object* l_Lean_Parser_mkAntiquotSplice_formatter___closed__5; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__60; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__109; 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*); lean_object* l_Lean_PrettyPrinter_Formatter_group(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_scientificLit_formatter___closed__2; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__9; 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_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__119; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__53; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__81; static lean_object* l_Lean_Parser_strLit_formatter___closed__1; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__24; LEAN_EXPORT lean_object* l_Lean_Parser_ppLine_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__81; LEAN_EXPORT lean_object* l_Lean_Parser_withoutForbidden_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PrettyPrinter_Parenthesizer_scientificLitNoAntiquot_parenthesizer___boxed(lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); lean_object* l_Lean_PrettyPrinter_Formatter_visitArgs(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__31; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__3; static lean_object* l_Lean_Parser_antiquotNestedExpr_parenthesizer___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_ppAllowUngrouped_parenthesizer___rarg(lean_object*); lean_object* l_Lean_PrettyPrinter_Formatter_fold(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_dbgTraceState_formatter___boxed(lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__108; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__100; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__62; lean_object* lean_nat_mod(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__83; LEAN_EXPORT lean_object* l_Lean_Parser_withCache_formatter___boxed(lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__64; LEAN_EXPORT lean_object* l_Lean_Parser_sepByElemParser_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_2047____closed__46; -static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__25; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__70; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_evalInsideQuot_formatter___boxed(lean_object*); lean_object* l_Lean_PrettyPrinter_Formatter_checkNoImmediateColon_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__3; static lean_object* l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_atomic_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_ident_formatter___closed__3; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); 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___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__55; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__56; extern lean_object* l_Lean_PrettyPrinter_Parenthesizer_parenthesizerAliasesRef; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__31; lean_object* l_Lean_Parser_withAntiquotSpliceAndSuffix(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_nameLit_formatter___closed__4; extern lean_object* l_Lean_Parser_skip; LEAN_EXPORT lean_object* l_Lean_Parser_decQuotDepth_formatter(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__24; static lean_object* l_Lean_Parser_charLit___closed__2; 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___regBuiltin_Lean_Parser_antiquotNestedExpr_parenthesizer___closed__4; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__89; LEAN_EXPORT lean_object* l_Lean_ppHardLineUnlessUngrouped_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_withoutPosition_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047_(lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__14; static lean_object* l_Lean_Parser_mkAntiquot_formatter___closed__16; -static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__12; extern lean_object* l_Lean_Parser_scientificLitNoAntiquot; LEAN_EXPORT 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_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__35; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__1; -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_2065____closed__112; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__9; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__13; extern lean_object* l_Lean_Parser_strLitNoAntiquot; lean_object* l_Lean_Syntax_node1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_leadingNode_formatter___closed__2; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__65; lean_object* l_Lean_Parser_registerAlias(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__1___closed__65; +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*); 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_2047____closed__4; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__1; +static lean_object* l_Lean_Parser_hygieneInfo___closed__1; LEAN_EXPORT lean_object* l_Lean_ppAllowUngrouped_formatter___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_patternIgnore(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_dbgTraceState_parenthesizer(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_2065____closed__33; static lean_object* l_Lean_Parser_mkAntiquotSplice_parenthesizer___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__45; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__59; LEAN_EXPORT lean_object* l_Lean_Parser_many1(lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__80; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__34; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__7; uint8_t l_Lean_Syntax_isNone(lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__39; lean_object* l_Lean_PrettyPrinter_Formatter_withAntiquotSuffixSplice_formatter___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_group_formatter___closed__2; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__124; +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_2065____closed__25; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__114; static lean_object* l_Lean_Parser_strLit_parenthesizer___closed__2; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__118; LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_addQuotDepth_parenthesizer___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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_charLit_parenthesizer___closed__1; static lean_object* l_Lean_Parser_commandParser_formatter___rarg___closed__1; lean_object* lean_nat_sub(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__21; lean_object* l_Std_Format_getIndent(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__67; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__70; static lean_object* l_Lean_Parser_mkAntiquot_formatter___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__44; 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___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__69; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__88; extern lean_object* l_Lean_PrettyPrinter_backtrackExceptionId; LEAN_EXPORT lean_object* l_Lean_Parser_termParser_parenthesizer(lean_object*, 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*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__75; LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_addQuotDepth_formatter(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ppRealGroup(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__17; lean_object* l_Lean_PrettyPrinter_Formatter_andthen_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_setExpected_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__1___closed__11; LEAN_EXPORT lean_object* l_Lean_Parser_ppDedent(lean_object*); 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*); lean_object* l_Lean_PrettyPrinter_Formatter_checkColGe_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_numLit_formatter___closed__2; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__104; extern lean_object* l_Lean_Parser_maxPrec; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__47; lean_object* l_Lean_Parser_sepBy(lean_object*, lean_object*, lean_object*, uint8_t); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__87; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__49; lean_object* l_Lean_PrettyPrinter_Formatter_strLitNoAntiquot_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Parser_identNoAntiquot; -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_2065____closed__84; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__60; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__71; static lean_object* l_Lean_Parser_nameLit_parenthesizer___closed__2; lean_object* l_Lean_PrettyPrinter_Parenthesizer_charLitNoAntiquot_parenthesizer___boxed(lean_object*); static lean_object* l_Lean_Parser_antiquotExpr_parenthesizer___closed__1; lean_object* l_List_reverse___rarg(lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__50; 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___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__25; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__35; 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_object* l_Lean_Syntax_MonadTraverser_goLeft___at_Lean_PrettyPrinter_Formatter_visitArgs___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__74; lean_object* l_String_intercalate(lean_object*, lean_object*); static lean_object* l_Lean_Parser_mkAntiquot_formatter___closed__7; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__120; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__14; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__30; 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*); -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__75; lean_object* l_Lean_PrettyPrinter_Parenthesizer_tokenWithAntiquot_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_2065____closed__19; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__26; size_t lean_usize_add(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Parser_sepByIndent(lean_object*, lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Parser_ppDedentIfGrouped(lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____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__6; 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*, lean_object*); static lean_object* l_Lean_Parser_many_parenthesizer___closed__1; LEAN_EXPORT lean_object* l_Lean_ppRealFill_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_dbgTraceState_parenthesizer___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__97; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__105; lean_object* l_Lean_Parser_mkAntiquot(lean_object*, lean_object*, uint8_t, uint8_t); 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_many_formatter___closed__1; static lean_object* l_Lean_Parser_optional_formatter___closed__4; lean_object* lean_array_uget(lean_object*, size_t); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__33; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__111; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__79; 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_numLit_formatter___closed__4; 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*); @@ -707,9 +733,9 @@ LEAN_EXPORT lean_object* l_Lean_Parser_incQuotDepth_parenthesizer(lean_object*, static lean_object* l_Lean_Parser_sepByIndent_formatter___closed__1; lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_ident_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__1___closed__29; LEAN_EXPORT lean_object* l_Lean_Parser_numLit; 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*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__34; LEAN_EXPORT lean_object* l_Lean_Parser_ppIndent___boxed(lean_object*); extern lean_object* l_Lean_PrettyPrinter_formatterAttribute; static lean_object* l_Lean_Parser_mkAntiquotSplice_formatter___closed__1; @@ -717,101 +743,93 @@ LEAN_EXPORT lean_object* l_Lean_ppIndent_formatter(lean_object*, lean_object*, l lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); 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_scientificLit_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__1___closed__23; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__63; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__29; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__47; static lean_object* l_Lean_Parser_optional_formatter___closed__2; lean_object* l_Lean_PrettyPrinter_Parenthesizer_rawIdentNoAntiquot_parenthesizer___boxed(lean_object*); 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_mkAntiquot_formatter___closed__15; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__84; -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__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__1___closed__64; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__30; lean_object* lean_string_append(lean_object*, lean_object*); 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_2047____closed__21; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__83; LEAN_EXPORT lean_object* l_Lean_Parser_group_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_2047____closed__102; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__81; LEAN_EXPORT lean_object* l_Lean_Parser_numLit_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__96; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__43; LEAN_EXPORT lean_object* l___private_Lean_Parser_Basic_0__Lean_Parser_addQuotDepth_formatter___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_sepByIndent___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__24; lean_object* l_Lean_PrettyPrinter_Formatter_symbolNoAntiquot_formatter(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__1___closed__81; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__61; lean_object* lean_array_get_size(lean_object*); -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_2065____closed__120; static lean_object* l_Lean_Parser_mkAntiquotSplice_formatter___closed__3; static lean_object* l_Lean_Parser_numLit___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__101; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__56; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__15; static lean_object* l_Lean_Parser_mkAntiquot_parenthesizer___closed__3; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__72; uint8_t lean_nat_dec_le(lean_object*, lean_object*); 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___lambda__1___closed__46; lean_object* l_Lean_PrettyPrinter_Formatter_categoryParser_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_withOpenDecl_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_manyNoAntiquot(lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__70; 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___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__53; 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_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__3; LEAN_EXPORT lean_object* l_Lean_ppAllowUngrouped_formatter(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ppDedentIfGrouped_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_2047____closed__22; static lean_object* l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__4; static lean_object* l_Lean_ppDedentIfGrouped_formatter___closed__1; 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_2065____closed__94; lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_evalInsideQuot_formatter___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_mkAntiquotSplice_formatter___closed__6; -static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__15; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__31; static lean_object* l_Lean_Parser_numLit_formatter___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_sepByIndent___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__111; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__69; -static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__34; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__41; static lean_object* l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_patternIgnore_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PrettyPrinter_Parenthesizer_numLitNoAntiquot_parenthesizer___boxed(lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__88; +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_many1Indent_parenthesizer___closed__1; static lean_object* l_Lean_Parser_antiquotNestedExpr_formatter___closed__10; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__35; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__92; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__82; -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__1___closed__15; +static lean_object* l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__34; +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; lean_object* l_String_toSubstring_x27(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__20; LEAN_EXPORT lean_object* l_Lean_Parser_withoutForbidden_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_2047____closed__99; static lean_object* l_Lean_Parser_mkAntiquot_formatter___closed__9; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__36; LEAN_EXPORT lean_object* l_Lean_Parser_dbgTraceState_formatter(lean_object*); static lean_object* l_Lean_Parser_sepByIndent___closed__5; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__49; +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_antiquotNestedExpr_parenthesizer___closed__1; +lean_object* l_Lean_PrettyPrinter_Parenthesizer_hygieneInfoNoAntiquot_parenthesizer___boxed(lean_object*); static lean_object* l_Lean_Parser_mkAntiquot_parenthesizer___closed__6; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__78; static lean_object* l_Lean_Parser_mkAntiquot_formatter___closed__17; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__50; +static lean_object* l_Lean_Parser_hygieneInfo_formatter___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_ppDedentIfGrouped___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_adaptCacheableContext_parenthesizer___boxed(lean_object*); 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_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__30; +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_initFn____x40_Lean_Parser_Extra___hyg_2047____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__72; static lean_object* l_Lean_Parser_antiquotExpr_formatter___closed__2; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__73; +static lean_object* l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__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__1___closed__7; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__60; static lean_object* l_Lean_Parser_antiquotExpr_formatter___closed__1; static lean_object* l_Lean_Parser_antiquotNestedExpr_formatter___closed__5; 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___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__4; 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_Formatter_node_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__61; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__59; lean_object* l_Lean_PrettyPrinter_Parenthesizer_nameLitNoAntiquot_parenthesizer___boxed(lean_object*); static lean_object* l_Lean_Parser_mkAntiquot_formatter___closed__3; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__113; lean_object* l_Lean_PrettyPrinter_Parenthesizer_pushNone_parenthesizer___boxed(lean_object*); static lean_object* l_Lean_Parser_numLit_parenthesizer___closed__2; static lean_object* _init_l_Lean_Parser_leadingNode_formatter___closed__1() { @@ -3242,6 +3260,123 @@ x_1 = l_Lean_Parser_rawIdent___closed__1; return x_1; } } +static lean_object* _init_l_Lean_Parser_hygieneInfo_formatter___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("hygieneInfo", 11); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_hygieneInfo_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_hygieneInfo_formatter___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_hygieneInfo_formatter___closed__3() { +_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_hygieneInfo_formatter___closed__1; +x_2 = l_Lean_Parser_hygieneInfo_formatter___closed__2; +x_3 = 0; +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_hygieneInfo_formatter___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_hygieneInfoNoAntiquot_formatter___boxed), 1, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Parser_hygieneInfo_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_hygieneInfo_formatter___closed__3; +x_7 = l_Lean_Parser_hygieneInfo_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_hygieneInfo_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_hygieneInfo_formatter___closed__1; +x_2 = l_Lean_Parser_hygieneInfo_formatter___closed__2; +x_3 = 0; +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_hygieneInfo_parenthesizer___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_hygieneInfoNoAntiquot_parenthesizer___boxed), 1, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Parser_hygieneInfo_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_hygieneInfo_parenthesizer___closed__1; +x_7 = l_Lean_Parser_hygieneInfo_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_hygieneInfo___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; uint8_t x_3; lean_object* x_4; +x_1 = l_Lean_Parser_hygieneInfo_formatter___closed__1; +x_2 = l_Lean_Parser_hygieneInfo_formatter___closed__2; +x_3 = 0; +x_4 = l_Lean_Parser_mkAntiquot(x_1, x_2, x_3, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_hygieneInfo___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_hygieneInfo___closed__1; +x_2 = l_Lean_Parser_hygieneInfoNoAntiquot; +x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_hygieneInfo() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Parser_hygieneInfo___closed__2; +return x_1; +} +} static lean_object* _init_l_Lean_Parser_numLit_formatter___closed__1() { _start: { @@ -6047,26 +6182,26 @@ 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:=_)____", 38); +x_1 = lean_mk_string_from_bytes("termRegister_parser_alias(Kind:=_)______", 40); 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; lean_object* x_4; x_1 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__1; x_2 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__2; -x_3 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__1; +x_3 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -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; @@ -6074,35 +6209,35 @@ 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 = l_Lean_Name_str___override(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; -x_1 = lean_mk_string_from_bytes("register_parser_alias", 21); +x_1 = lean_mk_string_from_bytes("register_parser_alias ", 22); 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() { +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; @@ -6112,7 +6247,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__8() { +static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__8() { _start: { lean_object* x_1; @@ -6120,11 +6255,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__9() { +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__8; +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); @@ -6132,13 +6267,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__10() { +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__7; -x_3 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__9; +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); @@ -6146,7 +6281,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__11() { +static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__11() { _start: { lean_object* x_1; @@ -6154,23 +6289,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__12() { +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__11; +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__13() { +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__10; -x_3 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__12; +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); @@ -6178,7 +6313,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__14() { +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; @@ -6190,13 +6325,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__15() { +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__13; -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__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); @@ -6204,23 +6339,31 @@ 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__16() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(") ", 2); +return x_1; +} +} +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; -x_1 = l_Lean_Parser_antiquotNestedExpr_formatter___closed__7; +x_1 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________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_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__18() { _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__15; +x_3 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________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); @@ -6228,37 +6371,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__18() { +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_group_formatter___closed__2; -x_2 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__17; +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__19() { +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; x_1 = l_Lean_Parser_optional_formatter___closed__2; -x_2 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__18; +x_2 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________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_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__21() { _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__19; +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__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); @@ -6266,7 +6409,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__21() { +static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__22() { _start: { lean_object* x_1; @@ -6274,46 +6417,89 @@ 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__22() { +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; lean_object* x_3; lean_object* x_4; x_1 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__1; x_2 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__2; -x_3 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__21; +x_3 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__22; x_4 = l_Lean_Name_mkStr3(x_1, x_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__24() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__22; +x_1 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__23; 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__24() { +static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__25() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("ppSpace", 7); +return x_1; +} +} +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; lean_object* x_4; +x_1 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__1; +x_2 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__2; +x_3 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__25; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +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__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__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__24; +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; 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__23; +x_2 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__28; 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__25() { +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; 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__20; -x_3 = 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__4; +x_2 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__21; +x_3 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________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); @@ -6321,7 +6507,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__26() { +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; lean_object* x_3; lean_object* x_4; @@ -6332,23 +6518,23 @@ x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -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__32() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = 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__31; 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__28() { +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; 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_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__30; +x_3 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__32; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -6356,7 +6542,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__29() { +static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__34() { _start: { lean_object* x_1; @@ -6364,33 +6550,33 @@ 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() { +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; x_1 = lean_box(0); -x_2 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__29; +x_2 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__34; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -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__36() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__30; +x_1 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________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_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__32() { +static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__37() { _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_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__36; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -6398,25 +6584,39 @@ 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() { +static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__38() { +_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__37; +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__39() { _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_2 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__38; 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() { +static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__40() { _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__28; -x_3 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29___________closed__33; +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__33; +x_3 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__39; x_4 = lean_alloc_ctor(2, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -6424,13 +6624,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__35() { +static lean_object* _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__41() { _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__34; +x_3 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__40; x_4 = lean_alloc_ctor(3, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -6438,15 +6638,15 @@ 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__35; +x_1 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__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___lambda__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; @@ -6454,7 +6654,7 @@ 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___lambda__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; @@ -6462,7 +6662,7 @@ 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___lambda__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__3() { _start: { lean_object* x_1; @@ -6470,19 +6670,19 @@ 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___lambda__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__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___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__1; x_2 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__2; -x_3 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__2; -x_4 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__3; +x_3 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__2; +x_4 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__3; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_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__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__5() { _start: { lean_object* x_1; @@ -6490,19 +6690,19 @@ 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___lambda__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__6() { _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_Parser_antiquotNestedExpr_formatter___closed__1; x_2 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__2; -x_3 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__2; -x_4 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__5; +x_3 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__2; +x_4 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__5; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_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__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__7() { _start: { lean_object* x_1; @@ -6510,19 +6710,19 @@ 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___lambda__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__8() { _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_Parser_antiquotNestedExpr_formatter___closed__1; x_2 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__2; -x_3 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__2; -x_4 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__7; +x_3 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__2; +x_4 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__7; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_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__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__9() { _start: { lean_object* x_1; @@ -6530,19 +6730,19 @@ 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__1___closed__10() { +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; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__1; x_2 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__2; -x_3 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__2; -x_4 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__9; +x_3 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__2; +x_4 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__9; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_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__11() { +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; @@ -6550,19 +6750,19 @@ 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___lambda__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__1___closed__12() { _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_Parser_antiquotNestedExpr_formatter___closed__1; x_2 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__2; -x_3 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__2; -x_4 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__11; +x_3 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__2; +x_4 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__11; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_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__13() { +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__13() { _start: { lean_object* x_1; @@ -6570,16 +6770,16 @@ 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___lambda__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__1___closed__14() { _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___lambda__1___closed__13; +x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__13; x_2 = l_String_toSubstring_x27(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__15() { +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__15() { _start: { lean_object* x_1; @@ -6587,74 +6787,74 @@ 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___lambda__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__1___closed__16() { _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___lambda__1___closed__15; +x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__15; x_3 = l_Lean_Name_mkStr2(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__17() { +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__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__1; x_2 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__2; -x_3 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__15; +x_3 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__15; x_4 = l_Lean_Name_mkStr3(x_1, x_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___lambda__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__1___closed__18() { _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__17; +x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___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_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__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__1___closed__19() { _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___lambda__1___closed__17; +x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___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___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__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__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___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__19; +x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___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_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__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__1___closed__21() { _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___lambda__1___closed__18; -x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__20; +x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__18; +x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___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_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__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__1___closed__22() { _start: { lean_object* x_1; @@ -6662,19 +6862,19 @@ x_1 = lean_mk_string_from_bytes("doubleQuotedName", 16); 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__23() { +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__23() { _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_Parser_antiquotNestedExpr_formatter___closed__1; x_2 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__2; -x_3 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__2; -x_4 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__22; +x_3 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__2; +x_4 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__22; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_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__24() { +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__24() { _start: { lean_object* x_1; @@ -6682,7 +6882,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___lambda__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__1___closed__25() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; @@ -6692,7 +6892,7 @@ x_3 = l_Lean_SourceInfo_fromRef(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__26() { +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__26() { _start: { lean_object* x_1; @@ -6700,17 +6900,17 @@ 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___lambda__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__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_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__26; +x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__26; x_3 = l_Lean_Name_str___override(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__28() { +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__28() { _start: { lean_object* x_1; @@ -6718,17 +6918,17 @@ 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___lambda__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__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___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__28; +x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__28; x_3 = l_Lean_Name_str___override(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__30() { +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__30() { _start: { lean_object* x_1; @@ -6736,19 +6936,19 @@ 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__31() { +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__31() { _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___lambda__1___closed__25; -x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__30; +x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__25; +x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__30; 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_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__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__1___closed__32() { _start: { lean_object* x_1; @@ -6756,31 +6956,31 @@ 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__33() { +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__33() { _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___lambda__1___closed__25; -x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__32; +x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__25; +x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__32; 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_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__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__1___closed__34() { _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___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__25; -x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__29; -x_3 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__31; -x_4 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__33; +x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__25; +x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__29; +x_3 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__31; +x_4 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__33; x_5 = l_Lean_Syntax_node2(x_1, x_2, x_3, x_4); return x_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__35() { +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__35() { _start: { lean_object* x_1; @@ -6788,19 +6988,19 @@ 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___lambda__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__1___closed__36() { _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_Parser_antiquotNestedExpr_formatter___closed__1; x_2 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__2; -x_3 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__2; -x_4 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__35; +x_3 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__2; +x_4 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__35; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_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__37() { +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__37() { _start: { lean_object* x_1; lean_object* x_2; @@ -6809,13 +7009,13 @@ 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__38() { +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__38() { _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___lambda__1___closed__25; +x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__25; 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__37; +x_3 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__37; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -6823,7 +7023,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___lambda__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__1___closed__39() { _start: { lean_object* x_1; @@ -6831,56 +7031,56 @@ 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___lambda__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__1___closed__40() { _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_Parser_antiquotNestedExpr_formatter___closed__1; x_2 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__2; -x_3 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__2; -x_4 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__39; +x_3 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__2; +x_4 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__39; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_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__41() { +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__41() { _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___lambda__1___closed__25; -x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__40; -x_3 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__38; +x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__25; +x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__40; +x_3 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__38; x_4 = l_Lean_Syntax_node1(x_1, x_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___lambda__1___closed__42() { +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__42() { _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___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__25; -x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__36; -x_3 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__31; -x_4 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__38; -x_5 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__41; -x_6 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__33; +x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__25; +x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__36; +x_3 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__31; +x_4 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__38; +x_5 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__41; +x_6 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__33; x_7 = l_Lean_Syntax_node6(x_1, x_2, x_3, x_4, x_4, x_5, x_4, x_6); return x_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__43() { +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__43() { _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___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__25; -x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__27; -x_3 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__34; -x_4 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__42; +x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__25; +x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__27; +x_3 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__34; +x_4 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__42; x_5 = l_Lean_Syntax_node2(x_1, x_2, x_3, x_4); return x_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__44() { +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__44() { _start: { lean_object* x_1; @@ -6888,19 +7088,19 @@ 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___lambda__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__1___closed__45() { _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_Parser_antiquotNestedExpr_formatter___closed__1; x_2 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__2; -x_3 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__2; -x_4 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__44; +x_3 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__2; +x_4 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__44; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_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__46() { +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__46() { _start: { lean_object* x_1; @@ -6908,26 +7108,26 @@ 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___lambda__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__1___closed__47() { _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___lambda__1___closed__46; +x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__46; x_2 = l_String_toSubstring_x27(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__48() { +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__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___lambda__1___closed__46; +x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__46; x_3 = l_Lean_Name_str___override(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__49() { +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__49() { _start: { lean_object* x_1; @@ -6935,7 +7135,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___lambda__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__1___closed__50() { _start: { lean_object* x_1; @@ -6943,26 +7143,26 @@ 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___lambda__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__1___closed__51() { _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___lambda__1___closed__50; +x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__50; x_2 = l_String_toSubstring_x27(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__52() { +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__52() { _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__50; +x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__50; x_3 = l_Lean_Name_str___override(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__53() { +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__53() { _start: { lean_object* x_1; @@ -6970,41 +7170,41 @@ 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___lambda__1___closed__54() { +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__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___lambda__1___closed__53; -x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__50; +x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__53; +x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__50; x_3 = l_Lean_Name_mkStr2(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__55() { +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__55() { _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__54; +x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__54; 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___lambda__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__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___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__55; +x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___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_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__57() { +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__57() { _start: { lean_object* x_1; @@ -7012,16 +7212,16 @@ 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___lambda__1___closed__58() { +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__58() { _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___lambda__1___closed__57; +x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__57; x_2 = l_String_toSubstring_x27(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__59() { +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__59() { _start: { lean_object* x_1; @@ -7029,7 +7229,7 @@ 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___lambda__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__1___closed__60() { _start: { lean_object* x_1; @@ -7037,76 +7237,76 @@ 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___lambda__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__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___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__59; -x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__60; -x_3 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__15; +x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__59; +x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__60; +x_3 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__15; x_4 = l_Lean_Name_mkStr3(x_1, x_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___lambda__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__1___closed__62() { _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_Parser_antiquotNestedExpr_formatter___closed__1; -x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__59; -x_3 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__60; -x_4 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__15; +x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__59; +x_3 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__60; +x_4 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__15; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_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__63() { +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__63() { _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__62; +x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__62; 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___lambda__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__1___closed__64() { _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___lambda__1___closed__62; +x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__62; 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___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__65() { +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__65() { _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__64; +x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__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___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__66() { +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__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___lambda__1___closed__63; -x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__65; +x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__63; +x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__65; 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___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__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__1___closed__67() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -7116,7 +7316,7 @@ x_3 = l_Lean_Name_str___override(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__68() { +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__68() { _start: { lean_object* x_1; @@ -7124,16 +7324,16 @@ 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___lambda__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__1___closed__69() { _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___lambda__1___closed__68; +x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__68; x_2 = l_String_toSubstring_x27(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__70() { +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__70() { _start: { lean_object* x_1; @@ -7141,76 +7341,76 @@ 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___lambda__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__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___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__59; -x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__70; -x_3 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__15; +x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__59; +x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__70; +x_3 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__15; x_4 = l_Lean_Name_mkStr3(x_1, x_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___lambda__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__1___closed__72() { _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_Parser_antiquotNestedExpr_formatter___closed__1; -x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__59; -x_3 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__70; -x_4 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__15; +x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__59; +x_3 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__70; +x_4 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__15; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_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__73() { +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__73() { _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__72; +x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___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_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__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__1___closed__74() { _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___lambda__1___closed__72; +x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__72; 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___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__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__1___closed__75() { _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__74; +x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__74; 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___lambda__1___closed__76() { +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__76() { _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___lambda__1___closed__73; -x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__75; +x_1 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__73; +x_2 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__75; 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___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__77() { +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__77() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -7220,7 +7420,7 @@ x_3 = l_Lean_Name_str___override(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__78() { +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__78() { _start: { lean_object* x_1; @@ -7228,19 +7428,19 @@ 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___lambda__1___closed__79() { +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__79() { _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_Parser_antiquotNestedExpr_formatter___closed__1; x_2 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__2; -x_3 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__2; -x_4 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__78; +x_3 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__2; +x_4 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__78; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_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__80() { +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__80() { _start: { lean_object* x_1; @@ -7248,7 +7448,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___lambda__1___closed__81() { +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__81() { _start: { lean_object* x_1; lean_object* x_2; @@ -7257,7 +7457,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -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) { +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; @@ -7427,7 +7627,7 @@ lean_dec(x_16); lean_dec(x_14); lean_dec(x_9); lean_dec(x_3); -x_20 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__1; +x_20 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__1; x_21 = l_Lean_Macro_throwError___rarg(x_20, x_4, x_19); lean_dec(x_4); return x_21; @@ -7481,7 +7681,7 @@ lean_dec(x_16); lean_dec(x_14); lean_dec(x_9); lean_dec(x_3); -x_169 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__1; +x_169 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__1; x_170 = l_Lean_Macro_throwError___rarg(x_169, x_4, x_19); lean_dec(x_4); return x_170; @@ -7498,7 +7698,7 @@ lean_dec(x_16); lean_dec(x_14); lean_dec(x_9); lean_dec(x_3); -x_171 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__1; +x_171 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__1; x_172 = l_Lean_Macro_throwError___rarg(x_171, x_4, x_19); lean_dec(x_4); return x_172; @@ -7515,30 +7715,30 @@ lean_inc(x_30); x_31 = lean_ctor_get(x_4, 1); lean_inc(x_31); lean_dec(x_4); -x_32 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__3; +x_32 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__3; lean_inc(x_29); x_33 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_33, 0, x_29); lean_ctor_set(x_33, 1, x_32); -x_34 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__16; +x_34 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__16; lean_inc(x_30); lean_inc(x_31); x_35 = l_Lean_addMacroScope(x_31, x_34, x_30); x_36 = lean_box(0); -x_37 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__14; -x_38 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__21; +x_37 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__14; +x_38 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__21; lean_inc(x_29); x_39 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_39, 0, x_29); lean_ctor_set(x_39, 1, x_37); lean_ctor_set(x_39, 2, x_35); lean_ctor_set(x_39, 3, x_38); -x_40 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__24; +x_40 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__24; lean_inc(x_29); x_41 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_41, 0, x_29); lean_ctor_set(x_41, 1, x_40); -x_42 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__23; +x_42 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__23; lean_inc(x_9); lean_inc(x_41); lean_inc(x_29); @@ -7548,28 +7748,28 @@ lean_inc(x_29); x_45 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_45, 0, x_29); lean_ctor_set(x_45, 1, x_44); -x_46 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__48; +x_46 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__48; lean_inc(x_30); lean_inc(x_31); x_47 = l_Lean_addMacroScope(x_31, x_46, x_30); -x_48 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__47; +x_48 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__47; lean_inc(x_29); x_49 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_49, 0, x_29); lean_ctor_set(x_49, 1, x_48); lean_ctor_set(x_49, 2, x_47); lean_ctor_set(x_49, 3, x_36); -x_50 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__49; +x_50 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__49; lean_inc(x_29); x_51 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_51, 0, x_29); lean_ctor_set(x_51, 1, x_50); -x_52 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__52; +x_52 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__52; lean_inc(x_30); lean_inc(x_31); x_53 = l_Lean_addMacroScope(x_31, x_52, x_30); -x_54 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__51; -x_55 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__56; +x_54 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__51; +x_55 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__56; lean_inc(x_29); x_56 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_56, 0, x_29); @@ -7584,25 +7784,25 @@ 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_mkAntiquotSplice_formatter___closed__6; -x_61 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__37; +x_61 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__37; lean_inc(x_29); x_62 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_62, 0, x_29); lean_ctor_set(x_62, 1, x_60); lean_ctor_set(x_62, 2, x_61); -x_63 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__61; +x_63 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__61; lean_inc(x_30); lean_inc(x_31); x_64 = l_Lean_addMacroScope(x_31, x_63, x_30); -x_65 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__58; -x_66 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__66; +x_65 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__58; +x_66 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__66; lean_inc(x_29); x_67 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_67, 0, x_29); lean_ctor_set(x_67, 1, x_65); lean_ctor_set(x_67, 2, x_64); lean_ctor_set(x_67, 3, x_66); -x_68 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__67; +x_68 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__67; lean_inc(x_14); x_69 = l_Lean_Name_append(x_14, x_68); lean_inc(x_9); @@ -7610,27 +7810,27 @@ x_70 = l_Lean_mkIdentFrom(x_9, x_69, x_28); lean_inc(x_26); lean_inc(x_29); x_71 = l_Lean_Syntax_node2(x_29, x_60, x_26, x_70); -x_72 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__12; +x_72 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__12; lean_inc(x_29); x_73 = l_Lean_Syntax_node2(x_29, x_72, x_67, x_71); -x_74 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__10; +x_74 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__10; lean_inc(x_29); x_75 = l_Lean_Syntax_node1(x_29, x_74, x_73); -x_76 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__8; +x_76 = 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_62); lean_inc(x_29); x_77 = l_Lean_Syntax_node2(x_29, x_76, x_75, x_62); -x_78 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__71; +x_78 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__71; x_79 = l_Lean_addMacroScope(x_31, x_78, x_30); -x_80 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__69; -x_81 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__76; +x_80 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__69; +x_81 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__76; lean_inc(x_29); x_82 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_82, 0, x_29); 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_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__77; +x_83 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__77; x_84 = l_Lean_Name_append(x_14, x_83); lean_inc(x_9); x_85 = l_Lean_mkIdentFrom(x_9, x_84, x_28); @@ -7647,7 +7847,7 @@ x_89 = l_Lean_Syntax_node2(x_29, x_76, x_88, x_62); if (lean_obj_tag(x_16) == 0) { lean_object* x_161; -x_161 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__43; +x_161 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__43; x_90 = x_161; goto block_160; } @@ -7672,7 +7872,7 @@ lean_inc(x_29); x_92 = l_Lean_Syntax_node1(x_29, x_60, x_91); lean_inc(x_29); x_93 = l_Lean_Syntax_node2(x_29, x_72, x_56, x_92); -x_94 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__45; +x_94 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__45; lean_inc(x_29); x_95 = l_Lean_Syntax_node5(x_29, x_94, x_45, x_49, x_51, x_93, x_59); lean_inc(x_29); @@ -7685,10 +7885,10 @@ lean_inc(x_29); x_99 = l_Lean_Syntax_node2(x_29, x_76, x_98, x_62); lean_inc(x_29); x_100 = l_Lean_Syntax_node3(x_29, x_60, x_99, x_77, x_89); -x_101 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__6; +x_101 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__6; lean_inc(x_29); x_102 = l_Lean_Syntax_node1(x_29, x_101, x_100); -x_103 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__4; +x_103 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__4; x_104 = l_Lean_Syntax_node2(x_29, x_103, x_33, x_102); x_105 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_105, 0, x_104); @@ -7706,7 +7906,7 @@ lean_inc(x_29); x_107 = l_Lean_Syntax_node1(x_29, x_60, x_106); lean_inc(x_29); x_108 = l_Lean_Syntax_node2(x_29, x_72, x_56, x_107); -x_109 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__45; +x_109 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__45; lean_inc(x_29); x_110 = l_Lean_Syntax_node5(x_29, x_109, x_45, x_49, x_51, x_108, x_59); lean_inc(x_29); @@ -7719,10 +7919,10 @@ lean_inc(x_29); x_114 = l_Lean_Syntax_node2(x_29, x_76, x_113, x_62); lean_inc(x_29); x_115 = l_Lean_Syntax_node3(x_29, x_60, x_114, x_77, x_89); -x_116 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__6; +x_116 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__6; lean_inc(x_29); x_117 = l_Lean_Syntax_node1(x_29, x_116, x_115); -x_118 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__4; +x_118 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__4; x_119 = l_Lean_Syntax_node2(x_29, x_118, x_33, x_117); x_120 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_120, 0, x_119); @@ -7739,15 +7939,15 @@ lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; x_121 = lean_ctor_get(x_57, 0); lean_inc(x_121); lean_dec(x_57); -x_122 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__80; +x_122 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__80; x_123 = l_String_intercalate(x_122, x_121); x_124 = lean_string_append(x_40, x_123); lean_dec(x_123); x_125 = lean_box(2); x_126 = l_Lean_Syntax_mkNameLit(x_124, x_125); -x_127 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__81; +x_127 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__81; x_128 = lean_array_push(x_127, x_126); -x_129 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__79; +x_129 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__79; x_130 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_130, 0, x_125); lean_ctor_set(x_130, 1, x_129); @@ -7756,7 +7956,7 @@ lean_inc(x_29); x_131 = l_Lean_Syntax_node1(x_29, x_60, x_130); lean_inc(x_29); x_132 = l_Lean_Syntax_node2(x_29, x_72, x_56, x_131); -x_133 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__45; +x_133 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__45; lean_inc(x_29); x_134 = l_Lean_Syntax_node5(x_29, x_133, x_45, x_49, x_51, x_132, x_59); lean_inc(x_29); @@ -7769,10 +7969,10 @@ lean_inc(x_29); x_138 = l_Lean_Syntax_node2(x_29, x_76, x_137, x_62); lean_inc(x_29); x_139 = l_Lean_Syntax_node3(x_29, x_60, x_138, x_77, x_89); -x_140 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__6; +x_140 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__6; lean_inc(x_29); x_141 = l_Lean_Syntax_node1(x_29, x_140, x_139); -x_142 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__4; +x_142 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__4; x_143 = l_Lean_Syntax_node2(x_29, x_142, x_33, x_141); x_144 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_144, 0, x_143); @@ -7790,7 +7990,7 @@ lean_inc(x_29); x_146 = l_Lean_Syntax_node1(x_29, x_60, x_145); lean_inc(x_29); x_147 = l_Lean_Syntax_node2(x_29, x_72, x_56, x_146); -x_148 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__45; +x_148 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__45; lean_inc(x_29); x_149 = l_Lean_Syntax_node5(x_29, x_148, x_45, x_49, x_51, x_147, x_59); lean_inc(x_29); @@ -7803,10 +8003,10 @@ lean_inc(x_29); x_153 = l_Lean_Syntax_node2(x_29, x_76, x_152, x_62); lean_inc(x_29); x_154 = l_Lean_Syntax_node3(x_29, x_60, x_153, x_77, x_89); -x_155 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__6; +x_155 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__6; lean_inc(x_29); x_156 = l_Lean_Syntax_node1(x_29, x_155, x_154); -x_157 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__4; +x_157 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__4; x_158 = l_Lean_Syntax_node2(x_29, x_157, x_33, x_156); x_159 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_159, 0, x_158); @@ -7821,11 +8021,11 @@ return x_159; } } } -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(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; +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) @@ -7892,7 +8092,7 @@ 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_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1(x_1, x_23, x_22, x_2, x_3); +x_24 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1(x_1, x_23, x_22, x_2, x_3); lean_dec(x_1); return x_24; } @@ -7904,24 +8104,24 @@ lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_dec(x_9); x_25 = lean_box(0); x_26 = lean_box(0); -x_27 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1(x_1, x_26, x_25, x_2, x_3); +x_27 = l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1(x_1, x_26, x_25, x_2, x_3); lean_dec(x_1); return x_27; } } } } -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) { +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); +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_2); lean_dec(x_1); return x_6; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__1() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; @@ -7932,7 +8132,7 @@ x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__2() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__2() { _start: { lean_object* x_1; @@ -7940,27 +8140,27 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_patternIgnore), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__3() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__3() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__2; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____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_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__4() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__1; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____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_2047____closed__5() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__5() { _start: { lean_object* x_1; lean_object* x_2; @@ -7970,12 +8170,12 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__6() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__6() { _start: { lean_object* x_1; lean_object* x_2; uint8_t x_3; lean_object* x_4; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__5; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__5; x_3 = 0; x_4 = lean_alloc_ctor(0, 2, 1); lean_ctor_set(x_4, 0, x_1); @@ -7984,7 +8184,7 @@ 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_Extra___hyg_2047____closed__7() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__7() { _start: { lean_object* x_1; @@ -7992,17 +8192,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_patternIgnore_formatter), 6, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__8() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__7; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____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_2047____closed__9() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__9() { _start: { lean_object* x_1; @@ -8010,7 +8210,7 @@ x_1 = l_Lean_PrettyPrinter_Formatter_formatterAliasesRef; return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__10() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__10() { _start: { lean_object* x_1; @@ -8018,17 +8218,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_patternIgnore_parenthesizer), 6, return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__11() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__11() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__10; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____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_2047____closed__12() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__12() { _start: { lean_object* x_1; @@ -8036,7 +8236,7 @@ x_1 = l_Lean_PrettyPrinter_Parenthesizer_parenthesizerAliasesRef; return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__13() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; @@ -8047,7 +8247,7 @@ x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__14() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__14() { _start: { lean_object* x_1; @@ -8055,27 +8255,27 @@ 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_2047____closed__15() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__15() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__14; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__14; 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_2047____closed__16() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__16() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__13; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__13; 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_2047____closed__17() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__17() { _start: { lean_object* x_1; @@ -8083,17 +8283,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_2047____closed__18() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__18() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__17; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__17; 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_2047____closed__19() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__19() { _start: { lean_object* x_1; @@ -8101,17 +8301,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_2047____closed__20() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__20() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__19; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__19; 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_2047____closed__21() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__21() { _start: { lean_object* x_1; @@ -8119,28 +8319,28 @@ 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_2047____closed__22() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__22() { _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_2047____closed__21; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__21; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__23() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__1; x_2 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__2; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__21; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__21; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__24() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__24() { _start: { lean_object* x_1; lean_object* x_2; @@ -8150,17 +8350,17 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__25() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__25() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__23; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__23; 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_2047____closed__26() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__26() { _start: { lean_object* x_1; lean_object* x_2; @@ -8170,12 +8370,12 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__27() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__27() { _start: { lean_object* x_1; lean_object* x_2; uint8_t x_3; lean_object* x_4; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__26; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__26; x_3 = 1; x_4 = lean_alloc_ctor(0, 2, 1); lean_ctor_set(x_4, 0, x_1); @@ -8184,7 +8384,7 @@ 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_Extra___hyg_2047____closed__28() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__28() { _start: { lean_object* x_1; @@ -8192,17 +8392,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_2047____closed__29() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__29() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__28; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____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_Extra___hyg_2047____closed__30() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__30() { _start: { lean_object* x_1; @@ -8210,56 +8410,37 @@ 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_2047____closed__31() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__31() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__30; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____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_Extra___hyg_2047____closed__32() { -_start: -{ -lean_object* x_1; -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_2047____closed__33() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____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_2047____closed__32; +x_2 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__25; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__34() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__1; -x_2 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__2; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__32; -x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); -return x_4; -} -} -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__35() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__33() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__34; +x_1 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________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_Extra___hyg_2047____closed__36() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__34() { _start: { lean_object* x_1; @@ -8267,17 +8448,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_2047____closed__37() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__35() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__36; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__34; 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_2047____closed__38() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__36() { _start: { lean_object* x_1; @@ -8285,17 +8466,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_2047____closed__39() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__37() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__38; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__36; 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_2047____closed__40() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__38() { _start: { lean_object* x_1; @@ -8303,38 +8484,38 @@ 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_2047____closed__41() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____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_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__40; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__38; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__42() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__40() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__1; x_2 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__2; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__40; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__38; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__43() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__41() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__42; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__40; 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_2047____closed__44() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__42() { _start: { lean_object* x_1; @@ -8342,17 +8523,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_2047____closed__45() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__43() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__44; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____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_Extra___hyg_2047____closed__46() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__44() { _start: { lean_object* x_1; @@ -8360,17 +8541,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_2047____closed__47() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__45() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__46; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__44; 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_2047____closed__48() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__46() { _start: { lean_object* x_1; @@ -8378,28 +8559,28 @@ 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_2047____closed__49() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____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_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__48; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__46; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__50() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__48() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__1; x_2 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__2; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__48; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__46; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__51() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__49() { _start: { lean_object* x_1; @@ -8407,27 +8588,27 @@ 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_2047____closed__52() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__50() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__51; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__49; 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_2047____closed__53() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__51() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__50; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__48; 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_2047____closed__54() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__52() { _start: { lean_object* x_1; lean_object* x_2; uint8_t x_3; lean_object* x_4; @@ -8441,7 +8622,7 @@ 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_Extra___hyg_2047____closed__55() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__53() { _start: { lean_object* x_1; @@ -8449,17 +8630,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_2047____closed__56() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__54() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__55; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____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_2047____closed__57() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__55() { _start: { lean_object* x_1; @@ -8467,17 +8648,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_2047____closed__58() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__56() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__57; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____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_2047____closed__59() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__57() { _start: { lean_object* x_1; @@ -8485,28 +8666,28 @@ 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_2047____closed__60() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____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_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__59; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__57; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__61() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__59() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__1; x_2 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__2; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__59; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__57; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__62() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__60() { _start: { lean_object* x_1; @@ -8514,27 +8695,27 @@ 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_2047____closed__63() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__61() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__62; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__60; 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_2047____closed__64() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__62() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__61; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__59; 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_2047____closed__65() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__63() { _start: { lean_object* x_1; @@ -8542,17 +8723,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_2047____closed__66() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__64() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__65; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____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_2047____closed__67() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__65() { _start: { lean_object* x_1; @@ -8560,17 +8741,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_2047____closed__68() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__66() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__67; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____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_2047____closed__69() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__67() { _start: { lean_object* x_1; @@ -8578,28 +8759,28 @@ 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_2047____closed__70() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__68() { _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_2047____closed__69; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__67; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__71() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__69() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__1; x_2 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__2; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__69; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__67; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__72() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__70() { _start: { lean_object* x_1; @@ -8607,27 +8788,27 @@ 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_2047____closed__73() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__71() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__72; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__70; 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_2047____closed__74() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__72() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__71; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__69; 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_2047____closed__75() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__73() { _start: { lean_object* x_1; @@ -8635,17 +8816,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_2047____closed__76() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__74() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__75; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____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_2047____closed__77() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__75() { _start: { lean_object* x_1; @@ -8653,17 +8834,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_2047____closed__78() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__76() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__77; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____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_2047____closed__79() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__77() { _start: { lean_object* x_1; @@ -8671,28 +8852,28 @@ 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_2047____closed__80() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____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_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__79; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__77; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__81() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__79() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__1; x_2 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__2; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__79; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__77; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__82() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__80() { _start: { lean_object* x_1; @@ -8700,27 +8881,27 @@ 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_2047____closed__83() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__81() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__82; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__80; 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_2047____closed__84() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__82() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__81; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____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_Extra___hyg_2047____closed__85() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__83() { _start: { lean_object* x_1; @@ -8728,17 +8909,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_2047____closed__86() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__84() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__85; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____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_2047____closed__87() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__85() { _start: { lean_object* x_1; @@ -8746,17 +8927,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_2047____closed__88() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__86() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__87; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____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_2047____closed__89() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__87() { _start: { lean_object* x_1; @@ -8764,28 +8945,28 @@ 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_2047____closed__90() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__88() { _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_2047____closed__89; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__87; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__91() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__89() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__1; x_2 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__2; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__89; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__87; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__92() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__90() { _start: { lean_object* x_1; @@ -8793,27 +8974,27 @@ 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_2047____closed__93() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__91() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__92; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__90; 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_2047____closed__94() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__92() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__91; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__89; 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_2047____closed__95() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__93() { _start: { lean_object* x_1; @@ -8821,17 +9002,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_2047____closed__96() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__94() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__95; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__93; 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_2047____closed__97() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__95() { _start: { lean_object* x_1; @@ -8839,17 +9020,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_2047____closed__98() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__96() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__97; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__95; 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_2047____closed__99() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__97() { _start: { lean_object* x_1; @@ -8857,28 +9038,28 @@ 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_2047____closed__100() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____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_2047____closed__99; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__97; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__101() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__99() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__1; x_2 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__2; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__99; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__97; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__102() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__100() { _start: { lean_object* x_1; @@ -8886,27 +9067,27 @@ 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_2047____closed__103() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__101() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__102; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__100; 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_2047____closed__104() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__102() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__101; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____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_2047____closed__105() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__103() { _start: { lean_object* x_1; @@ -8914,17 +9095,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_2047____closed__106() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__104() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__105; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____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_2047____closed__107() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__105() { _start: { lean_object* x_1; @@ -8932,17 +9113,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_2047____closed__108() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__106() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__107; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____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_2047____closed__109() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__107() { _start: { lean_object* x_1; @@ -8950,38 +9131,38 @@ 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_2047____closed__110() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____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_2047____closed__109; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__107; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__111() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__109() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__1; x_2 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__2; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__109; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__107; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__112() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__110() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__111; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____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_2047____closed__113() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__111() { _start: { lean_object* x_1; @@ -8989,17 +9170,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_2047____closed__114() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__112() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__113; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____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_2047____closed__115() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__113() { _start: { lean_object* x_1; @@ -9007,17 +9188,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_2047____closed__116() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__114() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__115; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__113; 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_2047____closed__117() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__115() { _start: { lean_object* x_1; @@ -9025,38 +9206,38 @@ 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_2047____closed__118() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__116() { _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_2047____closed__117; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__115; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__119() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__117() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__1; x_2 = l___regBuiltin_Lean_Parser_antiquotNestedExpr_formatter___closed__2; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__117; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__115; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__120() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__118() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__119; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__117; 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_2047____closed__121() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__119() { _start: { lean_object* x_1; @@ -9064,17 +9245,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_2047____closed__122() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__120() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__121; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__119; 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_2047____closed__123() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__121() { _start: { lean_object* x_1; @@ -9082,25 +9263,25 @@ 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_2047____closed__124() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__122() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__123; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__121; 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_2047_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065_(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; x_2 = l_Lean_Parser_patternIgnore_formatter___closed__2; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__1; -x_4 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__3; -x_5 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__4; -x_6 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__6; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__1; +x_4 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__3; +x_5 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__4; +x_6 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__6; x_7 = l_Lean_Parser_registerAlias(x_2, x_3, x_4, x_5, x_6, x_1); if (lean_obj_tag(x_7) == 0) { @@ -9108,8 +9289,8 @@ 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_Extra___hyg_2047____closed__9; -x_10 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__8; +x_9 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__9; +x_10 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__8; x_11 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_2, x_10, x_8); if (lean_obj_tag(x_11) == 0) { @@ -9117,8 +9298,8 @@ 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_Extra___hyg_2047____closed__12; -x_14 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__11; +x_13 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__12; +x_14 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__11; x_15 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_2, x_14, x_12); if (lean_obj_tag(x_15) == 0) { @@ -9127,9 +9308,9 @@ x_16 = lean_ctor_get(x_15, 1); lean_inc(x_16); lean_dec(x_15); x_17 = l_Lean_Parser_group_formatter___closed__2; -x_18 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__13; -x_19 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__15; -x_20 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__16; +x_18 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__13; +x_19 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__15; +x_20 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__16; x_21 = l_Lean_Parser_registerAlias(x_17, x_18, x_19, x_20, x_6, x_16); if (lean_obj_tag(x_21) == 0) { @@ -9137,7 +9318,7 @@ 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_2047____closed__18; +x_23 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__18; x_24 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_17, x_23, x_22); if (lean_obj_tag(x_24) == 0) { @@ -9145,7 +9326,7 @@ lean_object* x_25; lean_object* x_26; lean_object* x_27; 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_2047____closed__20; +x_26 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__20; x_27 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_17, x_26, x_25); if (lean_obj_tag(x_27) == 0) { @@ -9153,11 +9334,11 @@ lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean x_28 = lean_ctor_get(x_27, 1); lean_inc(x_28); lean_dec(x_27); -x_29 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__22; -x_30 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__23; -x_31 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__24; -x_32 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__25; -x_33 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__27; +x_29 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__22; +x_30 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__23; +x_31 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__24; +x_32 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__25; +x_33 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__27; x_34 = l_Lean_Parser_registerAlias(x_29, x_30, x_31, x_32, x_33, x_28); if (lean_obj_tag(x_34) == 0) { @@ -9165,7 +9346,7 @@ 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_Extra___hyg_2047____closed__29; +x_36 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__29; x_37 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_29, x_36, x_35); if (lean_obj_tag(x_37) == 0) { @@ -9173,7 +9354,7 @@ lean_object* x_38; lean_object* x_39; lean_object* x_40; 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_Extra___hyg_2047____closed__31; +x_39 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__31; x_40 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_29, x_39, x_38); if (lean_obj_tag(x_40) == 0) { @@ -9181,9 +9362,9 @@ lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean 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_2047____closed__33; -x_43 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__34; -x_44 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__35; +x_42 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__32; +x_43 = l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__26; +x_44 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__33; x_45 = l_Lean_Parser_registerAlias(x_42, x_43, x_31, x_44, x_33, x_41); if (lean_obj_tag(x_45) == 0) { @@ -9191,7 +9372,7 @@ lean_object* x_46; lean_object* x_47; lean_object* x_48; x_46 = lean_ctor_get(x_45, 1); lean_inc(x_46); lean_dec(x_45); -x_47 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__37; +x_47 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__35; x_48 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_42, x_47, x_46); if (lean_obj_tag(x_48) == 0) { @@ -9199,7 +9380,7 @@ lean_object* x_49; lean_object* x_50; lean_object* x_51; x_49 = lean_ctor_get(x_48, 1); lean_inc(x_49); lean_dec(x_48); -x_50 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__39; +x_50 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__37; x_51 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_42, x_50, x_49); if (lean_obj_tag(x_51) == 0) { @@ -9207,9 +9388,9 @@ lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean 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_2047____closed__41; -x_54 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__42; -x_55 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__43; +x_53 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__39; +x_54 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__40; +x_55 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__41; x_56 = l_Lean_Parser_registerAlias(x_53, x_54, x_31, x_55, x_33, x_52); if (lean_obj_tag(x_56) == 0) { @@ -9217,7 +9398,7 @@ lean_object* x_57; lean_object* x_58; lean_object* x_59; x_57 = lean_ctor_get(x_56, 1); lean_inc(x_57); lean_dec(x_56); -x_58 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__45; +x_58 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__43; x_59 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_53, x_58, x_57); if (lean_obj_tag(x_59) == 0) { @@ -9225,7 +9406,7 @@ lean_object* x_60; lean_object* x_61; lean_object* x_62; x_60 = lean_ctor_get(x_59, 1); lean_inc(x_60); lean_dec(x_59); -x_61 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__47; +x_61 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__45; x_62 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_53, x_61, x_60); if (lean_obj_tag(x_62) == 0) { @@ -9233,11 +9414,11 @@ lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean 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_2047____closed__49; -x_65 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__50; -x_66 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__52; -x_67 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__53; -x_68 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__54; +x_64 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__47; +x_65 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__48; +x_66 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__50; +x_67 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__51; +x_68 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__52; x_69 = l_Lean_Parser_registerAlias(x_64, x_65, x_66, x_67, x_68, x_63); if (lean_obj_tag(x_69) == 0) { @@ -9245,7 +9426,7 @@ lean_object* x_70; lean_object* x_71; lean_object* x_72; 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_Extra___hyg_2047____closed__56; +x_71 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__54; x_72 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_64, x_71, x_70); if (lean_obj_tag(x_72) == 0) { @@ -9253,7 +9434,7 @@ lean_object* x_73; lean_object* x_74; lean_object* x_75; x_73 = lean_ctor_get(x_72, 1); lean_inc(x_73); lean_dec(x_72); -x_74 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__58; +x_74 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__56; x_75 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_64, x_74, x_73); if (lean_obj_tag(x_75) == 0) { @@ -9261,10 +9442,10 @@ lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean x_76 = lean_ctor_get(x_75, 1); lean_inc(x_76); lean_dec(x_75); -x_77 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__60; -x_78 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__61; -x_79 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__63; -x_80 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__64; +x_77 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__58; +x_78 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__59; +x_79 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__61; +x_80 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__62; x_81 = l_Lean_Parser_registerAlias(x_77, x_78, x_79, x_80, x_68, x_76); if (lean_obj_tag(x_81) == 0) { @@ -9272,7 +9453,7 @@ lean_object* x_82; lean_object* x_83; lean_object* x_84; x_82 = lean_ctor_get(x_81, 1); lean_inc(x_82); lean_dec(x_81); -x_83 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__66; +x_83 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__64; x_84 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_77, x_83, x_82); if (lean_obj_tag(x_84) == 0) { @@ -9280,7 +9461,7 @@ 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_2047____closed__68; +x_86 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__66; x_87 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_77, x_86, x_85); if (lean_obj_tag(x_87) == 0) { @@ -9288,10 +9469,10 @@ lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean 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_2047____closed__70; -x_90 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__71; -x_91 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__73; -x_92 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__74; +x_89 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__68; +x_90 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__69; +x_91 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__71; +x_92 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__72; x_93 = l_Lean_Parser_registerAlias(x_89, x_90, x_91, x_92, x_68, x_88); if (lean_obj_tag(x_93) == 0) { @@ -9299,7 +9480,7 @@ 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_2047____closed__76; +x_95 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__74; x_96 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_89, x_95, x_94); if (lean_obj_tag(x_96) == 0) { @@ -9307,7 +9488,7 @@ 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_2047____closed__78; +x_98 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__76; x_99 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_89, x_98, x_97); if (lean_obj_tag(x_99) == 0) { @@ -9315,10 +9496,10 @@ 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_2047____closed__80; -x_102 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__81; -x_103 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__83; -x_104 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__84; +x_101 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__78; +x_102 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__79; +x_103 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__81; +x_104 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__82; x_105 = l_Lean_Parser_registerAlias(x_101, x_102, x_103, x_104, x_68, x_100); if (lean_obj_tag(x_105) == 0) { @@ -9326,7 +9507,7 @@ 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_2047____closed__86; +x_107 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__84; x_108 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_101, x_107, x_106); if (lean_obj_tag(x_108) == 0) { @@ -9334,7 +9515,7 @@ 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_2047____closed__88; +x_110 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__86; x_111 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_101, x_110, x_109); if (lean_obj_tag(x_111) == 0) { @@ -9342,10 +9523,10 @@ lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; 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_2047____closed__90; -x_114 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__91; -x_115 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__93; -x_116 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__94; +x_113 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__88; +x_114 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__89; +x_115 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__91; +x_116 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__92; x_117 = l_Lean_Parser_registerAlias(x_113, x_114, x_115, x_116, x_68, x_112); if (lean_obj_tag(x_117) == 0) { @@ -9353,7 +9534,7 @@ 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_2047____closed__96; +x_119 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__94; x_120 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_113, x_119, x_118); if (lean_obj_tag(x_120) == 0) { @@ -9361,7 +9542,7 @@ 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_Extra___hyg_2047____closed__98; +x_122 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__96; x_123 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_113, x_122, x_121); if (lean_obj_tag(x_123) == 0) { @@ -9369,10 +9550,10 @@ lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; 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_Extra___hyg_2047____closed__100; -x_126 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__101; -x_127 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__103; -x_128 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__104; +x_125 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__98; +x_126 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__99; +x_127 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__101; +x_128 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__102; x_129 = l_Lean_Parser_registerAlias(x_125, x_126, x_127, x_128, x_68, x_124); if (lean_obj_tag(x_129) == 0) { @@ -9380,7 +9561,7 @@ 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_2047____closed__106; +x_131 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__104; x_132 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_125, x_131, x_130); if (lean_obj_tag(x_132) == 0) { @@ -9388,7 +9569,7 @@ 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_Extra___hyg_2047____closed__108; +x_134 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__106; x_135 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_125, x_134, x_133); if (lean_obj_tag(x_135) == 0) { @@ -9396,9 +9577,9 @@ lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; 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_Extra___hyg_2047____closed__110; -x_138 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__111; -x_139 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__112; +x_137 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__108; +x_138 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__109; +x_139 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__110; x_140 = l_Lean_Parser_registerAlias(x_137, x_138, x_31, x_139, x_33, x_136); if (lean_obj_tag(x_140) == 0) { @@ -9406,7 +9587,7 @@ lean_object* x_141; lean_object* x_142; lean_object* x_143; x_141 = lean_ctor_get(x_140, 1); lean_inc(x_141); lean_dec(x_140); -x_142 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__114; +x_142 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__112; x_143 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_137, x_142, x_141); if (lean_obj_tag(x_143) == 0) { @@ -9414,7 +9595,7 @@ 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_Extra___hyg_2047____closed__116; +x_145 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__114; x_146 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_137, x_145, x_144); if (lean_obj_tag(x_146) == 0) { @@ -9422,9 +9603,9 @@ lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; 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_Extra___hyg_2047____closed__118; -x_149 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__119; -x_150 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__120; +x_148 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__116; +x_149 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__117; +x_150 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__118; x_151 = l_Lean_Parser_registerAlias(x_148, x_149, x_31, x_150, x_33, x_147); if (lean_obj_tag(x_151) == 0) { @@ -9432,7 +9613,7 @@ lean_object* x_152; lean_object* x_153; lean_object* x_154; x_152 = lean_ctor_get(x_151, 1); lean_inc(x_152); lean_dec(x_151); -x_153 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__122; +x_153 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__120; x_154 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_148, x_153, x_152); if (lean_obj_tag(x_154) == 0) { @@ -9440,7 +9621,7 @@ 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_Extra___hyg_2047____closed__124; +x_156 = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__122; x_157 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_148, x_156, x_155); return x_157; } @@ -10550,6 +10731,24 @@ l_Lean_Parser_rawIdent___closed__1 = _init_l_Lean_Parser_rawIdent___closed__1(); lean_mark_persistent(l_Lean_Parser_rawIdent___closed__1); l_Lean_Parser_rawIdent = _init_l_Lean_Parser_rawIdent(); lean_mark_persistent(l_Lean_Parser_rawIdent); +l_Lean_Parser_hygieneInfo_formatter___closed__1 = _init_l_Lean_Parser_hygieneInfo_formatter___closed__1(); +lean_mark_persistent(l_Lean_Parser_hygieneInfo_formatter___closed__1); +l_Lean_Parser_hygieneInfo_formatter___closed__2 = _init_l_Lean_Parser_hygieneInfo_formatter___closed__2(); +lean_mark_persistent(l_Lean_Parser_hygieneInfo_formatter___closed__2); +l_Lean_Parser_hygieneInfo_formatter___closed__3 = _init_l_Lean_Parser_hygieneInfo_formatter___closed__3(); +lean_mark_persistent(l_Lean_Parser_hygieneInfo_formatter___closed__3); +l_Lean_Parser_hygieneInfo_formatter___closed__4 = _init_l_Lean_Parser_hygieneInfo_formatter___closed__4(); +lean_mark_persistent(l_Lean_Parser_hygieneInfo_formatter___closed__4); +l_Lean_Parser_hygieneInfo_parenthesizer___closed__1 = _init_l_Lean_Parser_hygieneInfo_parenthesizer___closed__1(); +lean_mark_persistent(l_Lean_Parser_hygieneInfo_parenthesizer___closed__1); +l_Lean_Parser_hygieneInfo_parenthesizer___closed__2 = _init_l_Lean_Parser_hygieneInfo_parenthesizer___closed__2(); +lean_mark_persistent(l_Lean_Parser_hygieneInfo_parenthesizer___closed__2); +l_Lean_Parser_hygieneInfo___closed__1 = _init_l_Lean_Parser_hygieneInfo___closed__1(); +lean_mark_persistent(l_Lean_Parser_hygieneInfo___closed__1); +l_Lean_Parser_hygieneInfo___closed__2 = _init_l_Lean_Parser_hygieneInfo___closed__2(); +lean_mark_persistent(l_Lean_Parser_hygieneInfo___closed__2); +l_Lean_Parser_hygieneInfo = _init_l_Lean_Parser_hygieneInfo(); +lean_mark_persistent(l_Lean_Parser_hygieneInfo); l_Lean_Parser_numLit_formatter___closed__1 = _init_l_Lean_Parser_numLit_formatter___closed__1(); lean_mark_persistent(l_Lean_Parser_numLit_formatter___closed__1); l_Lean_Parser_numLit_formatter___closed__2 = _init_l_Lean_Parser_numLit_formatter___closed__2(); @@ -10702,489 +10901,497 @@ 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___________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__1___closed__13 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__13(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__13); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__14 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__14(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__14); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__15 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__15(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__15); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__16 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__16(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__16); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__17 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__17(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__17); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__18 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__18(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__18); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__19 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__19(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__19); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__20 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__20(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__20); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__21 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__21(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__21); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__22 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__22(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__22); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__23 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__23(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__23); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__24 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__24(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__24); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__25 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__25(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__25); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__26 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__26(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__26); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__27 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__27(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__27); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__28 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__28(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__28); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__29 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__29(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__29); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__30 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__30(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__30); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__31 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__31(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__31); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__32 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__32(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__32); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__33 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__33(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__33); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__34 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__34(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__34); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__35 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__35(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__35); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__36 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__36(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__36); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__37 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__37(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__37); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__38 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__38(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__38); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__39 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__39(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__39); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__40 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__40(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__40); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__41 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__41(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__41); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__42 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__42(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__42); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__43 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__43(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__43); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__44 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__44(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__44); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__45 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__45(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__45); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__46 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__46(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__46); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__47 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__47(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__47); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__48 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__48(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__48); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__49 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__49(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__49); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__50 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__50(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__50); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__51 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__51(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__51); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__52 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__52(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__52); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__53 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__53(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__53); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__54 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__54(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__54); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__55 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__55(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__55); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__56 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__56(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__56); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__57 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__57(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__57); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__58 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__58(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__58); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__59 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__59(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__59); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__60 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__60(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__60); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__61 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__61(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__61); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__62 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__62(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__62); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__63 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__63(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__63); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__64 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__64(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__64); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__65 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__65(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__65); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__66 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__66(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__66); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__67 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__67(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__67); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__68 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__68(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__68); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__69 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__69(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__69); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__70 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__70(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__70); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__71 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__71(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__71); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__72 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__72(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__72); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__73 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__73(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__73); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__74 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__74(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__74); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__75 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__75(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__75); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__76 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__76(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__76); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__77 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__77(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__77); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__78 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__78(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__78); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__79 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__79(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__79); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__80 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__80(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__80); -l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__81 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__81(); -lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29__________1___lambda__1___closed__81); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__1(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__1); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__2 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__2(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__2); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__3 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__3(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__3); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__4 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__4(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__4); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__5 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__5(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__5); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__6 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__6(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__6); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__7 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__7(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__7); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__8 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__8(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__8); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__9 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__9(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__9); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__10 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__10(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__10); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__11 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__11(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__11); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__12 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__12(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__12); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__13 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__13(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__13); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__14 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__14(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__14); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__15 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__15(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__15); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__16 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__16(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__16); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__17 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__17(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__17); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__18 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__18(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__18); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__19 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__19(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__19); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__20 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__20(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__20); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__21 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__21(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__21); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__22 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__22(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__22); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__23 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__23(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__23); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__24 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__24(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__24); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__25 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__25(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__25); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__26 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__26(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__26); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__27 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__27(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__27); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__28 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__28(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__28); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__29 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__29(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__29); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__30 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__30(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__30); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__31 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__31(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__31); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__32 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__32(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__32); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__33 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__33(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__33); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__34 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__34(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__34); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__35 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__35(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__35); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__36 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__36(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__36); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__37 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__37(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__37); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__38 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__38(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__38); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__39 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__39(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__39); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__40 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__40(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__40); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__41 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__41(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__41); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__42 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__42(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__42); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__43 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__43(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__43); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__44 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__44(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__44); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__45 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__45(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__45); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__46 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__46(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__46); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__47 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__47(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__47); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__48 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__48(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__48); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__49 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__49(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__49); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__50 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__50(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__50); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__51 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__51(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__51); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__52 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__52(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__52); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__53 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__53(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__53); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__54 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__54(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__54); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__55 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__55(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__55); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__56 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__56(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__56); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__57 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__57(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__57); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__58 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__58(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__58); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__59 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__59(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__59); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__60 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__60(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__60); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__61 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__61(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__61); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__62 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__62(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__62); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__63 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__63(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__63); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__64 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__64(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__64); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__65 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__65(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__65); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__66 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__66(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__66); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__67 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__67(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__67); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__68 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__68(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__68); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__69 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__69(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__69); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__70 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__70(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__70); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__71 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__71(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__71); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__72 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__72(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__72); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__73 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__73(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__73); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__74 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__74(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__74); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__75 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__75(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__75); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__76 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__76(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__76); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__77 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__77(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__77); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__78 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__78(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__78); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__79 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__79(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__79); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__80 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__80(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__80); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__81 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__81(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__81); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__82 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__82(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__82); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__83 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__83(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__83); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__84 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__84(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__84); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__85 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__85(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__85); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__86 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__86(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__86); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__87 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__87(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__87); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__88 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__88(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__88); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__89 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__89(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__89); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__90 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__90(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__90); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__91 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__91(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__91); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__92 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__92(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__92); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__93 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__93(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__93); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__94 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__94(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__94); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__95 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__95(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__95); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__96 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__96(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__96); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__97 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__97(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__97); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__98 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__98(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__98); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__99 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__99(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__99); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__100 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__100(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__100); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__101 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__101(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__101); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__102 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__102(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__102); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__103 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__103(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__103); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__104 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__104(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__104); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__105 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__105(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__105); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__106 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__106(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__106); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__107 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__107(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__107); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__108 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__108(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__108); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__109 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__109(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__109); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__110 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__110(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__110); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__111 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__111(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__111); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__112 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__112(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__112); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__113 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__113(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__113); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__114 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__114(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__114); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__115 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__115(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__115); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__116 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__116(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__116); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__117 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__117(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__117); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__118 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__118(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__118); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__119 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__119(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__119); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__120 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__120(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__120); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__121 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__121(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__121); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__122 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__122(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__122); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__123 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__123(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__123); -l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__124 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__124(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047____closed__124); -res = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2047_(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_______________closed__36 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__36(); +lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__36); +l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__37 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__37(); +lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__37); +l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__38 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__38(); +lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__38); +l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__39 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__39(); +lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__39); +l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__40 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__40(); +lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__40); +l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__41 = _init_l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__41(); +lean_mark_persistent(l_Lean_Parser_termRegister__parser__alias_x28Kind_x3a_x3d___x29_______________closed__41); +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__1___closed__13 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__13(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__13); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__14 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__14(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__14); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__15 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__15(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__15); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__16 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__16(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__16); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__17 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__17(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__17); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__18 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__18(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__18); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__19 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__19(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__19); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__20 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__20(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__20); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__21 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__21(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__21); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__22 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__22(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__22); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__23 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__23(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__23); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__24 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__24(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__24); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__25 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__25(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__25); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__26 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__26(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__26); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__27 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__27(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__27); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__28 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__28(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__28); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__29 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__29(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__29); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__30 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__30(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__30); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__31 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__31(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__31); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__32 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__32(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__32); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__33 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__33(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__33); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__34 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__34(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__34); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__35 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__35(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__35); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__36 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__36(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__36); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__37 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__37(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__37); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__38 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__38(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__38); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__39 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__39(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__39); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__40 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__40(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__40); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__41 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__41(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__41); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__42 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__42(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__42); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__43 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__43(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__43); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__44 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__44(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__44); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__45 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__45(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__45); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__46 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__46(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__46); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__47 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__47(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__47); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__48 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__48(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__48); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__49 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__49(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__49); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__50 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__50(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__50); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__51 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__51(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__51); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__52 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__52(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__52); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__53 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__53(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__53); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__54 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__54(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__54); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__55 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__55(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__55); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__56 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__56(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__56); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__57 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__57(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__57); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__58 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__58(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__58); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__59 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__59(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__59); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__60 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__60(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__60); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__61 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__61(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__61); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__62 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__62(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__62); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__63 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__63(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__63); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__64 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__64(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__64); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__65 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__65(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__65); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__66 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__66(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__66); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__67 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__67(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__67); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__68 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__68(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__68); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__69 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__69(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__69); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__70 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__70(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__70); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__71 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__71(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__71); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__72 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__72(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__72); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__73 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__73(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__73); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__74 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__74(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__74); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__75 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__75(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__75); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__76 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__76(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__76); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__77 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__77(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__77); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__78 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__78(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__78); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__79 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__79(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__79); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__80 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__80(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__80); +l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__81 = _init_l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__81(); +lean_mark_persistent(l_Lean_Parser___aux__Lean__Parser__Extra______macroRules__Lean__Parser__termRegister__parser__alias_x28Kind_x3a_x3d___x29______________1___lambda__1___closed__81); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__1(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__1); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__2 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__2(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__2); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__3 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__3(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__3); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__4 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__4(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__4); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__5 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__5(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__5); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__6 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__6(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__6); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__7 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__7(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__7); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__8 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__8(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__8); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__9 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__9(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__9); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__10 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__10(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__10); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__11 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__11(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__11); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__12 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__12(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__12); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__13 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__13(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__13); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__14 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__14(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__14); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__15 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__15(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__15); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__16 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__16(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__16); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__17 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__17(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__17); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__18 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__18(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__18); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__19 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__19(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__19); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__20 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__20(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__20); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__21 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__21(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__21); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__22 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__22(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__22); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__23 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__23(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__23); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__24 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__24(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__24); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__25 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__25(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__25); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__26 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__26(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__26); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__27 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__27(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__27); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__28 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__28(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__28); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__29 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__29(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__29); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__30 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__30(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__30); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__31 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__31(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__31); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__32 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__32(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__32); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__33 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__33(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__33); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__34 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__34(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__34); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__35 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__35(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__35); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__36 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__36(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__36); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__37 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__37(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__37); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__38 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__38(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__38); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__39 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__39(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__39); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__40 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__40(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__40); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__41 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__41(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__41); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__42 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__42(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__42); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__43 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__43(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__43); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__44 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__44(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__44); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__45 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__45(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__45); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__46 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__46(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__46); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__47 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__47(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__47); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__48 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__48(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__48); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__49 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__49(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__49); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__50 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__50(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__50); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__51 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__51(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__51); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__52 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__52(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__52); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__53 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__53(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__53); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__54 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__54(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__54); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__55 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__55(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__55); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__56 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__56(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__56); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__57 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__57(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__57); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__58 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__58(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__58); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__59 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__59(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__59); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__60 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__60(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__60); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__61 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__61(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__61); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__62 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__62(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__62); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__63 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__63(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__63); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__64 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__64(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__64); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__65 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__65(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__65); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__66 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__66(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__66); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__67 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__67(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__67); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__68 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__68(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__68); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__69 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__69(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__69); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__70 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__70(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__70); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__71 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__71(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__71); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__72 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__72(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__72); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__73 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__73(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__73); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__74 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__74(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__74); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__75 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__75(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__75); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__76 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__76(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__76); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__77 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__77(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__77); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__78 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__78(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__78); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__79 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__79(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__79); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__80 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__80(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__80); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__81 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__81(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__81); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__82 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__82(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__82); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__83 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__83(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__83); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__84 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__84(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__84); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__85 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__85(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__85); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__86 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__86(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__86); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__87 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__87(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__87); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__88 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__88(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__88); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__89 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__89(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__89); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__90 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__90(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__90); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__91 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__91(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__91); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__92 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__92(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__92); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__93 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__93(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__93); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__94 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__94(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__94); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__95 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__95(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__95); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__96 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__96(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__96); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__97 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__97(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__97); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__98 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__98(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__98); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__99 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__99(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__99); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__100 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__100(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__100); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__101 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__101(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__101); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__102 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__102(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__102); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__103 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__103(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__103); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__104 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__104(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__104); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__105 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__105(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__105); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__106 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__106(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__106); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__107 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__107(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__107); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__108 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__108(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__108); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__109 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__109(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__109); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__110 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__110(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__110); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__111 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__111(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__111); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__112 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__112(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__112); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__113 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__113(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__113); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__114 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__114(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__114); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__115 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__115(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__115); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__116 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__116(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__116); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__117 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__117(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__117); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__118 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__118(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__118); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__119 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__119(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__119); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__120 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__120(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__120); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__121 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__121(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__121); +l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__122 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__122(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065____closed__122); +res = l_Lean_Parser_initFn____x40_Lean_Parser_Extra___hyg_2065_(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/Syntax.c b/stage0/stdlib/Lean/Parser/Syntax.c index f6386f59017f..7d9c97034910 100644 --- a/stage0/stdlib/Lean/Parser/Syntax.c +++ b/stage0/stdlib/Lean/Parser/Syntax.c @@ -279,6 +279,7 @@ static lean_object* l_Lean_Parser_Syntax_sepBy_parenthesizer___closed__14; static lean_object* l_Lean_Parser_Command_mixfix_formatter___closed__7; static lean_object* l_Lean_Parser_Command_elab_formatter___closed__12; static lean_object* l___regBuiltin_Lean_Parser_Command_mixfix_parenthesizer___closed__1; +static lean_object* l_Lean_Parser_Command_elab_formatter___closed__15; static lean_object* l___regBuiltin_Lean_Parser_Syntax_unary_declRange___closed__3; static lean_object* l_Lean_Parser_Command_mixfix___closed__5; static lean_object* l_Lean_Parser_Syntax_binary___closed__8; @@ -305,6 +306,7 @@ static lean_object* l_Lean_Parser_Syntax_cat_parenthesizer___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Command_mixfix_declRange___closed__5; static lean_object* l_Lean_Parser_Command_mixfix_formatter___closed__5; static lean_object* l_Lean_Parser_Command_catBehaviorBoth_parenthesizer___closed__2; +static lean_object* l_Lean_Parser_Command_syntax_parenthesizer___closed__16; static lean_object* l_Lean_Parser_Command_optKind_formatter___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_Command_elab_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Command_elab__rules_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -543,6 +545,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Syntax_sepBy_declRange(lean_ static lean_object* l___regBuiltin_Lean_Parser_Command_catBehaviorBoth_formatter___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Syntax_atom; static lean_object* l___regBuiltin_Lean_Parser_Syntax_nonReserved_declRange___closed__2; +static lean_object* l_Lean_Parser_Command_syntax___closed__22; static lean_object* l_Lean_Parser_Command_syntax_formatter___closed__11; 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_Command_macroArg_formatter___closed__6; @@ -612,6 +615,8 @@ lean_object* l_Lean_PrettyPrinter_Parenthesizer_atomic_parenthesizer(lean_object static lean_object* l_Lean_Parser_Command_mixfixKind___closed__2; static lean_object* l_Lean_Parser_Syntax_sepBy___closed__8; static lean_object* l___regBuiltin_Lean_Parser_Command_mixfix___closed__2; +static lean_object* l_Lean_Parser_Command_namedName___closed__18; +static lean_object* l_Lean_Parser_Command_elab_parenthesizer___closed__15; static lean_object* l_Lean_Parser_Command_syntax___closed__21; LEAN_EXPORT lean_object* l_Lean_Parser_Command_mixfixKind_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_optKind___closed__8; @@ -705,9 +710,11 @@ static lean_object* l_Lean_Parser_Syntax_sepBy_parenthesizer___closed__2; lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Syntax___hyg_63____closed__3; static lean_object* l___regBuiltin_Lean_Parser_Syntax_nonReserved_declRange___closed__1; +static lean_object* l_Lean_Parser_Command_macro_formatter___closed__14; static lean_object* l_Lean_Parser_Command_infix_formatter___closed__2; static lean_object* l_Lean_Parser_Syntax_sepBy_parenthesizer___closed__9; static lean_object* l___regBuiltin_Lean_Parser_Syntax_paren_declRange___closed__6; +static lean_object* l_Lean_Parser_Command_namedName_formatter___closed__10; static lean_object* l_Lean_Parser_Command_notationItem___closed__4; static lean_object* l_Lean_Parser_Syntax_binary_formatter___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_Command_elab; @@ -915,6 +922,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Command_macroArg_parenthesizer___ LEAN_EXPORT lean_object* l_Lean_Parser_Command_optNamedName; static lean_object* l_Lean_Parser_Command_elab_parenthesizer___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Command_macro_parenthesizer___closed__1; +static lean_object* l_Lean_Parser_Command_syntax_formatter___closed__14; static lean_object* l___regBuiltin_Lean_Parser_Command_syntaxAbbrev_formatter___closed__1; static lean_object* l_Lean_Parser_Syntax_atom_parenthesizer___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Syntax_cat_declRange___closed__1; @@ -949,6 +957,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_infixr_formatter(lean_object*, le static lean_object* l_Lean_Parser_Command_syntaxAbbrev___closed__8; static lean_object* l_Lean_Parser_Command_macro_parenthesizer___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Command_macro_declRange___closed__5; +static lean_object* l_Lean_Parser_Command_syntax_formatter___closed__15; static lean_object* l___regBuiltin_Lean_Parser_Command_macroTail_formatter___closed__1; static lean_object* l_Lean_Parser_Command_mixfix___closed__13; static lean_object* l_Lean_Parser_Command_macroArg___closed__13; @@ -1238,9 +1247,11 @@ static lean_object* l_Lean_Parser_Command_macro__rules___closed__14; LEAN_EXPORT lean_object* l_Lean_Parser_Command_macroRhs_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Command_prefix___closed__7; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Syntax___hyg_5____closed__1; +static lean_object* l_Lean_Parser_Command_namedName___closed__17; static lean_object* l_Lean_Parser_Command_notation_parenthesizer___closed__12; static lean_object* l_Lean_Parser_Command_macro__rules_formatter___closed__2; static lean_object* l_Lean_Parser_Command_mixfix_parenthesizer___closed__9; +static lean_object* l_Lean_Parser_Command_macro_parenthesizer___closed__14; static lean_object* l___regBuiltin_Lean_Parser_Command_notation_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_mixfix_formatter___closed__15; static lean_object* l_Lean_Parser_Command_infixl_parenthesizer___closed__1; @@ -1285,6 +1296,7 @@ static lean_object* l_Lean_Parser_Command_prefix_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Command_macro_formatter___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Command_syntaxAbbrev_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Syntax_paren_parenthesizer___closed__1; +static lean_object* l_Lean_Parser_Command_namedName_parenthesizer___closed__10; static lean_object* l_Lean_Parser_Command_infix_parenthesizer___closed__3; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Command_syntaxAbbrev_declRange(lean_object*); static lean_object* l_Lean_Parser_Command_elab_formatter___closed__4; @@ -6584,40 +6596,57 @@ static lean_object* _init_l_Lean_Parser_Command_namedName___closed__5() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("name", 4); +x_1 = lean_mk_string_from_bytes(" (", 2); return x_1; } } static lean_object* _init_l_Lean_Parser_Command_namedName___closed__6() { _start: { -lean_object* x_1; uint8_t x_2; lean_object* x_3; +lean_object* x_1; lean_object* x_2; x_1 = l_Lean_Parser_Command_namedName___closed__5; +x_2 = l_Lean_Parser_symbol(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Command_namedName___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_Parser_Command_namedName___closed__8() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Command_namedName___closed__7; x_2 = 0; x_3 = l_Lean_Parser_nonReservedSymbol(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_namedName___closed__7() { +static lean_object* _init_l_Lean_Parser_Command_namedName___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Syntax_paren___closed__5; -x_2 = l_Lean_Parser_Command_namedName___closed__6; +x_1 = l_Lean_Parser_Command_namedName___closed__6; +x_2 = l_Lean_Parser_Command_namedName___closed__8; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_namedName___closed__8() { +static lean_object* _init_l_Lean_Parser_Command_namedName___closed__10() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_namedName___closed__7; +x_1 = l_Lean_Parser_Command_namedName___closed__9; x_2 = l_Lean_Parser_atomic(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_namedName___closed__9() { +static lean_object* _init_l_Lean_Parser_Command_namedName___closed__11() { _start: { lean_object* x_1; @@ -6625,16 +6654,16 @@ x_1 = lean_mk_string_from_bytes(" := ", 4); return x_1; } } -static lean_object* _init_l_Lean_Parser_Command_namedName___closed__10() { +static lean_object* _init_l_Lean_Parser_Command_namedName___closed__12() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_namedName___closed__9; +x_1 = l_Lean_Parser_Command_namedName___closed__11; x_2 = l_Lean_Parser_symbol(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_namedName___closed__11() { +static lean_object* _init_l_Lean_Parser_Command_namedName___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -6644,53 +6673,53 @@ x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_namedName___closed__12() { +static lean_object* _init_l_Lean_Parser_Command_namedName___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_namedName___closed__10; -x_2 = l_Lean_Parser_Command_namedName___closed__11; +x_1 = l_Lean_Parser_Command_namedName___closed__12; +x_2 = l_Lean_Parser_Command_namedName___closed__13; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_namedName___closed__13() { +static lean_object* _init_l_Lean_Parser_Command_namedName___closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_namedName___closed__8; -x_2 = l_Lean_Parser_Command_namedName___closed__12; +x_1 = l_Lean_Parser_Command_namedName___closed__10; +x_2 = l_Lean_Parser_Command_namedName___closed__14; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_namedName___closed__14() { +static lean_object* _init_l_Lean_Parser_Command_namedName___closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_namedName___closed__3; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_namedName___closed__13; +x_3 = l_Lean_Parser_Command_namedName___closed__15; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Command_namedName___closed__15() { +static lean_object* _init_l_Lean_Parser_Command_namedName___closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_namedName___closed__4; -x_2 = l_Lean_Parser_Command_namedName___closed__14; +x_2 = l_Lean_Parser_Command_namedName___closed__16; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_namedName___closed__16() { +static lean_object* _init_l_Lean_Parser_Command_namedName___closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_namedName___closed__3; -x_2 = l_Lean_Parser_Command_namedName___closed__15; +x_2 = l_Lean_Parser_Command_namedName___closed__17; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -6699,7 +6728,7 @@ static lean_object* _init_l_Lean_Parser_Command_namedName() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Command_namedName___closed__16; +x_1 = l_Lean_Parser_Command_namedName___closed__18; return x_1; } } @@ -7991,8 +8020,18 @@ return x_7; static lean_object* _init_l_Lean_Parser_Command_namedName_formatter___closed__2() { _start: { -lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; +lean_object* x_1; lean_object* x_2; x_1 = l_Lean_Parser_Command_namedName___closed__5; +x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_formatter), 6, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Command_namedName_formatter___closed__3() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Command_namedName___closed__7; x_2 = 0; x_3 = lean_box(x_2); x_4 = lean_alloc_closure((void*)(l_Lean_Parser_nonReservedSymbol_formatter___boxed), 7, 2); @@ -8001,39 +8040,39 @@ lean_closure_set(x_4, 1, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Command_namedName_formatter___closed__3() { +static lean_object* _init_l_Lean_Parser_Command_namedName_formatter___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Syntax_paren_formatter___closed__2; -x_2 = l_Lean_Parser_Command_namedName_formatter___closed__2; +x_1 = l_Lean_Parser_Command_namedName_formatter___closed__2; +x_2 = l_Lean_Parser_Command_namedName_formatter___closed__3; 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; } } -static lean_object* _init_l_Lean_Parser_Command_namedName_formatter___closed__4() { +static lean_object* _init_l_Lean_Parser_Command_namedName_formatter___closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_namedName_formatter___closed__3; +x_1 = l_Lean_Parser_Command_namedName_formatter___closed__4; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_atomic_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_namedName_formatter___closed__5() { +static lean_object* _init_l_Lean_Parser_Command_namedName_formatter___closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_namedName___closed__9; +x_1 = l_Lean_Parser_Command_namedName___closed__11; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_namedName_formatter___closed__6() { +static lean_object* _init_l_Lean_Parser_Command_namedName_formatter___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -8045,37 +8084,37 @@ lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_namedName_formatter___closed__7() { +static lean_object* _init_l_Lean_Parser_Command_namedName_formatter___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_namedName_formatter___closed__5; -x_2 = l_Lean_Parser_Command_namedName_formatter___closed__6; +x_1 = l_Lean_Parser_Command_namedName_formatter___closed__6; +x_2 = l_Lean_Parser_Command_namedName_formatter___closed__7; 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; } } -static lean_object* _init_l_Lean_Parser_Command_namedName_formatter___closed__8() { +static lean_object* _init_l_Lean_Parser_Command_namedName_formatter___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_namedName_formatter___closed__4; -x_2 = l_Lean_Parser_Command_namedName_formatter___closed__7; +x_1 = l_Lean_Parser_Command_namedName_formatter___closed__5; +x_2 = l_Lean_Parser_Command_namedName_formatter___closed__8; 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; } } -static lean_object* _init_l_Lean_Parser_Command_namedName_formatter___closed__9() { +static lean_object* _init_l_Lean_Parser_Command_namedName_formatter___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_namedName___closed__3; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_namedName_formatter___closed__8; +x_3 = l_Lean_Parser_Command_namedName_formatter___closed__9; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -8088,7 +8127,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_namedName_formatter(lean_object* { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Command_namedName_formatter___closed__1; -x_7 = l_Lean_Parser_Command_namedName_formatter___closed__9; +x_7 = l_Lean_Parser_Command_namedName_formatter___closed__10; 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; } @@ -8914,8 +8953,18 @@ return x_7; static lean_object* _init_l_Lean_Parser_Command_namedName_parenthesizer___closed__2() { _start: { -lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; +lean_object* x_1; lean_object* x_2; x_1 = l_Lean_Parser_Command_namedName___closed__5; +x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_parenthesizer), 6, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Command_namedName_parenthesizer___closed__3() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Command_namedName___closed__7; x_2 = 0; x_3 = lean_box(x_2); x_4 = lean_alloc_closure((void*)(l_Lean_Parser_nonReservedSymbol_parenthesizer___boxed), 7, 2); @@ -8924,39 +8973,39 @@ lean_closure_set(x_4, 1, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Command_namedName_parenthesizer___closed__3() { +static lean_object* _init_l_Lean_Parser_Command_namedName_parenthesizer___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Syntax_paren_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Command_namedName_parenthesizer___closed__2; +x_1 = l_Lean_Parser_Command_namedName_parenthesizer___closed__2; +x_2 = l_Lean_Parser_Command_namedName_parenthesizer___closed__3; 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; } } -static lean_object* _init_l_Lean_Parser_Command_namedName_parenthesizer___closed__4() { +static lean_object* _init_l_Lean_Parser_Command_namedName_parenthesizer___closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_namedName_parenthesizer___closed__3; +x_1 = l_Lean_Parser_Command_namedName_parenthesizer___closed__4; x_2 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_atomic_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_namedName_parenthesizer___closed__5() { +static lean_object* _init_l_Lean_Parser_Command_namedName_parenthesizer___closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_namedName___closed__9; +x_1 = l_Lean_Parser_Command_namedName___closed__11; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_namedName_parenthesizer___closed__6() { +static lean_object* _init_l_Lean_Parser_Command_namedName_parenthesizer___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -8968,37 +9017,37 @@ lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_namedName_parenthesizer___closed__7() { +static lean_object* _init_l_Lean_Parser_Command_namedName_parenthesizer___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_namedName_parenthesizer___closed__5; -x_2 = l_Lean_Parser_Command_namedName_parenthesizer___closed__6; +x_1 = l_Lean_Parser_Command_namedName_parenthesizer___closed__6; +x_2 = l_Lean_Parser_Command_namedName_parenthesizer___closed__7; 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; } } -static lean_object* _init_l_Lean_Parser_Command_namedName_parenthesizer___closed__8() { +static lean_object* _init_l_Lean_Parser_Command_namedName_parenthesizer___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_namedName_parenthesizer___closed__4; -x_2 = l_Lean_Parser_Command_namedName_parenthesizer___closed__7; +x_1 = l_Lean_Parser_Command_namedName_parenthesizer___closed__5; +x_2 = l_Lean_Parser_Command_namedName_parenthesizer___closed__8; 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; } } -static lean_object* _init_l_Lean_Parser_Command_namedName_parenthesizer___closed__9() { +static lean_object* _init_l_Lean_Parser_Command_namedName_parenthesizer___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_namedName___closed__3; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_namedName_parenthesizer___closed__8; +x_3 = l_Lean_Parser_Command_namedName_parenthesizer___closed__9; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -9011,7 +9060,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_namedName_parenthesizer(lean_obje { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Command_namedName_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Command_namedName_parenthesizer___closed__9; +x_7 = l_Lean_Parser_Command_namedName_parenthesizer___closed__10; 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; } @@ -9458,7 +9507,7 @@ static lean_object* _init_l_Lean_Parser_Command_optKind___closed__5() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_optKind___closed__4; -x_2 = l_Lean_Parser_Command_namedName___closed__11; +x_2 = l_Lean_Parser_Command_namedName___closed__13; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -9477,7 +9526,7 @@ static lean_object* _init_l_Lean_Parser_Command_optKind___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Syntax_paren___closed__5; +x_1 = l_Lean_Parser_Command_namedName___closed__6; x_2 = l_Lean_Parser_Command_optKind___closed__6; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; @@ -10693,7 +10742,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_macro__rules_declRa { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_unsigned_to_nat(80u); -x_2 = lean_unsigned_to_nat(45u); +x_2 = lean_unsigned_to_nat(44u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -10707,7 +10756,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Parser_Command_macro__rules_declRange___closed__1; x_2 = lean_unsigned_to_nat(26u); x_3 = l___regBuiltin_Lean_Parser_Command_macro__rules_declRange___closed__2; -x_4 = lean_unsigned_to_nat(45u); +x_4 = lean_unsigned_to_nat(44u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -10806,7 +10855,7 @@ static lean_object* _init_l_Lean_Parser_Command_optKind_formatter___closed__3() { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_optKind_formatter___closed__2; -x_2 = l_Lean_Parser_Command_namedName_formatter___closed__6; +x_2 = l_Lean_Parser_Command_namedName_formatter___closed__7; 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); @@ -10829,7 +10878,7 @@ static lean_object* _init_l_Lean_Parser_Command_optKind_formatter___closed__5() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Syntax_paren_formatter___closed__2; +x_1 = l_Lean_Parser_Command_namedName_formatter___closed__2; x_2 = l_Lean_Parser_Command_optKind_formatter___closed__4; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -11037,7 +11086,7 @@ static lean_object* _init_l_Lean_Parser_Command_optKind_parenthesizer___closed__ { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_optKind_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Command_namedName_parenthesizer___closed__6; +x_2 = l_Lean_Parser_Command_namedName_parenthesizer___closed__7; 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); @@ -11060,7 +11109,7 @@ static lean_object* _init_l_Lean_Parser_Command_optKind_parenthesizer___closed__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Syntax_paren_parenthesizer___closed__2; +x_1 = l_Lean_Parser_Command_namedName_parenthesizer___closed__2; x_2 = l_Lean_Parser_Command_optKind_parenthesizer___closed__4; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -11302,13 +11351,23 @@ return x_3; static lean_object* _init_l_Lean_Parser_Command_syntax___closed__7() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_skip; +x_2 = l_Lean_Parser_Command_syntax___closed__6; +x_3 = l_Lean_Parser_andthen(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Command_syntax___closed__8() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_syntax___closed__6; +x_1 = l_Lean_Parser_Command_syntax___closed__7; x_2 = l_Lean_Parser_many1(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_syntax___closed__8() { +static lean_object* _init_l_Lean_Parser_Command_syntax___closed__9() { _start: { lean_object* x_1; @@ -11316,31 +11375,21 @@ x_1 = lean_mk_string_from_bytes(" : ", 3); return x_1; } } -static lean_object* _init_l_Lean_Parser_Command_syntax___closed__9() { +static lean_object* _init_l_Lean_Parser_Command_syntax___closed__10() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_syntax___closed__8; +x_1 = l_Lean_Parser_Command_syntax___closed__9; x_2 = l_Lean_Parser_symbol(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_syntax___closed__10() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_syntax___closed__9; -x_2 = l_Lean_Parser_ident; -x_3 = l_Lean_Parser_andthen(x_1, x_2); -return x_3; -} -} static lean_object* _init_l_Lean_Parser_Command_syntax___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_syntax___closed__7; -x_2 = l_Lean_Parser_Command_syntax___closed__10; +x_1 = l_Lean_Parser_Command_syntax___closed__10; +x_2 = l_Lean_Parser_ident; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -11349,7 +11398,7 @@ static lean_object* _init_l_Lean_Parser_Command_syntax___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_optNamedPrio; +x_1 = l_Lean_Parser_Command_syntax___closed__8; x_2 = l_Lean_Parser_Command_syntax___closed__11; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; @@ -11359,7 +11408,7 @@ static lean_object* _init_l_Lean_Parser_Command_syntax___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_optNamedName; +x_1 = l_Lean_Parser_Command_optNamedPrio; x_2 = l_Lean_Parser_Command_syntax___closed__12; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; @@ -11369,7 +11418,7 @@ static lean_object* _init_l_Lean_Parser_Command_syntax___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_optPrecedence; +x_1 = l_Lean_Parser_Command_optNamedName; x_2 = l_Lean_Parser_Command_syntax___closed__13; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; @@ -11379,7 +11428,7 @@ static lean_object* _init_l_Lean_Parser_Command_syntax___closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_syntax___closed__5; +x_1 = l_Lean_Parser_optPrecedence; x_2 = l_Lean_Parser_Command_syntax___closed__14; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; @@ -11389,7 +11438,7 @@ static lean_object* _init_l_Lean_Parser_Command_syntax___closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_attrKind; +x_1 = l_Lean_Parser_Command_syntax___closed__5; x_2 = l_Lean_Parser_Command_syntax___closed__15; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; @@ -11399,7 +11448,7 @@ static lean_object* _init_l_Lean_Parser_Command_syntax___closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_mixfix___closed__5; +x_1 = l_Lean_Parser_Term_attrKind; x_2 = l_Lean_Parser_Command_syntax___closed__16; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; @@ -11409,7 +11458,7 @@ static lean_object* _init_l_Lean_Parser_Command_syntax___closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_mixfix___closed__4; +x_1 = l_Lean_Parser_Command_mixfix___closed__5; x_2 = l_Lean_Parser_Command_syntax___closed__17; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; @@ -11418,30 +11467,40 @@ return x_3; static lean_object* _init_l_Lean_Parser_Command_syntax___closed__19() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Command_mixfix___closed__4; +x_2 = l_Lean_Parser_Command_syntax___closed__18; +x_3 = l_Lean_Parser_andthen(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Command_syntax___closed__20() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_syntax___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_syntax___closed__18; +x_3 = l_Lean_Parser_Command_syntax___closed__19; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Command_syntax___closed__20() { +static lean_object* _init_l_Lean_Parser_Command_syntax___closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_syntax___closed__3; -x_2 = l_Lean_Parser_Command_syntax___closed__19; +x_2 = l_Lean_Parser_Command_syntax___closed__20; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_syntax___closed__21() { +static lean_object* _init_l_Lean_Parser_Command_syntax___closed__22() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_syntax___closed__2; -x_2 = l_Lean_Parser_Command_syntax___closed__20; +x_2 = l_Lean_Parser_Command_syntax___closed__21; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -11450,7 +11509,7 @@ static lean_object* _init_l_Lean_Parser_Command_syntax() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Command_syntax___closed__21; +x_1 = l_Lean_Parser_Command_syntax___closed__22; return x_1; } } @@ -11484,7 +11543,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_syntax_declRange___ { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_unsigned_to_nat(83u); -x_2 = lean_unsigned_to_nat(110u); +x_2 = lean_unsigned_to_nat(121u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -11498,7 +11557,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Parser_Command_syntax_declRange___closed__1; x_2 = lean_unsigned_to_nat(26u); x_3 = l___regBuiltin_Lean_Parser_Command_syntax_declRange___closed__2; -x_4 = lean_unsigned_to_nat(110u); +x_4 = lean_unsigned_to_nat(121u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -11600,18 +11659,40 @@ return x_2; static lean_object* _init_l_Lean_Parser_Command_syntax_formatter___closed__3() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Command_mixfix_formatter___closed__10; +x_2 = l_Lean_Parser_Syntax_paren_formatter___closed__3; +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; +} +} +static lean_object* _init_l_Lean_Parser_Command_syntax_formatter___closed__4() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_syntax___closed__8; +x_1 = l_Lean_Parser_Command_syntax_formatter___closed__3; +x_2 = lean_alloc_closure((void*)(l_Lean_Parser_many1_formatter), 6, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Command_syntax_formatter___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Command_syntax___closed__9; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_syntax_formatter___closed__4() { +static lean_object* _init_l_Lean_Parser_Command_syntax_formatter___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_syntax_formatter___closed__3; +x_1 = l_Lean_Parser_Command_syntax_formatter___closed__5; x_2 = l_Lean_Parser_Syntax_cat_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); @@ -11619,109 +11700,109 @@ lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_syntax_formatter___closed__5() { +static lean_object* _init_l_Lean_Parser_Command_syntax_formatter___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Syntax_paren_formatter___closed__4; -x_2 = l_Lean_Parser_Command_syntax_formatter___closed__4; +x_1 = l_Lean_Parser_Command_syntax_formatter___closed__4; +x_2 = l_Lean_Parser_Command_syntax_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; } } -static lean_object* _init_l_Lean_Parser_Command_syntax_formatter___closed__6() { +static lean_object* _init_l_Lean_Parser_Command_syntax_formatter___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_mixfix_formatter___closed__12; -x_2 = l_Lean_Parser_Command_syntax_formatter___closed__5; +x_2 = l_Lean_Parser_Command_syntax_formatter___closed__7; 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; } } -static lean_object* _init_l_Lean_Parser_Command_syntax_formatter___closed__7() { +static lean_object* _init_l_Lean_Parser_Command_syntax_formatter___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_mixfix_formatter___closed__14; -x_2 = l_Lean_Parser_Command_syntax_formatter___closed__6; +x_2 = l_Lean_Parser_Command_syntax_formatter___closed__8; 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; } } -static lean_object* _init_l_Lean_Parser_Command_syntax_formatter___closed__8() { +static lean_object* _init_l_Lean_Parser_Command_syntax_formatter___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Syntax_cat_formatter___closed__3; -x_2 = l_Lean_Parser_Command_syntax_formatter___closed__7; +x_2 = l_Lean_Parser_Command_syntax_formatter___closed__9; 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; } } -static lean_object* _init_l_Lean_Parser_Command_syntax_formatter___closed__9() { +static lean_object* _init_l_Lean_Parser_Command_syntax_formatter___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_syntax_formatter___closed__2; -x_2 = l_Lean_Parser_Command_syntax_formatter___closed__8; +x_2 = l_Lean_Parser_Command_syntax_formatter___closed__10; 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; } } -static lean_object* _init_l_Lean_Parser_Command_syntax_formatter___closed__10() { +static lean_object* _init_l_Lean_Parser_Command_syntax_formatter___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_mixfix_formatter___closed__19; -x_2 = l_Lean_Parser_Command_syntax_formatter___closed__9; +x_2 = l_Lean_Parser_Command_syntax_formatter___closed__11; 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; } } -static lean_object* _init_l_Lean_Parser_Command_syntax_formatter___closed__11() { +static lean_object* _init_l_Lean_Parser_Command_syntax_formatter___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_mixfix_formatter___closed__5; -x_2 = l_Lean_Parser_Command_syntax_formatter___closed__10; +x_2 = l_Lean_Parser_Command_syntax_formatter___closed__12; 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; } } -static lean_object* _init_l_Lean_Parser_Command_syntax_formatter___closed__12() { +static lean_object* _init_l_Lean_Parser_Command_syntax_formatter___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_mixfix_formatter___closed__3; -x_2 = l_Lean_Parser_Command_syntax_formatter___closed__11; +x_2 = l_Lean_Parser_Command_syntax_formatter___closed__13; 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; } } -static lean_object* _init_l_Lean_Parser_Command_syntax_formatter___closed__13() { +static lean_object* _init_l_Lean_Parser_Command_syntax_formatter___closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_syntax___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_syntax_formatter___closed__12; +x_3 = l_Lean_Parser_Command_syntax_formatter___closed__14; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -11734,7 +11815,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_syntax_formatter(lean_object* x_1 { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Command_syntax_formatter___closed__1; -x_7 = l_Lean_Parser_Command_syntax_formatter___closed__13; +x_7 = l_Lean_Parser_Command_syntax_formatter___closed__15; 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; } @@ -11813,28 +11894,40 @@ return x_2; static lean_object* _init_l_Lean_Parser_Command_syntax_parenthesizer___closed__4() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Command_mixfix_parenthesizer___closed__10; +x_2 = l_Lean_Parser_Command_syntax_parenthesizer___closed__3; +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; +} +} +static lean_object* _init_l_Lean_Parser_Command_syntax_parenthesizer___closed__5() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_syntax_parenthesizer___closed__3; +x_1 = l_Lean_Parser_Command_syntax_parenthesizer___closed__4; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_many1_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_syntax_parenthesizer___closed__5() { +static lean_object* _init_l_Lean_Parser_Command_syntax_parenthesizer___closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_syntax___closed__8; +x_1 = l_Lean_Parser_Command_syntax___closed__9; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_syntax_parenthesizer___closed__6() { +static lean_object* _init_l_Lean_Parser_Command_syntax_parenthesizer___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_syntax_parenthesizer___closed__5; +x_1 = l_Lean_Parser_Command_syntax_parenthesizer___closed__6; x_2 = l_Lean_Parser_Syntax_cat_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); @@ -11842,109 +11935,109 @@ lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_syntax_parenthesizer___closed__7() { +static lean_object* _init_l_Lean_Parser_Command_syntax_parenthesizer___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_syntax_parenthesizer___closed__4; -x_2 = l_Lean_Parser_Command_syntax_parenthesizer___closed__6; +x_1 = l_Lean_Parser_Command_syntax_parenthesizer___closed__5; +x_2 = l_Lean_Parser_Command_syntax_parenthesizer___closed__7; 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; } } -static lean_object* _init_l_Lean_Parser_Command_syntax_parenthesizer___closed__8() { +static lean_object* _init_l_Lean_Parser_Command_syntax_parenthesizer___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_mixfix_parenthesizer___closed__12; -x_2 = l_Lean_Parser_Command_syntax_parenthesizer___closed__7; +x_2 = l_Lean_Parser_Command_syntax_parenthesizer___closed__8; 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; } } -static lean_object* _init_l_Lean_Parser_Command_syntax_parenthesizer___closed__9() { +static lean_object* _init_l_Lean_Parser_Command_syntax_parenthesizer___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_mixfix_parenthesizer___closed__14; -x_2 = l_Lean_Parser_Command_syntax_parenthesizer___closed__8; +x_2 = l_Lean_Parser_Command_syntax_parenthesizer___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); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_syntax_parenthesizer___closed__10() { +static lean_object* _init_l_Lean_Parser_Command_syntax_parenthesizer___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Syntax_cat_parenthesizer___closed__3; -x_2 = l_Lean_Parser_Command_syntax_parenthesizer___closed__9; +x_2 = l_Lean_Parser_Command_syntax_parenthesizer___closed__10; 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; } } -static lean_object* _init_l_Lean_Parser_Command_syntax_parenthesizer___closed__11() { +static lean_object* _init_l_Lean_Parser_Command_syntax_parenthesizer___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_syntax_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Command_syntax_parenthesizer___closed__10; +x_2 = l_Lean_Parser_Command_syntax_parenthesizer___closed__11; 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; } } -static lean_object* _init_l_Lean_Parser_Command_syntax_parenthesizer___closed__12() { +static lean_object* _init_l_Lean_Parser_Command_syntax_parenthesizer___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_mixfix_parenthesizer___closed__19; -x_2 = l_Lean_Parser_Command_syntax_parenthesizer___closed__11; +x_2 = l_Lean_Parser_Command_syntax_parenthesizer___closed__12; 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; } } -static lean_object* _init_l_Lean_Parser_Command_syntax_parenthesizer___closed__13() { +static lean_object* _init_l_Lean_Parser_Command_syntax_parenthesizer___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_mixfix_parenthesizer___closed__5; -x_2 = l_Lean_Parser_Command_syntax_parenthesizer___closed__12; +x_2 = l_Lean_Parser_Command_syntax_parenthesizer___closed__13; 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; } } -static lean_object* _init_l_Lean_Parser_Command_syntax_parenthesizer___closed__14() { +static lean_object* _init_l_Lean_Parser_Command_syntax_parenthesizer___closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_mixfix_parenthesizer___closed__3; -x_2 = l_Lean_Parser_Command_syntax_parenthesizer___closed__13; +x_2 = l_Lean_Parser_Command_syntax_parenthesizer___closed__14; 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; } } -static lean_object* _init_l_Lean_Parser_Command_syntax_parenthesizer___closed__15() { +static lean_object* _init_l_Lean_Parser_Command_syntax_parenthesizer___closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_syntax___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_syntax_parenthesizer___closed__14; +x_3 = l_Lean_Parser_Command_syntax_parenthesizer___closed__15; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -11957,7 +12050,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_syntax_parenthesizer(lean_object* { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Command_syntax_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Command_syntax_parenthesizer___closed__15; +x_7 = l_Lean_Parser_Command_syntax_parenthesizer___closed__16; 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; } @@ -12031,7 +12124,7 @@ static lean_object* _init_l_Lean_Parser_Command_syntaxAbbrev___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_namedName___closed__10; +x_1 = l_Lean_Parser_Command_namedName___closed__12; x_2 = l_Lean_Parser_Syntax_paren___closed__7; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; @@ -12243,7 +12336,7 @@ static lean_object* _init_l_Lean_Parser_Command_syntaxAbbrev_formatter___closed_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_namedName_formatter___closed__5; +x_1 = l_Lean_Parser_Command_namedName_formatter___closed__6; x_2 = l_Lean_Parser_Syntax_paren_formatter___closed__4; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -12366,7 +12459,7 @@ static lean_object* _init_l_Lean_Parser_Command_syntaxAbbrev_parenthesizer___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_namedName_parenthesizer___closed__5; +x_1 = l_Lean_Parser_Command_namedName_parenthesizer___closed__6; x_2 = l_Lean_Parser_Syntax_paren_parenthesizer___closed__4; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -12687,7 +12780,7 @@ static lean_object* _init_l_Lean_Parser_Command_catBehavior___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_namedName___closed__10; +x_1 = l_Lean_Parser_Command_namedName___closed__12; x_2 = l_Lean_Parser_Command_catBehavior___closed__4; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; @@ -12707,7 +12800,7 @@ static lean_object* _init_l_Lean_Parser_Command_catBehavior___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Syntax_paren___closed__5; +x_1 = l_Lean_Parser_Command_namedName___closed__6; x_2 = l_Lean_Parser_Command_catBehavior___closed__6; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; @@ -13180,7 +13273,7 @@ static lean_object* _init_l_Lean_Parser_Command_catBehavior_formatter___closed__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_namedName_formatter___closed__5; +x_1 = l_Lean_Parser_Command_namedName_formatter___closed__6; x_2 = l_Lean_Parser_Command_catBehavior_formatter___closed__3; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -13204,7 +13297,7 @@ static lean_object* _init_l_Lean_Parser_Command_catBehavior_formatter___closed__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Syntax_paren_formatter___closed__2; +x_1 = l_Lean_Parser_Command_namedName_formatter___closed__2; x_2 = l_Lean_Parser_Command_catBehavior_formatter___closed__5; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -13567,7 +13660,7 @@ static lean_object* _init_l_Lean_Parser_Command_catBehavior_parenthesizer___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_namedName_parenthesizer___closed__5; +x_1 = l_Lean_Parser_Command_namedName_parenthesizer___closed__6; x_2 = l_Lean_Parser_Command_catBehavior_parenthesizer___closed__3; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -13591,7 +13684,7 @@ static lean_object* _init_l_Lean_Parser_Command_catBehavior_parenthesizer___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Syntax_paren_parenthesizer___closed__2; +x_1 = l_Lean_Parser_Command_namedName_parenthesizer___closed__2; x_2 = l_Lean_Parser_Command_catBehavior_parenthesizer___closed__5; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -13989,7 +14082,7 @@ static lean_object* _init_l_Lean_Parser_Command_macroTail___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_syntax___closed__10; +x_1 = l_Lean_Parser_Command_syntax___closed__11; x_2 = l_Lean_Parser_atomic(x_1); return x_2; } @@ -14088,25 +14181,27 @@ return x_5; static lean_object* _init_l_Lean_Parser_Command_macro___closed__4() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("macro ", 6); -return x_1; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Command_macro___closed__1; +x_2 = l_Lean_Parser_symbol(x_1); +return x_2; } } static lean_object* _init_l_Lean_Parser_Command_macro___closed__5() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_macro___closed__4; -x_2 = l_Lean_Parser_symbol(x_1); -return x_2; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_skip; +x_2 = l_Lean_Parser_Command_macroArg; +x_3 = l_Lean_Parser_andthen(x_1, x_2); +return x_3; } } static lean_object* _init_l_Lean_Parser_Command_macro___closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_macroArg; +x_1 = l_Lean_Parser_Command_macro___closed__5; x_2 = l_Lean_Parser_many1(x_1); return x_2; } @@ -14155,7 +14250,7 @@ static lean_object* _init_l_Lean_Parser_Command_macro___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_macro___closed__5; +x_1 = l_Lean_Parser_Command_macro___closed__4; x_2 = l_Lean_Parser_Command_macro___closed__10; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; @@ -14270,7 +14365,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_macro_declRange___c { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_unsigned_to_nat(97u); -x_2 = lean_unsigned_to_nat(90u); +x_2 = lean_unsigned_to_nat(102u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -14284,7 +14379,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Parser_Command_macro_declRange___closed__1; x_2 = lean_unsigned_to_nat(26u); x_3 = l___regBuiltin_Lean_Parser_Command_macro_declRange___closed__2; -x_4 = lean_unsigned_to_nat(90u); +x_4 = lean_unsigned_to_nat(102u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -14593,7 +14688,7 @@ static lean_object* _init_l_Lean_Parser_Command_macroTail_formatter___closed__2( _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_syntax_formatter___closed__4; +x_1 = l_Lean_Parser_Command_syntax_formatter___closed__6; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_atomic_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; @@ -14702,7 +14797,7 @@ static lean_object* _init_l_Lean_Parser_Command_macro_formatter___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_macro___closed__4; +x_1 = l_Lean_Parser_Command_macro___closed__1; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; @@ -14711,18 +14806,30 @@ return x_2; static lean_object* _init_l_Lean_Parser_Command_macro_formatter___closed__3() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Command_mixfix_formatter___closed__10; +x_2 = l___regBuiltin_Lean_Parser_Command_macroArg_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; +} +} +static lean_object* _init_l_Lean_Parser_Command_macro_formatter___closed__4() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l___regBuiltin_Lean_Parser_Command_macroArg_formatter___closed__2; +x_1 = l_Lean_Parser_Command_macro_formatter___closed__3; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_many1_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_macro_formatter___closed__4() { +static lean_object* _init_l_Lean_Parser_Command_macro_formatter___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_macro_formatter___closed__3; +x_1 = l_Lean_Parser_Command_macro_formatter___closed__4; x_2 = l___regBuiltin_Lean_Parser_Command_macroTail_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); @@ -14730,107 +14837,107 @@ lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_macro_formatter___closed__5() { +static lean_object* _init_l_Lean_Parser_Command_macro_formatter___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_mixfix_formatter___closed__12; -x_2 = l_Lean_Parser_Command_macro_formatter___closed__4; +x_2 = l_Lean_Parser_Command_macro_formatter___closed__5; 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; } } -static lean_object* _init_l_Lean_Parser_Command_macro_formatter___closed__6() { +static lean_object* _init_l_Lean_Parser_Command_macro_formatter___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_mixfix_formatter___closed__14; -x_2 = l_Lean_Parser_Command_macro_formatter___closed__5; +x_2 = l_Lean_Parser_Command_macro_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; } } -static lean_object* _init_l_Lean_Parser_Command_macro_formatter___closed__7() { +static lean_object* _init_l_Lean_Parser_Command_macro_formatter___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Syntax_cat_formatter___closed__3; -x_2 = l_Lean_Parser_Command_macro_formatter___closed__6; +x_2 = l_Lean_Parser_Command_macro_formatter___closed__7; 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; } } -static lean_object* _init_l_Lean_Parser_Command_macro_formatter___closed__8() { +static lean_object* _init_l_Lean_Parser_Command_macro_formatter___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_macro_formatter___closed__2; -x_2 = l_Lean_Parser_Command_macro_formatter___closed__7; +x_2 = l_Lean_Parser_Command_macro_formatter___closed__8; 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; } } -static lean_object* _init_l_Lean_Parser_Command_macro_formatter___closed__9() { +static lean_object* _init_l_Lean_Parser_Command_macro_formatter___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_mixfix_formatter___closed__19; -x_2 = l_Lean_Parser_Command_macro_formatter___closed__8; +x_2 = l_Lean_Parser_Command_macro_formatter___closed__9; 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; } } -static lean_object* _init_l_Lean_Parser_Command_macro_formatter___closed__10() { +static lean_object* _init_l_Lean_Parser_Command_macro_formatter___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_mixfix_formatter___closed__5; -x_2 = l_Lean_Parser_Command_macro_formatter___closed__9; +x_2 = l_Lean_Parser_Command_macro_formatter___closed__10; 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; } } -static lean_object* _init_l_Lean_Parser_Command_macro_formatter___closed__11() { +static lean_object* _init_l_Lean_Parser_Command_macro_formatter___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_mixfix_formatter___closed__3; -x_2 = l_Lean_Parser_Command_macro_formatter___closed__10; +x_2 = l_Lean_Parser_Command_macro_formatter___closed__11; 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; } } -static lean_object* _init_l_Lean_Parser_Command_macro_formatter___closed__12() { +static lean_object* _init_l_Lean_Parser_Command_macro_formatter___closed__13() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_macro_formatter___closed__11; +x_1 = l_Lean_Parser_Command_macro_formatter___closed__12; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_suppressInsideQuot_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_macro_formatter___closed__13() { +static lean_object* _init_l_Lean_Parser_Command_macro_formatter___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_macro___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_macro_formatter___closed__12; +x_3 = l_Lean_Parser_Command_macro_formatter___closed__13; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -14843,7 +14950,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_macro_formatter(lean_object* x_1, { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Command_macro_formatter___closed__1; -x_7 = l_Lean_Parser_Command_macro_formatter___closed__13; +x_7 = l_Lean_Parser_Command_macro_formatter___closed__14; 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; } @@ -15119,7 +15226,7 @@ static lean_object* _init_l_Lean_Parser_Command_macroTail_parenthesizer___closed _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_syntax_parenthesizer___closed__6; +x_1 = l_Lean_Parser_Command_syntax_parenthesizer___closed__7; x_2 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_atomic_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; @@ -15228,7 +15335,7 @@ static lean_object* _init_l_Lean_Parser_Command_macro_parenthesizer___closed__2( _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_macro___closed__4; +x_1 = l_Lean_Parser_Command_macro___closed__1; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; @@ -15237,18 +15344,30 @@ return x_2; static lean_object* _init_l_Lean_Parser_Command_macro_parenthesizer___closed__3() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Command_mixfix_parenthesizer___closed__10; +x_2 = l___regBuiltin_Lean_Parser_Command_macroArg_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; +} +} +static lean_object* _init_l_Lean_Parser_Command_macro_parenthesizer___closed__4() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l___regBuiltin_Lean_Parser_Command_macroArg_parenthesizer___closed__2; +x_1 = l_Lean_Parser_Command_macro_parenthesizer___closed__3; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_many1_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_macro_parenthesizer___closed__4() { +static lean_object* _init_l_Lean_Parser_Command_macro_parenthesizer___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_macro_parenthesizer___closed__3; +x_1 = l_Lean_Parser_Command_macro_parenthesizer___closed__4; x_2 = l___regBuiltin_Lean_Parser_Command_macroTail_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); @@ -15256,107 +15375,107 @@ lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_macro_parenthesizer___closed__5() { +static lean_object* _init_l_Lean_Parser_Command_macro_parenthesizer___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_mixfix_parenthesizer___closed__12; -x_2 = l_Lean_Parser_Command_macro_parenthesizer___closed__4; +x_2 = l_Lean_Parser_Command_macro_parenthesizer___closed__5; 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; } } -static lean_object* _init_l_Lean_Parser_Command_macro_parenthesizer___closed__6() { +static lean_object* _init_l_Lean_Parser_Command_macro_parenthesizer___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_mixfix_parenthesizer___closed__14; -x_2 = l_Lean_Parser_Command_macro_parenthesizer___closed__5; +x_2 = l_Lean_Parser_Command_macro_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; } } -static lean_object* _init_l_Lean_Parser_Command_macro_parenthesizer___closed__7() { +static lean_object* _init_l_Lean_Parser_Command_macro_parenthesizer___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Syntax_cat_parenthesizer___closed__3; -x_2 = l_Lean_Parser_Command_macro_parenthesizer___closed__6; +x_2 = l_Lean_Parser_Command_macro_parenthesizer___closed__7; 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; } } -static lean_object* _init_l_Lean_Parser_Command_macro_parenthesizer___closed__8() { +static lean_object* _init_l_Lean_Parser_Command_macro_parenthesizer___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_macro_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Command_macro_parenthesizer___closed__7; +x_2 = l_Lean_Parser_Command_macro_parenthesizer___closed__8; 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; } } -static lean_object* _init_l_Lean_Parser_Command_macro_parenthesizer___closed__9() { +static lean_object* _init_l_Lean_Parser_Command_macro_parenthesizer___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_mixfix_parenthesizer___closed__19; -x_2 = l_Lean_Parser_Command_macro_parenthesizer___closed__8; +x_2 = l_Lean_Parser_Command_macro_parenthesizer___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); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_macro_parenthesizer___closed__10() { +static lean_object* _init_l_Lean_Parser_Command_macro_parenthesizer___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_mixfix_parenthesizer___closed__5; -x_2 = l_Lean_Parser_Command_macro_parenthesizer___closed__9; +x_2 = l_Lean_Parser_Command_macro_parenthesizer___closed__10; 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; } } -static lean_object* _init_l_Lean_Parser_Command_macro_parenthesizer___closed__11() { +static lean_object* _init_l_Lean_Parser_Command_macro_parenthesizer___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_mixfix_parenthesizer___closed__3; -x_2 = l_Lean_Parser_Command_macro_parenthesizer___closed__10; +x_2 = l_Lean_Parser_Command_macro_parenthesizer___closed__11; 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; } } -static lean_object* _init_l_Lean_Parser_Command_macro_parenthesizer___closed__12() { +static lean_object* _init_l_Lean_Parser_Command_macro_parenthesizer___closed__13() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_macro_parenthesizer___closed__11; +x_1 = l_Lean_Parser_Command_macro_parenthesizer___closed__12; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_suppressInsideQuot_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_macro_parenthesizer___closed__13() { +static lean_object* _init_l_Lean_Parser_Command_macro_parenthesizer___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_macro___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_macro_parenthesizer___closed__12; +x_3 = l_Lean_Parser_Command_macro_parenthesizer___closed__13; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -15369,7 +15488,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_macro_parenthesizer(lean_object* { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Command_macro_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Command_macro_parenthesizer___closed__13; +x_7 = l_Lean_Parser_Command_macro_parenthesizer___closed__14; 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; } @@ -15452,7 +15571,7 @@ static lean_object* _init_l_Lean_Parser_Command_elab__rules___closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_syntax___closed__10; +x_1 = l_Lean_Parser_Command_syntax___closed__11; x_2 = l_Lean_Parser_optional(x_1); return x_2; } @@ -15642,7 +15761,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_elab__rules_declRan { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_unsigned_to_nat(100u); -x_2 = lean_unsigned_to_nat(103u); +x_2 = lean_unsigned_to_nat(102u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -15656,7 +15775,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Parser_Command_elab__rules_declRange___closed__1; x_2 = lean_unsigned_to_nat(26u); x_3 = l___regBuiltin_Lean_Parser_Command_elab__rules_declRange___closed__2; -x_4 = lean_unsigned_to_nat(103u); +x_4 = lean_unsigned_to_nat(102u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -15759,7 +15878,7 @@ static lean_object* _init_l_Lean_Parser_Command_elab__rules_formatter___closed__ _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_syntax_formatter___closed__4; +x_1 = l_Lean_Parser_Command_syntax_formatter___closed__6; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_optional_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; @@ -15980,7 +16099,7 @@ static lean_object* _init_l_Lean_Parser_Command_elab__rules_parenthesizer___clos _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_syntax_parenthesizer___closed__6; +x_1 = l_Lean_Parser_Command_syntax_parenthesizer___closed__7; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_optional_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; @@ -16223,7 +16342,7 @@ static lean_object* _init_l_Lean_Parser_Command_elabTail___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_syntax___closed__9; +x_1 = l_Lean_Parser_Command_syntax___closed__10; x_2 = l_Lean_Parser_Command_elabTail___closed__4; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; @@ -16332,25 +16451,27 @@ return x_5; static lean_object* _init_l_Lean_Parser_Command_elab___closed__4() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("elab ", 5); -return x_1; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Command_elab___closed__1; +x_2 = l_Lean_Parser_symbol(x_1); +return x_2; } } static lean_object* _init_l_Lean_Parser_Command_elab___closed__5() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_elab___closed__4; -x_2 = l_Lean_Parser_symbol(x_1); -return x_2; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_skip; +x_2 = l_Lean_Parser_Command_elabArg; +x_3 = l_Lean_Parser_andthen(x_1, x_2); +return x_3; } } static lean_object* _init_l_Lean_Parser_Command_elab___closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_elabArg; +x_1 = l_Lean_Parser_Command_elab___closed__5; x_2 = l_Lean_Parser_many1(x_1); return x_2; } @@ -16399,7 +16520,7 @@ static lean_object* _init_l_Lean_Parser_Command_elab___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_elab___closed__5; +x_1 = l_Lean_Parser_Command_elab___closed__4; x_2 = l_Lean_Parser_Command_elab___closed__10; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; @@ -16514,7 +16635,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Command_elab_declRange___cl { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_unsigned_to_nat(105u); -x_2 = lean_unsigned_to_nat(87u); +x_2 = lean_unsigned_to_nat(99u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -16528,7 +16649,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Parser_Command_elab_declRange___closed__1; x_2 = lean_unsigned_to_nat(26u); x_3 = l___regBuiltin_Lean_Parser_Command_elab_declRange___closed__2; -x_4 = lean_unsigned_to_nat(87u); +x_4 = lean_unsigned_to_nat(99u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -16641,7 +16762,7 @@ static lean_object* _init_l_Lean_Parser_Command_elabTail_formatter___closed__3() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_syntax_formatter___closed__3; +x_1 = l_Lean_Parser_Command_syntax_formatter___closed__5; x_2 = l_Lean_Parser_Command_elabTail_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); @@ -16762,7 +16883,7 @@ static lean_object* _init_l_Lean_Parser_Command_elab_formatter___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_elab___closed__4; +x_1 = l_Lean_Parser_Command_elab___closed__1; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; @@ -16779,18 +16900,30 @@ return x_1; static lean_object* _init_l_Lean_Parser_Command_elab_formatter___closed__4() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Command_mixfix_formatter___closed__10; +x_2 = l_Lean_Parser_Command_elab_formatter___closed__3; +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; +} +} +static lean_object* _init_l_Lean_Parser_Command_elab_formatter___closed__5() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_elab_formatter___closed__3; +x_1 = l_Lean_Parser_Command_elab_formatter___closed__4; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_many1_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_elab_formatter___closed__5() { +static lean_object* _init_l_Lean_Parser_Command_elab_formatter___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_elab_formatter___closed__4; +x_1 = l_Lean_Parser_Command_elab_formatter___closed__5; x_2 = l___regBuiltin_Lean_Parser_Command_elabTail_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); @@ -16798,107 +16931,107 @@ lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_elab_formatter___closed__6() { +static lean_object* _init_l_Lean_Parser_Command_elab_formatter___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_mixfix_formatter___closed__12; -x_2 = l_Lean_Parser_Command_elab_formatter___closed__5; +x_2 = l_Lean_Parser_Command_elab_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; } } -static lean_object* _init_l_Lean_Parser_Command_elab_formatter___closed__7() { +static lean_object* _init_l_Lean_Parser_Command_elab_formatter___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_mixfix_formatter___closed__14; -x_2 = l_Lean_Parser_Command_elab_formatter___closed__6; +x_2 = l_Lean_Parser_Command_elab_formatter___closed__7; 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; } } -static lean_object* _init_l_Lean_Parser_Command_elab_formatter___closed__8() { +static lean_object* _init_l_Lean_Parser_Command_elab_formatter___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Syntax_cat_formatter___closed__3; -x_2 = l_Lean_Parser_Command_elab_formatter___closed__7; +x_2 = l_Lean_Parser_Command_elab_formatter___closed__8; 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; } } -static lean_object* _init_l_Lean_Parser_Command_elab_formatter___closed__9() { +static lean_object* _init_l_Lean_Parser_Command_elab_formatter___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_elab_formatter___closed__2; -x_2 = l_Lean_Parser_Command_elab_formatter___closed__8; +x_2 = l_Lean_Parser_Command_elab_formatter___closed__9; 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; } } -static lean_object* _init_l_Lean_Parser_Command_elab_formatter___closed__10() { +static lean_object* _init_l_Lean_Parser_Command_elab_formatter___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_mixfix_formatter___closed__19; -x_2 = l_Lean_Parser_Command_elab_formatter___closed__9; +x_2 = l_Lean_Parser_Command_elab_formatter___closed__10; 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; } } -static lean_object* _init_l_Lean_Parser_Command_elab_formatter___closed__11() { +static lean_object* _init_l_Lean_Parser_Command_elab_formatter___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_mixfix_formatter___closed__5; -x_2 = l_Lean_Parser_Command_elab_formatter___closed__10; +x_2 = l_Lean_Parser_Command_elab_formatter___closed__11; 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; } } -static lean_object* _init_l_Lean_Parser_Command_elab_formatter___closed__12() { +static lean_object* _init_l_Lean_Parser_Command_elab_formatter___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_mixfix_formatter___closed__3; -x_2 = l_Lean_Parser_Command_elab_formatter___closed__11; +x_2 = l_Lean_Parser_Command_elab_formatter___closed__12; 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; } } -static lean_object* _init_l_Lean_Parser_Command_elab_formatter___closed__13() { +static lean_object* _init_l_Lean_Parser_Command_elab_formatter___closed__14() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_elab_formatter___closed__12; +x_1 = l_Lean_Parser_Command_elab_formatter___closed__13; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_suppressInsideQuot_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_elab_formatter___closed__14() { +static lean_object* _init_l_Lean_Parser_Command_elab_formatter___closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_elab___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_elab_formatter___closed__13; +x_3 = l_Lean_Parser_Command_elab_formatter___closed__14; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -16911,7 +17044,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_elab_formatter(lean_object* x_1, { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Command_elab_formatter___closed__1; -x_7 = l_Lean_Parser_Command_elab_formatter___closed__14; +x_7 = l_Lean_Parser_Command_elab_formatter___closed__15; 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; } @@ -16991,7 +17124,7 @@ static lean_object* _init_l_Lean_Parser_Command_elabTail_parenthesizer___closed_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_syntax_parenthesizer___closed__5; +x_1 = l_Lean_Parser_Command_syntax_parenthesizer___closed__6; x_2 = l_Lean_Parser_Command_elabTail_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); @@ -17112,7 +17245,7 @@ static lean_object* _init_l_Lean_Parser_Command_elab_parenthesizer___closed__2() _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_elab___closed__4; +x_1 = l_Lean_Parser_Command_elab___closed__1; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; @@ -17129,18 +17262,30 @@ return x_1; static lean_object* _init_l_Lean_Parser_Command_elab_parenthesizer___closed__4() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Command_mixfix_parenthesizer___closed__10; +x_2 = l_Lean_Parser_Command_elab_parenthesizer___closed__3; +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; +} +} +static lean_object* _init_l_Lean_Parser_Command_elab_parenthesizer___closed__5() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_elab_parenthesizer___closed__3; +x_1 = l_Lean_Parser_Command_elab_parenthesizer___closed__4; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_many1_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_elab_parenthesizer___closed__5() { +static lean_object* _init_l_Lean_Parser_Command_elab_parenthesizer___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Command_elab_parenthesizer___closed__4; +x_1 = l_Lean_Parser_Command_elab_parenthesizer___closed__5; x_2 = l___regBuiltin_Lean_Parser_Command_elabTail_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); @@ -17148,107 +17293,107 @@ lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_elab_parenthesizer___closed__6() { +static lean_object* _init_l_Lean_Parser_Command_elab_parenthesizer___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_mixfix_parenthesizer___closed__12; -x_2 = l_Lean_Parser_Command_elab_parenthesizer___closed__5; +x_2 = l_Lean_Parser_Command_elab_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; } } -static lean_object* _init_l_Lean_Parser_Command_elab_parenthesizer___closed__7() { +static lean_object* _init_l_Lean_Parser_Command_elab_parenthesizer___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_mixfix_parenthesizer___closed__14; -x_2 = l_Lean_Parser_Command_elab_parenthesizer___closed__6; +x_2 = l_Lean_Parser_Command_elab_parenthesizer___closed__7; 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; } } -static lean_object* _init_l_Lean_Parser_Command_elab_parenthesizer___closed__8() { +static lean_object* _init_l_Lean_Parser_Command_elab_parenthesizer___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Syntax_cat_parenthesizer___closed__3; -x_2 = l_Lean_Parser_Command_elab_parenthesizer___closed__7; +x_2 = l_Lean_Parser_Command_elab_parenthesizer___closed__8; 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; } } -static lean_object* _init_l_Lean_Parser_Command_elab_parenthesizer___closed__9() { +static lean_object* _init_l_Lean_Parser_Command_elab_parenthesizer___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_elab_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Command_elab_parenthesizer___closed__8; +x_2 = l_Lean_Parser_Command_elab_parenthesizer___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); return x_3; } } -static lean_object* _init_l_Lean_Parser_Command_elab_parenthesizer___closed__10() { +static lean_object* _init_l_Lean_Parser_Command_elab_parenthesizer___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_mixfix_parenthesizer___closed__19; -x_2 = l_Lean_Parser_Command_elab_parenthesizer___closed__9; +x_2 = l_Lean_Parser_Command_elab_parenthesizer___closed__10; 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; } } -static lean_object* _init_l_Lean_Parser_Command_elab_parenthesizer___closed__11() { +static lean_object* _init_l_Lean_Parser_Command_elab_parenthesizer___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_mixfix_parenthesizer___closed__5; -x_2 = l_Lean_Parser_Command_elab_parenthesizer___closed__10; +x_2 = l_Lean_Parser_Command_elab_parenthesizer___closed__11; 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; } } -static lean_object* _init_l_Lean_Parser_Command_elab_parenthesizer___closed__12() { +static lean_object* _init_l_Lean_Parser_Command_elab_parenthesizer___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Command_mixfix_parenthesizer___closed__3; -x_2 = l_Lean_Parser_Command_elab_parenthesizer___closed__11; +x_2 = l_Lean_Parser_Command_elab_parenthesizer___closed__12; 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; } } -static lean_object* _init_l_Lean_Parser_Command_elab_parenthesizer___closed__13() { +static lean_object* _init_l_Lean_Parser_Command_elab_parenthesizer___closed__14() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Command_elab_parenthesizer___closed__12; +x_1 = l_Lean_Parser_Command_elab_parenthesizer___closed__13; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_suppressInsideQuot_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Command_elab_parenthesizer___closed__14() { +static lean_object* _init_l_Lean_Parser_Command_elab_parenthesizer___closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Command_elab___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Command_elab_parenthesizer___closed__13; +x_3 = l_Lean_Parser_Command_elab_parenthesizer___closed__14; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -17261,7 +17406,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_elab_parenthesizer(lean_object* x { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Command_elab_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Command_elab_parenthesizer___closed__14; +x_7 = l_Lean_Parser_Command_elab_parenthesizer___closed__15; 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; } @@ -18212,6 +18357,10 @@ l_Lean_Parser_Command_namedName___closed__15 = _init_l_Lean_Parser_Command_named lean_mark_persistent(l_Lean_Parser_Command_namedName___closed__15); l_Lean_Parser_Command_namedName___closed__16 = _init_l_Lean_Parser_Command_namedName___closed__16(); lean_mark_persistent(l_Lean_Parser_Command_namedName___closed__16); +l_Lean_Parser_Command_namedName___closed__17 = _init_l_Lean_Parser_Command_namedName___closed__17(); +lean_mark_persistent(l_Lean_Parser_Command_namedName___closed__17); +l_Lean_Parser_Command_namedName___closed__18 = _init_l_Lean_Parser_Command_namedName___closed__18(); +lean_mark_persistent(l_Lean_Parser_Command_namedName___closed__18); l_Lean_Parser_Command_namedName = _init_l_Lean_Parser_Command_namedName(); lean_mark_persistent(l_Lean_Parser_Command_namedName); l_Lean_Parser_Command_optNamedName___closed__1 = _init_l_Lean_Parser_Command_optNamedName___closed__1(); @@ -18461,6 +18610,8 @@ l_Lean_Parser_Command_namedName_formatter___closed__8 = _init_l_Lean_Parser_Comm lean_mark_persistent(l_Lean_Parser_Command_namedName_formatter___closed__8); l_Lean_Parser_Command_namedName_formatter___closed__9 = _init_l_Lean_Parser_Command_namedName_formatter___closed__9(); lean_mark_persistent(l_Lean_Parser_Command_namedName_formatter___closed__9); +l_Lean_Parser_Command_namedName_formatter___closed__10 = _init_l_Lean_Parser_Command_namedName_formatter___closed__10(); +lean_mark_persistent(l_Lean_Parser_Command_namedName_formatter___closed__10); l___regBuiltin_Lean_Parser_Command_namedName_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_namedName_formatter___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_namedName_formatter___closed__1); l___regBuiltin_Lean_Parser_Command_namedName_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_namedName_formatter___closed__2(); @@ -18610,6 +18761,8 @@ l_Lean_Parser_Command_namedName_parenthesizer___closed__8 = _init_l_Lean_Parser_ lean_mark_persistent(l_Lean_Parser_Command_namedName_parenthesizer___closed__8); l_Lean_Parser_Command_namedName_parenthesizer___closed__9 = _init_l_Lean_Parser_Command_namedName_parenthesizer___closed__9(); lean_mark_persistent(l_Lean_Parser_Command_namedName_parenthesizer___closed__9); +l_Lean_Parser_Command_namedName_parenthesizer___closed__10 = _init_l_Lean_Parser_Command_namedName_parenthesizer___closed__10(); +lean_mark_persistent(l_Lean_Parser_Command_namedName_parenthesizer___closed__10); l___regBuiltin_Lean_Parser_Command_namedName_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_namedName_parenthesizer___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_namedName_parenthesizer___closed__1); l___regBuiltin_Lean_Parser_Command_namedName_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_namedName_parenthesizer___closed__2(); @@ -19038,6 +19191,8 @@ l_Lean_Parser_Command_syntax___closed__20 = _init_l_Lean_Parser_Command_syntax__ lean_mark_persistent(l_Lean_Parser_Command_syntax___closed__20); l_Lean_Parser_Command_syntax___closed__21 = _init_l_Lean_Parser_Command_syntax___closed__21(); lean_mark_persistent(l_Lean_Parser_Command_syntax___closed__21); +l_Lean_Parser_Command_syntax___closed__22 = _init_l_Lean_Parser_Command_syntax___closed__22(); +lean_mark_persistent(l_Lean_Parser_Command_syntax___closed__22); l_Lean_Parser_Command_syntax = _init_l_Lean_Parser_Command_syntax(); lean_mark_persistent(l_Lean_Parser_Command_syntax); res = l___regBuiltin_Lean_Parser_Command_syntax(lean_io_mk_world()); @@ -19086,6 +19241,10 @@ l_Lean_Parser_Command_syntax_formatter___closed__12 = _init_l_Lean_Parser_Comman lean_mark_persistent(l_Lean_Parser_Command_syntax_formatter___closed__12); l_Lean_Parser_Command_syntax_formatter___closed__13 = _init_l_Lean_Parser_Command_syntax_formatter___closed__13(); lean_mark_persistent(l_Lean_Parser_Command_syntax_formatter___closed__13); +l_Lean_Parser_Command_syntax_formatter___closed__14 = _init_l_Lean_Parser_Command_syntax_formatter___closed__14(); +lean_mark_persistent(l_Lean_Parser_Command_syntax_formatter___closed__14); +l_Lean_Parser_Command_syntax_formatter___closed__15 = _init_l_Lean_Parser_Command_syntax_formatter___closed__15(); +lean_mark_persistent(l_Lean_Parser_Command_syntax_formatter___closed__15); l___regBuiltin_Lean_Parser_Command_syntax_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_syntax_formatter___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_syntax_formatter___closed__1); l___regBuiltin_Lean_Parser_Command_syntax_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_syntax_formatter___closed__2(); @@ -19123,6 +19282,8 @@ l_Lean_Parser_Command_syntax_parenthesizer___closed__14 = _init_l_Lean_Parser_Co lean_mark_persistent(l_Lean_Parser_Command_syntax_parenthesizer___closed__14); l_Lean_Parser_Command_syntax_parenthesizer___closed__15 = _init_l_Lean_Parser_Command_syntax_parenthesizer___closed__15(); lean_mark_persistent(l_Lean_Parser_Command_syntax_parenthesizer___closed__15); +l_Lean_Parser_Command_syntax_parenthesizer___closed__16 = _init_l_Lean_Parser_Command_syntax_parenthesizer___closed__16(); +lean_mark_persistent(l_Lean_Parser_Command_syntax_parenthesizer___closed__16); l___regBuiltin_Lean_Parser_Command_syntax_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_syntax_parenthesizer___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_syntax_parenthesizer___closed__1); l___regBuiltin_Lean_Parser_Command_syntax_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_syntax_parenthesizer___closed__2(); @@ -19625,6 +19786,8 @@ l_Lean_Parser_Command_macro_formatter___closed__12 = _init_l_Lean_Parser_Command lean_mark_persistent(l_Lean_Parser_Command_macro_formatter___closed__12); l_Lean_Parser_Command_macro_formatter___closed__13 = _init_l_Lean_Parser_Command_macro_formatter___closed__13(); lean_mark_persistent(l_Lean_Parser_Command_macro_formatter___closed__13); +l_Lean_Parser_Command_macro_formatter___closed__14 = _init_l_Lean_Parser_Command_macro_formatter___closed__14(); +lean_mark_persistent(l_Lean_Parser_Command_macro_formatter___closed__14); l___regBuiltin_Lean_Parser_Command_macro_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_macro_formatter___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_macro_formatter___closed__1); l___regBuiltin_Lean_Parser_Command_macro_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_macro_formatter___closed__2(); @@ -19709,6 +19872,8 @@ l_Lean_Parser_Command_macro_parenthesizer___closed__12 = _init_l_Lean_Parser_Com lean_mark_persistent(l_Lean_Parser_Command_macro_parenthesizer___closed__12); l_Lean_Parser_Command_macro_parenthesizer___closed__13 = _init_l_Lean_Parser_Command_macro_parenthesizer___closed__13(); lean_mark_persistent(l_Lean_Parser_Command_macro_parenthesizer___closed__13); +l_Lean_Parser_Command_macro_parenthesizer___closed__14 = _init_l_Lean_Parser_Command_macro_parenthesizer___closed__14(); +lean_mark_persistent(l_Lean_Parser_Command_macro_parenthesizer___closed__14); l___regBuiltin_Lean_Parser_Command_macro_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_macro_parenthesizer___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_macro_parenthesizer___closed__1); l___regBuiltin_Lean_Parser_Command_macro_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_macro_parenthesizer___closed__2(); @@ -19985,6 +20150,8 @@ l_Lean_Parser_Command_elab_formatter___closed__13 = _init_l_Lean_Parser_Command_ lean_mark_persistent(l_Lean_Parser_Command_elab_formatter___closed__13); l_Lean_Parser_Command_elab_formatter___closed__14 = _init_l_Lean_Parser_Command_elab_formatter___closed__14(); lean_mark_persistent(l_Lean_Parser_Command_elab_formatter___closed__14); +l_Lean_Parser_Command_elab_formatter___closed__15 = _init_l_Lean_Parser_Command_elab_formatter___closed__15(); +lean_mark_persistent(l_Lean_Parser_Command_elab_formatter___closed__15); l___regBuiltin_Lean_Parser_Command_elab_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_elab_formatter___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_elab_formatter___closed__1); l___regBuiltin_Lean_Parser_Command_elab_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_elab_formatter___closed__2(); @@ -20041,6 +20208,8 @@ l_Lean_Parser_Command_elab_parenthesizer___closed__13 = _init_l_Lean_Parser_Comm lean_mark_persistent(l_Lean_Parser_Command_elab_parenthesizer___closed__13); l_Lean_Parser_Command_elab_parenthesizer___closed__14 = _init_l_Lean_Parser_Command_elab_parenthesizer___closed__14(); lean_mark_persistent(l_Lean_Parser_Command_elab_parenthesizer___closed__14); +l_Lean_Parser_Command_elab_parenthesizer___closed__15 = _init_l_Lean_Parser_Command_elab_parenthesizer___closed__15(); +lean_mark_persistent(l_Lean_Parser_Command_elab_parenthesizer___closed__15); l___regBuiltin_Lean_Parser_Command_elab_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Command_elab_parenthesizer___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Command_elab_parenthesizer___closed__1); l___regBuiltin_Lean_Parser_Command_elab_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Command_elab_parenthesizer___closed__2(); diff --git a/stage0/stdlib/Lean/Parser/Tactic.c b/stage0/stdlib/Lean/Parser/Tactic.c index 9ef263faa38c..28c0913f23c4 100644 --- a/stage0/stdlib/Lean/Parser/Tactic.c +++ b/stage0/stdlib/Lean/Parser/Tactic.c @@ -2465,7 +2465,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_introMatch___closed__4() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("intro ", 6); +x_1 = lean_mk_string_from_bytes("intro", 5); return x_1; } } @@ -2576,7 +2576,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Tactic_introMatch_declRange { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_unsigned_to_nat(66u); -x_2 = lean_unsigned_to_nat(41u); +x_2 = lean_unsigned_to_nat(40u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -2590,7 +2590,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Parser_Tactic_introMatch_declRange___closed__1; x_2 = lean_unsigned_to_nat(25u); x_3 = l___regBuiltin_Lean_Parser_Tactic_introMatch_declRange___closed__2; -x_4 = lean_unsigned_to_nat(41u); +x_4 = lean_unsigned_to_nat(40u); 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/Parser/Term.c b/stage0/stdlib/Lean/Parser/Term.c index cec02d9c859f..1ebc15f867b5 100644 --- a/stage0/stdlib/Lean/Parser/Term.c +++ b/stage0/stdlib/Lean/Parser/Term.c @@ -37,7 +37,6 @@ static lean_object* l_Lean_Parser_Term_app_formatter___closed__2; static lean_object* l_Lean_Parser_Term_num___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_attrInstance_formatter___closed__2; static lean_object* l_Lean_Parser_Tactic_quotSeq___closed__4; -static lean_object* l_Lean_Parser_Term_structInstLVal___closed__16; static lean_object* l_Lean_Parser_Term_let__delayed_formatter___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Term_binop_parenthesizer___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Tactic_quot_parenthesizer(lean_object*); @@ -61,7 +60,6 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_dotIdent_parenthesizer( static lean_object* l___regBuiltin_Lean_Parser_Term_let_formatter___closed__2; static lean_object* l_Lean_Parser_Term_strictImplicitLeftBracket___closed__4; static lean_object* l_Lean_Parser_Term_suffices_formatter___closed__3; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__42; static lean_object* l_Lean_Parser_Term_proj___closed__4; static lean_object* l_Lean_Parser_Term_have_formatter___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Tactic_quot_declRange___closed__4; @@ -102,6 +100,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_implicitBinder_parenthesizer___boxed static lean_object* l_Lean_Parser_Term_defaultOrOfNonempty_parenthesizer___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_assert_declRange___closed__1; static lean_object* l_Lean_Parser_Term_strictImplicitRightBracket_formatter___closed__2; +static lean_object* l_Lean_Parser_Term_stateRefT_parenthesizer___closed__7; static lean_object* l_Lean_Parser_Term_subst___closed__5; static lean_object* l_Lean_Parser_Term_unop___closed__2; static lean_object* l_Lean_Parser_Term_typeAscription___closed__22; @@ -137,6 +136,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Tactic_tacticSeq1Indented_formatt static lean_object* l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Term___hyg_195____closed__10; LEAN_EXPORT lean_object* l_Lean_Parser_Term_macroArg_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_anonymousCtor___closed__2; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__23; 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_Term_structInst_parenthesizer___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Term_letDecl_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -162,6 +162,7 @@ static lean_object* l_Lean_Parser_Term_whereDecls_parenthesizer___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Term_unreachable_docString___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_quotedName_docString___closed__1; static lean_object* l_Lean_Parser_Term_type_parenthesizer___closed__10; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__9; static lean_object* l_Lean_Parser_Term_let__fun___closed__8; static lean_object* l_Lean_Parser_Tactic_quot___closed__5; static lean_object* l_Lean_Parser_Term_withDeclName___closed__2; @@ -178,6 +179,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_binrel_formatter___closed__1 LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_namedPattern(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_sort_formatter___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_byTactic_declRange___closed__2; +static lean_object* l_Lean_Parser_Term_sufficesDecl___closed__13; static lean_object* l_Lean_Parser_Term_leading__parser_parenthesizer___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_scientific_declRange___closed__2; static lean_object* l_Lean_Parser_Term_strictImplicitBinder___closed__3; @@ -201,6 +203,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_char_declRange___closed__7; static lean_object* l_Lean_Parser_Term_structInst___closed__25; static lean_object* l_Lean_Parser_Term_basicFun___closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_Term_doubleQuotedName_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Term_attributes_formatter___closed__8; static lean_object* l_Lean_Parser_Term_inaccessible_parenthesizer___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_tuple_declRange___closed__1; static lean_object* l_Lean_Parser_Term_structInstLVal_parenthesizer___closed__11; @@ -286,6 +289,7 @@ static lean_object* l_Lean_Parser_Term_namedPattern_formatter___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Term_withAnonymousAntiquot; static lean_object* l_Lean_Parser_Term_dynamicQuot_parenthesizer___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_declName_declRange___closed__5; +static lean_object* l_Lean_Parser_Term_clear_parenthesizer___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Term_let__fun_parenthesizer___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_structInstField_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_Term_generalizingParam_formatter___closed__9; @@ -379,6 +383,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_typeOf_declRange___closed__3 static lean_object* l_Lean_Parser_Term_explicit___closed__10; static lean_object* l___regBuiltin_Lean_Parser_Term_forInMacro_x27_declRange___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Term_strictImplicitLeftBracket; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__44; static lean_object* l___regBuiltin_Lean_Parser_Term_nomatch_declRange___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Term_hole_formatter___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_matchAltsWhereDecls_formatter___closed__1; @@ -396,12 +401,10 @@ static lean_object* l_Lean_Parser_Term_borrowed___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_quotedName_declRange___closed__6; static lean_object* l_Lean_Parser_Term_structInstLVal___closed__1; static lean_object* l_Lean_Parser_Term_letMVar_parenthesizer___closed__6; -static lean_object* l_Lean_Parser_Term_basicFun_formatter___closed__11; static lean_object* l_Lean_Parser_Term_structInstArrayRef___closed__13; static lean_object* l_Lean_Parser_Term_structInstArrayRef___closed__5; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_borrowed_formatter(lean_object*); static lean_object* l_Lean_Parser_Term_typeAscription_parenthesizer___closed__6; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__5; static lean_object* l_Lean_Parser_Term_matchAlt_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_proj_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Term_ensureTypeOf_formatter___closed__5; @@ -428,6 +431,7 @@ static lean_object* l_Lean_Parser_Term_arrow___closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_Term_optIdent_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_subst(lean_object*); static lean_object* l_Lean_Parser_Term_match___closed__5; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__43; static lean_object* l___regBuiltin_Lean_Parser_Term_leading__parser_declRange___closed__6; static lean_object* l_Lean_Parser_Term_suffices___closed__12; static lean_object* l_Lean_Parser_Term_doubleQuotedName_parenthesizer___closed__3; @@ -447,7 +451,6 @@ static lean_object* l_Lean_Parser_Term_noErrorIfUnused___closed__5; static lean_object* l_Lean_Parser_Term_unop_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Term_withDeclName_formatter___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Term_explicitUniv_docString___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__24; static lean_object* l_Lean_Parser_Term_noindex_formatter___closed__1; static lean_object* l_Lean_Parser_Term_dbgTrace_formatter___closed__3; static lean_object* l_Lean_Parser_Term_withAnonymousAntiquot___closed__10; @@ -500,7 +503,9 @@ static lean_object* l_Lean_Parser_Term_trueVal___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_byTactic_x27_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Term_binop__lazy___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Term_typeAscription_declRange___closed__4; +static lean_object* l_Lean_Parser_Term_typeAscription_formatter___closed__15; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_syntheticHole_parenthesizer(lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__37; static lean_object* l_Lean_Parser_Term_typeAscription_formatter___closed__2; static lean_object* l_Lean_Parser_Term_basicFun_parenthesizer___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_Term_show_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -537,6 +542,7 @@ static lean_object* l_Lean_Parser_Term_sorry___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_have_declRange___closed__4; static lean_object* l_Lean_Parser_Term_structInst_formatter___closed__7; static lean_object* l_Lean_Parser_Term_binop_formatter___closed__4; +static lean_object* l_Lean_Parser_Term_ellipsis___closed__11; LEAN_EXPORT lean_object* l_Lean_Parser_Term_binop_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_strictImplicitRightBracket_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Term_dotIdent___closed__8; @@ -566,6 +572,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_binrel_declRange___closed__7 static lean_object* l_Lean_Parser_Term_structInstArrayRef_formatter___closed__6; static lean_object* l_Lean_Parser_Term_letIdLhs_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Tactic_tacticSeqBracketed___closed__3; +static lean_object* l_Lean_Parser_Term_match_formatter___closed__14; static lean_object* l___regBuiltin_Lean_Parser_Term_quotedName_declRange___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Term_typeOf_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_assert___closed__1; @@ -598,6 +605,7 @@ lean_object* l_Lean_Parser_many_parenthesizer(lean_object*, lean_object*, lean_o static lean_object* l_Lean_Parser_Term_letrec___closed__7; static lean_object* l_Lean_Parser_Command_docComment_formatter___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Term_num_declRange___closed__5; +static lean_object* l_Lean_Parser_Term_anonymousCtor_parenthesizer___closed__9; static lean_object* l_Lean_Parser_Term_optExprPrecedence_formatter___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_tacticSeqBracketed_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_let__fun___closed__14; @@ -643,7 +651,6 @@ static lean_object* l_Lean_Parser_Command_commentBody___closed__3; static lean_object* l_Lean_Parser_Term_prop_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Term_structInstArrayRef_formatter___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_waitIfContainsMVar_formatter___closed__2; -static lean_object* l_Lean_Parser_Term_explicitBinder_formatter___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_Term_basicFun_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_noErrorIfUnused_formatter___closed__2; static lean_object* l_Lean_Parser_Term_letPatDecl_parenthesizer___closed__4; @@ -664,6 +671,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_leading__parser_parenthesizer(lean_o static lean_object* l___regBuiltin_Lean_Parser_Term_nomatch_declRange___closed__5; static lean_object* l_Lean_Parser_Term_matchAltsWhereDecls___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Term_nomatch_formatter___closed__1; +static lean_object* l_Lean_Parser_Term_binrel_parenthesizer___closed__8; static lean_object* l_Lean_Parser_Term_letIdBinder___closed__4; static lean_object* l_Lean_Parser_Term_binrel__no__prop_parenthesizer___closed__3; LEAN_EXPORT uint8_t l_Lean_Parser_Term_isIdent(lean_object*); @@ -687,7 +695,6 @@ static lean_object* l_Lean_Parser_Term_letrec_formatter___closed__3; static lean_object* l_Lean_Parser_Term_dbgTrace___closed__8; LEAN_EXPORT lean_object* l_Lean_Parser_Term_attributes; static lean_object* l_Lean_Parser_Tactic_tacticSeq1Indented___closed__6; -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007_(lean_object*); static lean_object* l_Lean_Parser_Term_strictImplicitLeftBracket_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Term_argument_parenthesizer___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_letEqnsDecl_formatter(lean_object*); @@ -720,7 +727,6 @@ static lean_object* l_Lean_Parser_Term_paren___closed__7; static lean_object* l_Lean_Parser_Term_fromTerm___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Term_syntheticHole_formatter___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_letEqnsDecl_parenthesizer(lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__25; static lean_object* l_Lean_Parser_Term_forall_parenthesizer___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Tactic_tacticSeqBracketed_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_attrKind___closed__3; @@ -790,7 +796,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_completion_declRange___close static lean_object* l_Lean_Parser_Term_defaultOrOfNonempty___closed__10; static lean_object* l_Lean_Parser_Term_matchAlts___closed__3; static lean_object* l_Lean_Parser_Term_scoped___closed__6; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__30; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__36; static lean_object* l_Lean_Parser_Term_instBinder_formatter___closed__1; static lean_object* l_Lean_Parser_Term_nomatch_formatter___closed__2; lean_object* l_Lean_PrettyPrinter_Parenthesizer_fieldIdx_parenthesizer___boxed(lean_object*); @@ -807,6 +813,7 @@ static lean_object* l_Lean_Parser_Term_motive_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Term_typeAscription___closed__17; static lean_object* l_Lean_Parser_Term_stateRefT_formatter___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Term_letIdLhs; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__28; static lean_object* l_Lean_Parser_Term_motive___closed__5; static lean_object* l_Lean_Parser_Term_tuple_parenthesizer___closed__8; static lean_object* l_Lean_Parser_Term_funStrictImplicitBinder_formatter___closed__2; @@ -816,8 +823,10 @@ lean_object* l_Lean_Parser_nameLit_formatter(lean_object*, lean_object*, lean_ob lean_object* l_Lean_Parser_nonReservedSymbol(lean_object*, uint8_t); static lean_object* l___regBuiltin_Lean_Parser_Term_waitIfContainsMVar_declRange___closed__3; static lean_object* l_Lean_Parser_Term_structInstLVal___closed__10; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__6; static lean_object* l_Lean_Parser_Term_structInstFieldAbbrev_parenthesizer___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Term_nomatch_declRange___closed__4; +lean_object* l_Lean_Parser_hygieneInfo_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_tuple___closed__11; static lean_object* l___regBuiltin_Lean_Parser_Term_noImplicitLambda_declRange___closed__4; static lean_object* l_Lean_Parser_Term_namedArgument_formatter___closed__6; @@ -834,7 +843,6 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_defaultOrOfNonempty_formatter(lean_o LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_ensureExpectedType_declRange(lean_object*); static lean_object* l_Lean_Parser_Term_structInst___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Term_have_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__45; 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_ensureExpectedType_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_match_parenthesizer___closed__8; @@ -847,6 +855,7 @@ static lean_object* l_Lean_Parser_Term_tuple___closed__14; static lean_object* l_Lean_Parser_Term_scoped_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Term_waitIfTypeMVar_parenthesizer___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_arrow_declRange___closed__4; +static lean_object* l_Lean_Parser_Term_sufficesDecl___closed__12; static lean_object* l_Lean_Parser_Term_letIdLhs_parenthesizer___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_typeOf_declRange___closed__4; static lean_object* l_Lean_Parser_Term_panic_formatter___closed__4; @@ -855,6 +864,7 @@ static lean_object* l_Lean_Parser_Term_anonymousCtor___closed__13; static lean_object* l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__8; extern lean_object* l_Lean_Parser_strLit; static lean_object* l_Lean_Parser_Term_waitIfTypeContainsMVar___closed__6; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__3; static lean_object* l_Lean_Parser_Term_tuple___closed__6; static lean_object* l_Lean_Parser_Term_haveIdDecl___closed__2; static lean_object* l_Lean_Parser_Term_paren_formatter___closed__5; @@ -874,7 +884,6 @@ static lean_object* l_Lean_Parser_Term_letMVar_parenthesizer___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_byTactic_declRange___closed__7; static lean_object* l_Lean_Parser_Term_completion___closed__2; static lean_object* l_Lean_Parser_Term_structInst_formatter___closed__2; -static lean_object* l_Lean_Parser_Term_unop___closed__11; static lean_object* l_Lean_Parser_Term_dynamicQuot_parenthesizer___closed__7; static lean_object* l_Lean_Parser_Term_generalizingParam_parenthesizer___closed__8; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_typeSpec_parenthesizer(lean_object*); @@ -898,6 +907,7 @@ lean_object* l_Lean_Parser_attrParser_formatter___rarg(lean_object*, lean_object static lean_object* l___regBuiltin_Lean_Parser_Term_binop_declRange___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_darrow_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_assert_parenthesizer___closed__1; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__18; static lean_object* l_Lean_Parser_Term_forInMacro_x27___closed__7; static lean_object* l_Lean_Parser_Term_suffices___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Term_haveIdDecl_parenthesizer___closed__1; @@ -951,6 +961,7 @@ static lean_object* l_Lean_Parser_Term_binderTactic_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Term_instBinder___closed__3; static lean_object* l_Lean_Parser_Term_structInstField_formatter___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Term_let__tmp_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__10; static lean_object* l_Lean_Parser_Term_fun_parenthesizer___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Term_hole_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_tacticSeq1Indented_formatter___closed__3; @@ -972,6 +983,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_haveEqnsDecl_parenthesizer__ lean_object* l_Lean_Parser_orelse(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_unop___closed__8; static lean_object* l___regBuiltin_Lean_Parser_Term_assert_declRange___closed__6; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__34; static lean_object* l_Lean_Parser_Term_assert___closed__3; static lean_object* l_Lean_Parser_Term_dbgTrace___closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_Term_doubleQuotedName_formatter___closed__3___boxed__const__1; @@ -1000,6 +1012,7 @@ static lean_object* l_Lean_Parser_Term_letPatDecl_parenthesizer___closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_Term_unreachable; static lean_object* l_Lean_Parser_Term_inaccessible_formatter___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Term_binop_declRange___closed__1; +extern lean_object* l_Lean_Parser_hygieneInfo; static lean_object* l___regBuiltin_Lean_Parser_Term_inaccessible_formatter___closed__1; static lean_object* l_Lean_Parser_Term_forInMacro_formatter___closed__5; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_ensureTypeOf_parenthesizer(lean_object*); @@ -1042,6 +1055,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_assert_parenthesizer(le static lean_object* l_Lean_Parser_Term_let_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Tactic_tacticSeqBracketed_formatter___closed__9; static lean_object* l_Lean_Parser_Term_suffices___closed__9; +static lean_object* l_Lean_Parser_Term_sufficesDecl_formatter___closed__11; static lean_object* l_Lean_Parser_Term_nomatch___closed__3; static lean_object* l_Lean_Parser_Term_noImplicitLambda___closed__4; static lean_object* l_Lean_Parser_Term_whereDecls___closed__2; @@ -1053,6 +1067,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_haveDecl_parenthesizer___clo static lean_object* l___regBuiltin_Lean_Parser_Term_dotIdent_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_scoped_formatter___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_proj_declRange___closed__7; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__1; static lean_object* l_Lean_Parser_Term_structInstField___closed__10; static lean_object* l___regBuiltin_Lean_Parser_Term_panic_parenthesizer___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_letrec_declRange___closed__4; @@ -1082,11 +1097,13 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_app_declRange___closed__6; static lean_object* l_Lean_Parser_Term_prop___closed__5; static lean_object* l_Lean_Parser_Term_let__fun_formatter___closed__2; static lean_object* l_Lean_Parser_Term_suffices_parenthesizer___closed__3; +static lean_object* l_Lean_Parser_Term_instBinder_formatter___closed__9; static lean_object* l_Lean_Parser_Term_matchAlt_formatter___closed__6; static lean_object* l_Lean_Parser_Term_structInst___closed__27; static lean_object* l___regBuiltin_Lean_Parser_Term_namedPattern_docString___closed__1; static lean_object* l_Lean_Parser_Term_matchAlt_formatter___closed__2; static lean_object* l_Lean_Parser_Term_binop__lazy_formatter___closed__1; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__35; LEAN_EXPORT lean_object* l_Lean_Parser_Term_show_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_typeAscription___closed__16; LEAN_EXPORT lean_object* l_Lean_Parser_Term_waitIfTypeMVar_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1155,8 +1172,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_ensureTypeOf_formatter___clo static lean_object* l___regBuiltin_Lean_Parser_Term_ident_declRange___closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_Term_typeOf; static lean_object* l_Lean_Parser_semicolonOrLinebreak_formatter___closed__4; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__3; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_Term_binderType_parenthesizer(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Term_haveIdDecl; static lean_object* l_Lean_Parser_Term_letMVar_formatter___closed__4; @@ -1182,6 +1197,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_haveEqnsDecl; static lean_object* l_Lean_Parser_Tactic_quotSeq_formatter___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Term_pipeProj_declRange___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_char_declRange___closed__4; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__16; static lean_object* l_Lean_Parser_Tactic_tacticSeq1Indented_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Term_inaccessible_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Term_namedArgument___closed__9; @@ -1191,6 +1207,7 @@ static lean_object* l_Lean_Parser_Term_letPatDecl_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Term_structInstFieldAbbrev_parenthesizer___closed__9; static lean_object* l_Lean_Parser_Term_structInst_formatter___closed__18; static lean_object* l_Lean_Parser_Term_optEllipsis___closed__3; +static lean_object* l_Lean_Parser_Term_typeAscription___closed__23; static lean_object* l_Lean_Parser_Term_structInstField___closed__8; static lean_object* l_Lean_Parser_Term_attrKind_formatter___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_typeOf_formatter(lean_object*); @@ -1217,7 +1234,6 @@ static lean_object* l_Lean_Parser_Term_binop__lazy___closed__2; static lean_object* l_Lean_Parser_Term_match___closed__9; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_tuple_declRange(lean_object*); static lean_object* l_Lean_Parser_Term_dynamicQuot_parenthesizer___closed__6; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__37; LEAN_EXPORT lean_object* l_Lean_Parser_Term_strictImplicitLeftBracket_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_binrel_formatter___closed__6; static lean_object* l_Lean_Parser_Term_structInstFieldAbbrev_parenthesizer___closed__2; @@ -1266,6 +1282,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_doubleQuotedName_declRa static lean_object* l_Lean_Parser_Term_match_formatter___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Term_leading__parser_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Term_letPatDecl_parenthesizer___closed__6; +static lean_object* l_Lean_Parser_Term_clear_formatter___closed__7; static lean_object* l_Lean_Parser_Term_cdot_formatter___closed__4; static lean_object* l_Lean_Parser_Term_dbgTrace___closed__11; static lean_object* l_Lean_Parser_Term_namedArgument___closed__3; @@ -1318,7 +1335,6 @@ static lean_object* l_Lean_Parser_Term_app_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_subst___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Term_macroDollarArg; LEAN_EXPORT lean_object* l_Lean_Parser_Term_argument_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__31; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_let__fun_formatter(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Term_funBinder; static lean_object* l_Lean_Parser_Term_paren___closed__10; @@ -1374,12 +1390,12 @@ static lean_object* l_Lean_Parser_Term_paren___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Term_nomatch_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_sort_declRange(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_attrInstance_formatter___closed__1; +static lean_object* l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__9; static lean_object* l_Lean_Parser_Term_byTactic___closed__2; static lean_object* l_Lean_Parser_Term_sort_formatter___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_namedPattern_formatter(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_forInMacro_parenthesizer___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Tactic_tacticSeqBracketed_formatter___closed__3; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__28; static lean_object* l___regBuiltin_Lean_Parser_Term_suffices_declRange___closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_Term_bracketedBinder_parenthesizer(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_structInst___closed__5; @@ -1412,6 +1428,7 @@ static lean_object* l_Lean_Parser_Term_defaultOrOfNonempty_formatter___closed__5 static lean_object* l_Lean_Parser_Term_attrKind_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_paren_parenthesizer___closed__5; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_fromTerm_parenthesizer(lean_object*); +static lean_object* l_Lean_Parser_Term_withAnonymousAntiquot___closed__12; static lean_object* l___regBuiltin_Lean_Parser_Term_match_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_match___closed__3; static lean_object* l_Lean_Parser_Term_attrKind_parenthesizer___closed__4; @@ -1422,6 +1439,7 @@ static lean_object* l_Lean_Parser_Term_byTactic___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_tacticSeq1Indented_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_explicit___closed__6; static lean_object* l_Lean_Parser_Term_binderTactic___closed__4; +static lean_object* l_Lean_Parser_Term_haveDecl___closed__9; static lean_object* l_Lean_Parser_Term_match_parenthesizer___closed__7; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_match_docString(lean_object*); static lean_object* l_Lean_Parser_Term_explicitUniv_formatter___closed__5; @@ -1438,8 +1456,11 @@ static lean_object* l_Lean_Parser_Command_commentBody___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_sufficesDecl_formatter(lean_object*); static lean_object* l_Lean_Parser_Term_str_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Term_panic___closed__8; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__12; static lean_object* l_Lean_Parser_Term_have___closed__9; LEAN_EXPORT lean_object* l_Lean_Parser_Term_typeAscription; +static lean_object* l_Lean_Parser_Term_binrel___closed__13; +static lean_object* l_Lean_Parser_Term_whereDecls_parenthesizer___closed__8; LEAN_EXPORT lean_object* l_Lean_Parser_Term_proj; static lean_object* l_Lean_Parser_Term_declName___closed__3; static lean_object* l_Lean_Parser_Term_trueVal___closed__1; @@ -1455,6 +1476,7 @@ static lean_object* l_Lean_Parser_Term_haveIdDecl___closed__6; static lean_object* l_Lean_Parser_Term_forInMacro_x27___closed__8; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_attributes_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_Term_scoped___closed__7; +static lean_object* l_Lean_Parser_Term_match___closed__19; static lean_object* l___regBuiltin_Lean_Parser_Term_app_formatter___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_whereDecls_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_byTactic_formatter___closed__1; @@ -1472,12 +1494,12 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_paren_declRange___closed__3; static lean_object* l_Lean_Parser_Term_binrel_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Term_structInst___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_pipeProj_formatter___closed__2; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__15; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_quotedName_docString(lean_object*); static lean_object* l_Lean_Parser_Term_namedPattern_formatter___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Term_doubleQuotedName_declRange___closed__3; static lean_object* l_Lean_Parser_Term_matchAltsWhereDecls_formatter___closed__6; static lean_object* l_Lean_Parser_Term_binop___closed__3; +static lean_object* l_Lean_Parser_Term_withAnonymousAntiquot___closed__13; static lean_object* l_Lean_Parser_Term_dbgTrace_parenthesizer___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_dbgTrace_declRange___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Tactic_quotSeq_declRange___closed__3; @@ -1487,6 +1509,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_trailing__parser_declRange__ static lean_object* l_Lean_Parser_Term_waitIfTypeContainsMVar___closed__5; static lean_object* l_Lean_Parser_Term_clear___closed__6; static lean_object* l_Lean_Parser_Term_fromTerm___closed__1; +static lean_object* l_Lean_Parser_Term_sufficesDecl___closed__11; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_structInstLVal_formatter(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_doubleQuotedName_declRange___closed__2; static lean_object* l_Lean_Parser_Term_attrKind_formatter___closed__1; @@ -1526,8 +1549,8 @@ static lean_object* l_Lean_Parser_Term_binrel__no__prop___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_syntheticHole_declRange___closed__1; static lean_object* l_Lean_Parser_Term_letIdLhs___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Term_letRecDecl_parenthesizer___closed__2; +static lean_object* l_Lean_Parser_Term_attributes___closed__14; static lean_object* l_Lean_Parser_Term_type_formatter___closed__10; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__13; static lean_object* l_Lean_Parser_Term_forall___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_Command_docComment_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_implicitBinder___closed__3; @@ -1540,6 +1563,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_sorry_declRange___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_doubleQuotedName_formatter___closed__1; static lean_object* l_Lean_Parser_Term_structInstFieldAbbrev___closed__3; lean_object* l_Lean_Parser_strLit_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Term_sufficesDecl_formatter___closed__13; static lean_object* l___regBuiltin_Lean_Parser_Term_explicitUniv_declRange___closed__2; static lean_object* l_Lean_Parser_Term_funStrictImplicitBinder_formatter___closed__3; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_noErrorIfUnused_formatter(lean_object*); @@ -1569,6 +1593,7 @@ static lean_object* l_Lean_Parser_Term_tuple___closed__2; static lean_object* l_Lean_Parser_Term_unreachable___closed__8; static lean_object* l_Lean_Parser_Term_borrowed_formatter___closed__2; static lean_object* l_Lean_Parser_Term_pipeProj_parenthesizer___closed__3; +static lean_object* l_Lean_Parser_Term_haveDecl_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Term_tuple___closed__15; static lean_object* l_Lean_Parser_Term_ensureTypeOf_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Term_binrel___closed__9; @@ -1628,14 +1653,17 @@ static lean_object* l_Lean_Parser_Term_generalizingParam___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Term_forInMacro_declRange___closed__3; static lean_object* l_Lean_Parser_Term_binderTactic_formatter___closed__5; static lean_object* l_Lean_Parser_Term_binrel__no__prop___closed__8; +static lean_object* l_Lean_Parser_Term_optIdent___closed__5; static lean_object* l_Lean_Parser_Term_prop___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Term_explicitUniv_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Term_trailing__parser___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_sepBy1IndentSemicolon_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_5098____closed__29; static lean_object* l_Lean_Parser_Term_explicitBinder_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Term_let__fun___closed__1; static lean_object* l_Lean_Parser_Term_typeAscription_parenthesizer___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Term_unop_formatter___closed__2; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__11; static lean_object* l___regBuiltin_Lean_Parser_Term_pipeCompletion_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_letIdLhs_formatter___closed__7; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_byTactic(lean_object*); @@ -1661,7 +1689,6 @@ static lean_object* l_Lean_Parser_Term_funImplicitBinder_parenthesizer___closed_ static lean_object* l___regBuiltin_Lean_Parser_Term_motive_formatter___closed__2; static lean_object* l_Lean_Parser_Term_forall_formatter___closed__12; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_basicFun_parenthesizer(lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__6; static lean_object* l_Lean_Parser_Term_attributes___closed__10; static lean_object* l_Lean_Parser_Term_generalizingParam_formatter___closed__6; static lean_object* l_Lean_Parser_Term_letMVar___closed__9; @@ -1684,6 +1711,7 @@ static lean_object* l_Lean_Parser_Term_defaultOrOfNonempty_formatter___closed__1 static lean_object* l___regBuiltin_Lean_Parser_Term_show_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Term_type___closed__14; static lean_object* l___regBuiltin_Lean_Parser_Term_haveEqnsDecl_parenthesizer___closed__2; +static lean_object* l_Lean_Parser_Term_anonymousCtor_parenthesizer___closed__8; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_noErrorIfUnused_docString(lean_object*); lean_object* l_Lean_Parser_ident_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Term_fun_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1714,13 +1742,10 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_pipeCompletion_parenthesizer(lean_ob static lean_object* l_Lean_Parser_Term_noImplicitLambda_formatter___closed__3; static lean_object* l_Lean_Parser_semicolonOrLinebreak_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_structInst_formatter___closed__15; -static lean_object* l_Lean_Parser_Term_leading__parser___closed__12; static lean_object* l_Lean_Parser_Term_letrec_parenthesizer___closed__7; static lean_object* l_Lean_Parser_Term_waitIfContainsMVar_parenthesizer___closed__3; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__35; static lean_object* l___regBuiltin_Lean_Parser_Term_subst_declRange___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Term_tuple_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__39; static lean_object* l___regBuiltin_Lean_Parser_Term_attributes_parenthesizer___closed__1; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_94____closed__7; static lean_object* l_Lean_Parser_Term_macroDollarArg_formatter___closed__3; @@ -1749,7 +1774,6 @@ static lean_object* l_Lean_Parser_Term_haveIdLhs_parenthesizer___closed__2; lean_object* l_Lean_Parser_many1_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_structInst_parenthesizer___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Term_letIdLhs_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__41; LEAN_EXPORT lean_object* l_Lean_Parser_Term_completion_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_paren_parenthesizer___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Term_showRhs_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1784,6 +1808,7 @@ static lean_object* l_Lean_Parser_semicolonOrLinebreak___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_defaultOrOfNonempty_declRange___closed__4; static lean_object* l_Lean_Parser_Term_waitIfTypeMVar_parenthesizer___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_binderTactic_parenthesizer(lean_object*); +static lean_object* l_Lean_Parser_Term_withAnonymousAntiquot_formatter___closed__9; static lean_object* l_Lean_Parser_Term_explicitBinder_formatter___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Term_waitIfTypeContainsMVar_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_namedPattern_declRange(lean_object*); @@ -1801,6 +1826,7 @@ static lean_object* l_Lean_Parser_Term_doubleQuotedName___closed__5; static lean_object* l_Lean_Parser_Term_structInst___closed__16; static lean_object* l_Lean_Parser_Term_typeAscription___closed__1; static lean_object* l_Lean_Parser_Term_binop_formatter___closed__3; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__40; static lean_object* l_Lean_Parser_Term_sort___closed__3; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_have_formatter(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Term_char_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1823,7 +1849,6 @@ static lean_object* l_Lean_Parser_Term_letRecDecls_parenthesizer___closed__1; lean_object* l_Lean_ppAllowUngrouped_formatter___boxed(lean_object*); static lean_object* l_Lean_Parser_Term_namedPattern___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Term_letEqnsDecl; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__22; static lean_object* l_Lean_Parser_Term_dbgTrace_parenthesizer___closed__7; lean_object* l_Lean_Parser_adaptCacheableContext(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_char___closed__1; @@ -1868,7 +1893,6 @@ static lean_object* l_Lean_Parser_Term_argument___closed__6; static lean_object* l_Lean_Parser_Term_funImplicitBinder_formatter___closed__1; static lean_object* l_Lean_Parser_Term_inaccessible_formatter___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Term_implicitBinder(uint8_t); -static lean_object* l_Lean_Parser_Term_basicFun___closed__15; static lean_object* l_Lean_Parser_Term_structInstArrayRef_formatter___closed__2; static lean_object* l_Lean_Parser_Term_whereDecls___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_withDeclName_formatter___closed__1; @@ -1902,7 +1926,6 @@ static lean_object* l_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___clo static lean_object* l___regBuiltin_Lean_Parser_Term_subst_docString___closed__1; static lean_object* l_Lean_Parser_Term_letEqnsDecl___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Tactic_tacticSeqIndentGt_formatter___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__32; static lean_object* l_Lean_Parser_Term_prop___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_syntheticHole_declRange___closed__4; static lean_object* l_Lean_Parser_Tactic_tacticSeqBracketed_parenthesizer___closed__4; @@ -1916,7 +1939,6 @@ static lean_object* l_Lean_Parser_Term_funBinder_formatter___closed__2; static lean_object* l_Lean_Parser_Tactic_quot_formatter___closed__2; static lean_object* l_Lean_Parser_Term_letIdDecl_formatter___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_suffices_formatter___closed__2; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_have_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_structInst_formatter___closed__9; static lean_object* l___regBuiltin_Lean_Parser_Term_arrow_formatter___closed__2; @@ -1947,6 +1969,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Tactic_quotSeq_declRange___closed static lean_object* l_Lean_Parser_Tactic_tacticSeqBracketed_parenthesizer___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Term_sorry_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Term_ensureExpectedType_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___closed__8; static lean_object* l_Lean_Parser_Term_hole_formatter___closed__3; static lean_object* l_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___closed__7; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_arrow_formatter(lean_object*); @@ -2029,6 +2052,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_binderDefault_formatter static lean_object* l_Lean_Parser_Term_structInstFieldAbbrev_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Term_show_formatter___closed__3; static lean_object* l_Lean_Parser_Term_letEqnsDecl___closed__2; +static lean_object* l_Lean_Parser_Term_haveDecl_formatter___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Term_structInst_declRange___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Term_haveEqnsDecl_formatter___closed__2; static lean_object* l_Lean_Parser_Term_generalizingParam_parenthesizer___closed__1; @@ -2058,8 +2082,10 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_withDeclName_parenthesizer__ static lean_object* l_Lean_Parser_Term_motive_formatter___closed__5; lean_object* l_Lean_Parser_trailingNode(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_leading__parser___closed__7; +static lean_object* l_Lean_Parser_Term_anonymousCtor___closed__15; static lean_object* l_Lean_Parser_Term_doubleQuotedName_parenthesizer___closed__6; static lean_object* l_Lean_Parser_Term_completion___closed__1; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__2; static lean_object* l_Lean_Parser_Term_let__delayed___closed__3; static lean_object* l_Lean_Parser_Term_assert___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_char_declRange___closed__2; @@ -2118,7 +2144,6 @@ static lean_object* l_Lean_Parser_Term_typeAscription_formatter___closed__6; static lean_object* l_Lean_Parser_Term_argument___closed__1; static lean_object* l_Lean_Parser_Term_sufficesDecl_formatter___closed__7; static lean_object* l_Lean_Parser_Tactic_tacticSeqBracketed___closed__11; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__43; static lean_object* l_Lean_Parser_Term_letIdLhs___closed__2; static lean_object* l_Lean_Parser_Term_structInstArrayRef___closed__4; static lean_object* l_Lean_Parser_Term_let__fun___closed__2; @@ -2132,6 +2157,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_str_declRange___closed__6; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_94____closed__8; static lean_object* l_Lean_Parser_Term_letIdLhs_formatter___closed__3; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_suffices_formatter(lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__21; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_binrel_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_Tactic_quot_parenthesizer___closed__7; static lean_object* l_Lean_Parser_Term_optEllipsis___closed__1; @@ -2149,6 +2175,7 @@ static lean_object* l_Lean_Parser_Term_unreachable_parenthesizer___closed__1; lean_object* l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_sort_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Term_waitIfTypeContainsMVar_parenthesizer___closed__1; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__13; static lean_object* l___regBuiltin_Lean_Parser_Term_hole_declRange___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_show_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_instBinder___closed__7; @@ -2158,7 +2185,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_borrowed_parenthesizer___clo LEAN_EXPORT lean_object* l_Lean_Parser_Term_letMVar_formatter(lean_object*, lean_object*, lean_object*, 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_Term_typeAscription___closed__15; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__33; static lean_object* l___regBuiltin_Lean_Parser_Term_forInMacro_x27_formatter___closed__2; static lean_object* l_Lean_Parser_Term_subst___closed__7; static lean_object* l_Lean_Parser_Term_binop__lazy___closed__5; @@ -2186,13 +2212,14 @@ static lean_object* l_Lean_Parser_Term_letPatDecl___closed__7; static lean_object* l_Lean_Parser_Term_trailing__parser_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Term_whereDecls___closed__10; static lean_object* l___regBuiltin_Lean_Parser_Term_anonymousCtor_docString___closed__1; +static lean_object* l_Lean_Parser_Term_ellipsis_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Term_haveEqnsDecl___closed__3; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__44; static lean_object* l___regBuiltin_Lean_Parser_Term_paren_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Term_forall_formatter___closed__10; static lean_object* l___regBuiltin_Lean_Parser_Term_let__tmp_docString___closed__1; static lean_object* l_Lean_Parser_Term_typeOf___closed__7; static lean_object* l_Lean_Parser_Term_namedPattern_parenthesizer___closed__5; +static lean_object* l_Lean_Parser_Term_typeAscription_parenthesizer___closed__15; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_paren_declRange(lean_object*); static lean_object* l_Lean_Parser_Term_arrow___closed__5; static lean_object* l_Lean_Parser_Term_let_formatter___closed__5; @@ -2220,7 +2247,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_forInMacro_x27_declRange___c static lean_object* l_Lean_Parser_Term_byTactic_formatter___closed__6; static lean_object* l_Lean_Parser_Term_letEqnsDecl_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Term_defaultOrOfNonempty___closed__12; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__38; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_pipeProj_formatter(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Term_funBinder_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_tuple_parenthesizer___closed__2; @@ -2234,6 +2260,7 @@ static lean_object* l_Lean_Parser_Term_waitIfContainsMVar_formatter___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Term_binop_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_motive_formatter___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_explicit_declRange___closed__3; +static lean_object* l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__11; static lean_object* l___regBuiltin_Lean_Parser_Term_binop__lazy_declRange___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Term_unreachable_formatter___closed__2; static lean_object* l_Lean_Parser_Term_explicitBinder_formatter___closed__4; @@ -2354,7 +2381,6 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_dynamicQuot(lean_object static lean_object* l_Lean_Parser_Term_matchDiscr___closed__1; static lean_object* l_Lean_Parser_Term_instBinder___closed__2; static lean_object* l_Lean_Parser_Term_fun___closed__3; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__26; static lean_object* l_Lean_Parser_Term_haveDecl_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Term_haveIdLhs_formatter___closed__2; static lean_object* l_Lean_Parser_Term_unop_formatter___closed__2; @@ -2382,6 +2408,7 @@ static lean_object* l_Lean_Parser_Term_have_formatter___closed__3; static lean_object* l_Lean_Parser_Term_waitIfTypeContainsMVar___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_forall_parenthesizer(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_depArrow_declRange___closed__2; +static lean_object* l_Lean_Parser_Term_typeAscription_formatter___closed__14; static lean_object* l___regBuiltin_Lean_Parser_Term_scientific_declRange___closed__7; static lean_object* l_Lean_Parser_Term_syntheticHole_formatter___closed__2; static lean_object* l_Lean_Parser_Term_fromTerm_formatter___closed__3; @@ -2395,6 +2422,7 @@ static lean_object* l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Term_haveIdDecl___closed__5; static lean_object* l_Lean_Parser_Tactic_tacticSeqBracketed_formatter___closed__5; static lean_object* l_Lean_Parser_Term_unop_formatter___closed__3; +static lean_object* l_Lean_Parser_Term_whereDecls___closed__11; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_panic_formatter(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_match_declRange___closed__7; static lean_object* l_Lean_Parser_Term_panic_formatter___closed__2; @@ -2565,7 +2593,6 @@ static lean_object* l_Lean_Parser_Term_explicit_formatter___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Term_unop_declRange___closed__2; static lean_object* l_Lean_Parser_Term_generalizingParam_formatter___closed__5; static lean_object* l_Lean_Parser_Term_ellipsis_formatter___closed__2; -static lean_object* l_Lean_Parser_Term_withDeclName_formatter___closed__7; static lean_object* l_Lean_Parser_Term_binrel__no__prop___closed__1; static lean_object* l_Lean_Parser_Term_haveDecl___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Term_ident_declRange___closed__6; @@ -2656,9 +2683,11 @@ static lean_object* l_Lean_Parser_Term_match_parenthesizer___closed__12; static lean_object* l_Lean_Parser_Term_local___closed__2; static lean_object* l_Lean_Parser_Term_waitIfContainsMVar_formatter___closed__4; static lean_object* l_Lean_Parser_Term_let__fun_parenthesizer___closed__1; +static lean_object* l_Lean_Parser_Term_whereDecls_formatter___closed__11; static lean_object* l___regBuiltin_Lean_Parser_Term_noindex_declRange___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_noImplicitLambda_parenthesizer___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Term_letMVar_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__13; static lean_object* l___regBuiltin_Lean_Parser_Term_local_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_trueVal_formatter___closed__2; static lean_object* l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Term___hyg_195____closed__5; @@ -2699,6 +2728,7 @@ static lean_object* l_Lean_Parser_Term_letRecDecl_formatter___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Term_explicitUniv_formatter___closed__2; static lean_object* l_Lean_Parser_Term_macroLastArg___closed__1; static lean_object* l_Lean_Parser_Term_generalizingParam___closed__6; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__32; static lean_object* l_Lean_Parser_Term_dotIdent_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Term_match_parenthesizer___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Term_app_parenthesizer___closed__1; @@ -2736,6 +2766,7 @@ static lean_object* l_Lean_Parser_Tactic_sepByIndentSemicolon___closed__8; static lean_object* l_Lean_Parser_Term_byTactic___closed__7; static lean_object* l_Lean_Parser_Term_haveDecl___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_leading__parser_formatter___closed__1; +static lean_object* l_Lean_Parser_Term_ellipsis_formatter___closed__5; static lean_object* l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Term___hyg_195____closed__8; static lean_object* l_Lean_Parser_Term_generalizingParam_parenthesizer___closed__10; static lean_object* l_Lean_Parser_Term_byTactic_x27___closed__3; @@ -2753,8 +2784,8 @@ static lean_object* l_Lean_Parser_Term_letMVar___closed__8; static lean_object* l_Lean_Parser_Term_letRecDecl___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_binrel__no__prop_formatter___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_generalizingParam_formatter___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__27; static lean_object* l_Lean_Parser_Term_doubleQuotedName_formatter___closed__6; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__15; lean_object* l_Lean_Parser_charLit_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_cdot_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Term_byTactic_formatter___closed__4; @@ -2775,11 +2806,14 @@ static lean_object* l_Lean_Parser_Term_letDecl_parenthesizer___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Term_funImplicitBinder; static lean_object* l_Lean_Parser_Term_structInst_parenthesizer___closed__12; static lean_object* l_Lean_Parser_Term_letDecl_formatter___closed__3; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__19; static lean_object* l_Lean_Parser_Term_hole___closed__7; +static lean_object* l_Lean_Parser_Term_sufficesDecl___closed__15; static lean_object* l_Lean_Parser_Term_explicit___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_let_declRange___closed__4; static lean_object* l_Lean_Parser_Term_trailing__parser___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_arrow_declRange___closed__2; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__5; static lean_object* l_Lean_Parser_Term_dynamicQuot_formatter___closed__3; static lean_object* l_Lean_Parser_Term_letMVar_formatter___closed__3; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_explicitUniv_docString(lean_object*); @@ -2815,6 +2849,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Tactic_tacticSeqBracketed_formatt static lean_object* l___regBuiltin_Lean_Parser_Term_dbgTrace_declRange___closed__3; static lean_object* l_Lean_Parser_Term_matchAlt___closed__2; static lean_object* l_Lean_Parser_Term_letIdDecl___closed__9; +static lean_object* l_Lean_Parser_Term_attributes___closed__13; static lean_object* l_Lean_Parser_Term_whereDecls___closed__1; static lean_object* l_Lean_Parser_Term_structInstArrayRef___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_structInst_parenthesizer___closed__2; @@ -2830,12 +2865,12 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_doubleQuotedName_parenthesiz static lean_object* l___regBuiltin_Lean_Parser_Term_sort_docString___closed__1; static lean_object* l_Lean_Parser_Term_byTactic_parenthesizer___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_attributes_formatter___closed__1; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__24; static lean_object* l_Lean_Parser_Term_letMVar_parenthesizer___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_unreachable(lean_object*); static lean_object* l_Lean_Parser_Term_waitIfTypeMVar___closed__2; static lean_object* l_Lean_Parser_Term_strictImplicitLeftBracket___closed__2; static lean_object* l_Lean_Parser_Term_proj_parenthesizer___closed__5; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__18; static lean_object* l___regBuiltin_Lean_Parser_Term_instBinder_parenthesizer___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_letEqnsDecl_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Tactic_tacticSeq1Indented___closed__7; @@ -2843,16 +2878,17 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_pipeCompletion_formatter___c static lean_object* l___regBuiltin_Lean_Parser_Term_waitIfTypeMVar_declRange___closed__5; static lean_object* l_Lean_Parser_Term_macroDollarArg___closed__5; static lean_object* l_Lean_Parser_Term_subst_formatter___closed__1; -static lean_object* l_Lean_Parser_Term_have___closed__11; static lean_object* l___regBuiltin_Lean_Parser_Term_str_declRange___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Term_namedPattern_declRange___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_leading__parser_declRange___closed__3; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_binrel__no__prop_docString(lean_object*); +lean_object* l_Lean_Parser_hygieneInfo_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_doubleQuotedName___closed__1; static lean_object* l_Lean_Parser_Term_attributes_parenthesizer___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Tactic_quotSeq_formatter___closed__2; static lean_object* l_Lean_Parser_Term_leading__parser___closed__8; static lean_object* l___regBuiltin_Lean_Parser_Term_borrowed_declRange___closed__5; +static lean_object* l_Lean_Parser_Term_stateRefT___closed__11; static lean_object* l_Lean_Parser_Term_letEqnsDecl___closed__5; static lean_object* l_Lean_Parser_Term_basicFun_parenthesizer___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_binop_declRange___closed__3; @@ -2865,7 +2901,6 @@ static lean_object* l_Lean_Parser_Term_forall___closed__4; static lean_object* l_Lean_Parser_Term_basicFun___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_tacticSeqBracketed; LEAN_EXPORT lean_object* l_Lean_Parser_Term_isIdent___boxed(lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__10; static lean_object* l_Lean_Parser_Term_explicitUniv_parenthesizer___closed__8; static lean_object* l_Lean_Parser_Term_namedPattern_parenthesizer___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Term_unop_formatter___closed__1; @@ -2874,6 +2909,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_dotIdent(lean_object*); static lean_object* l_Lean_Parser_Term_local___closed__7; static lean_object* l_Lean_Parser_Term_noindex_formatter___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_binop__lazy_declRange___closed__2; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__42; static lean_object* l_Lean_Parser_Term_explicitUniv___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_dotIdent_declRange___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_proj_declRange___closed__5; @@ -2896,6 +2932,7 @@ static lean_object* l_Lean_Parser_Term_inaccessible___closed__7; static lean_object* l_Lean_Parser_Term_leading__parser___closed__11; static lean_object* l_Lean_Parser_Term_matchDiscr_parenthesizer___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_structInst_declRange___closed__1; +static lean_object* l_Lean_Parser_Term_match_parenthesizer___closed__14; static lean_object* l_Lean_Parser_Term_binderDefault_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Term_arrow_parenthesizer___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_dynamicQuot_formatter(lean_object*); @@ -2969,7 +3006,6 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_ensureExpectedType_formatter(lean_ob static lean_object* l_Lean_Parser_Term_sort_formatter___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_suffices_declRange___closed__6; static lean_object* l_Lean_Parser_Term_namedArgument_formatter___closed__2; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__17; LEAN_EXPORT lean_object* l_Lean_Parser_Term_optExprPrecedence; static lean_object* l_Lean_Parser_Term_letIdLhs_parenthesizer___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Term_prop_declRange___closed__2; @@ -2992,7 +3028,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_let__fun_declRange___closed_ static lean_object* l_Lean_Parser_Term_let_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Term_letPatDecl_formatter___closed__4; static lean_object* l_Lean_Parser_Term_let__tmp___closed__11; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__40; static lean_object* l_Lean_Parser_Term_fun_parenthesizer___closed__5; static lean_object* l_Lean_Parser_Term_haveEqnsDecl_formatter___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Term_scoped; @@ -3010,6 +3045,7 @@ static lean_object* l_Lean_Parser_Term_attrInstance___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Term_binrel__no__prop_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Term___hyg_195____closed__6; static lean_object* l_Lean_Parser_Term_syntheticHole___closed__9; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__20; static lean_object* l_Lean_Parser_Term_letRecDecl___closed__9; static lean_object* l_Lean_Parser_Term_ensureExpectedType___closed__7; static lean_object* l_Lean_Parser_Term_generalizingParam___closed__10; @@ -3021,7 +3057,6 @@ static lean_object* l_Lean_Parser_Term_funImplicitBinder_parenthesizer___closed_ static lean_object* l___regBuiltin_Lean_Parser_Term_let__tmp_parenthesizer___closed__2; lean_object* l_Lean_Parser_ident_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_tacticSeq___closed__7; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__4; static lean_object* l_Lean_Parser_Term_argument_parenthesizer___closed__4; static lean_object* l_Lean_Parser_Term_type___closed__15; static lean_object* l_Lean_Parser_Term_optExprPrecedence_parenthesizer___closed__1; @@ -3033,6 +3068,7 @@ static lean_object* l_Lean_Parser_Term_let___closed__10; LEAN_EXPORT lean_object* l_Lean_Parser_Term_letIdBinder; static lean_object* l_Lean_Parser_Term_explicitUniv_formatter___closed__2; static lean_object* l_Lean_Parser_Term_letRecDecl___closed__4; +static lean_object* l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__12; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_clear_parenthesizer(lean_object*); static lean_object* l_Lean_Parser_Term_motive___closed__3; static lean_object* l_Lean_Parser_Term_waitIfTypeContainsMVar_formatter___closed__1; @@ -3094,12 +3130,12 @@ static lean_object* l_Lean_Parser_Term_explicitUniv___closed__9; static lean_object* l___regBuiltin_Lean_Parser_Term_defaultOrOfNonempty_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_assert___closed__9; static lean_object* l___regBuiltin_Lean_Parser_Term_forInMacro_x27_declRange___closed__2; +static lean_object* l_Lean_Parser_Term_anonymousCtor_formatter___closed__9; static lean_object* l_Lean_Parser_darrow___closed__2; static lean_object* l_Lean_Parser_Term_funBinder_parenthesizer___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_byTactic_docString___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Term_strictImplicitRightBracket; LEAN_EXPORT lean_object* l_Lean_Parser_Term_structInstField_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Term_unop_parenthesizer___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Term_assert_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_fromTerm_formatter___closed__1; static lean_object* l_Lean_Parser_Term_assert_formatter___closed__1; @@ -3121,6 +3157,7 @@ static lean_object* l_Lean_Parser_Term_haveIdDecl_parenthesizer___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_letEqnsDecl_formatter___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Term_matchAlt_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_ensureTypeOf_declRange___closed__2; +static lean_object* l_Lean_Parser_Term_instBinder_parenthesizer___closed__7; static lean_object* l_Lean_Parser_Term_unop_formatter___closed__1; extern lean_object* l_Lean_Parser_scientificLit; static lean_object* l_Lean_Parser_Term_noErrorIfUnused___closed__3; @@ -3128,6 +3165,7 @@ static lean_object* l_Lean_Parser_Term_matchAlt___closed__10; LEAN_EXPORT lean_object* l_Lean_Parser_Term_quotedName_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_inaccessible_declRange___closed__5; static lean_object* l_Lean_Parser_Term_tuple_parenthesizer___closed__11; +static lean_object* l_Lean_Parser_Term_whereDecls_formatter___closed__10; static lean_object* l_Lean_Parser_Term_match_parenthesizer___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_waitIfTypeContainsMVar_declRange___closed__1; static lean_object* l_Lean_Parser_Term_waitIfTypeMVar___closed__10; @@ -3186,6 +3224,7 @@ static lean_object* l_Lean_Parser_Term_structInst_parenthesizer___closed__10; static lean_object* l_Lean_Parser_Term_instBinder_formatter___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_waitIfTypeMVar_declRange___closed__6; static lean_object* l_Lean_Parser_Term_letMVar_formatter___closed__6; +static lean_object* l_Lean_Parser_Term_withAnonymousAntiquot_formatter___closed__8; static lean_object* l_Lean_Parser_Term_type_formatter___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Term_anonymousCtor_declRange___closed__3; static lean_object* l_Lean_Parser_Term_whereDecls___closed__6; @@ -3202,7 +3241,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_noErrorIfUnused_declRange___ static lean_object* l_Lean_Parser_Term_attrKind___closed__6; static lean_object* l_Lean_Parser_Term_generalizingParam_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_assert_parenthesizer___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__12; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_let_declRange(lean_object*); static lean_object* l_Lean_Parser_Term_binop__lazy___closed__7; static lean_object* l_Lean_Parser_Tactic_quotSeq_parenthesizer___closed__6; @@ -3213,12 +3251,14 @@ static lean_object* l_Lean_Parser_Term_matchAltsWhereDecls___closed__6; static lean_object* l_Lean_Parser_Term_match_formatter___closed__2; static lean_object* l_Lean_Parser_Term_assert_parenthesizer___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Term_inaccessible_docString___closed__1; +static lean_object* l_Lean_Parser_Term_sufficesDecl_formatter___closed__10; static lean_object* l_Lean_Parser_Term_dynamicQuot___closed__8; static lean_object* l_Lean_Parser_Term_matchAlt_parenthesizer___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Term_typeAscription_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_withAnonymousAntiquot_formatter(lean_object*); static lean_object* l_Lean_Parser_Term_proj___closed__2; static lean_object* l_Lean_Parser_Term_depArrow_formatter___closed__6; +static lean_object* l_Lean_Parser_Term_sufficesDecl_formatter___closed__12; static lean_object* l_Lean_Parser_Tactic_seq1___closed__6; static lean_object* l_Lean_Parser_Term_matchAlt___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_haveEqnsDecl_formatter___closed__1; @@ -3233,16 +3273,15 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_bracketedBinder___boxed(lean_object* static lean_object* l_Lean_Parser_Term_typeAscription_formatter___closed__11; static lean_object* l_Lean_Parser_Term_motive_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_scientific_formatter___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__16; LEAN_EXPORT lean_object* l_Lean_Parser_Term_doubleQuotedName_parenthesizer___closed__3___boxed__const__1; static lean_object* l_Lean_Parser_Term_letIdDecl___closed__5; static lean_object* l_Lean_Parser_Term_binrel___closed__8; -static lean_object* l_Lean_Parser_Term_structInstLVal___closed__15; LEAN_EXPORT lean_object* l_Lean_Parser_Term_optEllipsis_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_noErrorIfUnused(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Term_syntheticHole; static lean_object* l_Lean_Parser_Term_namedArgument_parenthesizer___closed__6; static lean_object* l_Lean_Parser_Term_syntheticHole_formatter___closed__1; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__26; static lean_object* l_Lean_Parser_Term_stateRefT_formatter___closed__4; static lean_object* l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Term___hyg_195____closed__7; static lean_object* l_Lean_Parser_Term_optExprPrecedence___closed__3; @@ -3279,7 +3318,6 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_stateRefT_declRange(lea static lean_object* l_Lean_Parser_Term_basicFun_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Tactic_quotSeq_formatter___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Term_scientific___closed__2; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__36; LEAN_EXPORT lean_object* l_Lean_Parser_Term_type; static lean_object* l___regBuiltin_Lean_Parser_Term_letrec_declRange___closed__2; static lean_object* l_Lean_Parser_Term_motive___closed__6; @@ -3322,6 +3360,7 @@ static lean_object* l_Lean_Parser_Term_structInstFieldAbbrev___closed__9; static lean_object* l_Lean_Parser_Term_depArrow___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_trailing__parser_formatter(lean_object*); static lean_object* l_Lean_Parser_Term_forall___closed__13; +static lean_object* l_Lean_Parser_Term_anonymousCtor___closed__14; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_falseVal_parenthesizer(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Term_strictImplicitRightBracket_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_sufficesDecl___closed__2; @@ -3421,6 +3460,7 @@ static lean_object* l_Lean_Parser_Term_noImplicitLambda___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_Term_declName_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_structInst_formatter___closed__8; static lean_object* l_Lean_Parser_Term_binderDefault_formatter___closed__1; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__17; static lean_object* l___regBuiltin_Lean_Parser_Term_binrel_declRange___closed__1; static lean_object* l_Lean_Parser_Term_typeSpec_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_explicit_parenthesizer___closed__1; @@ -3435,6 +3475,7 @@ static lean_object* l_Lean_Parser_Term_binderTactic___closed__2; static lean_object* l_Lean_Parser_Term_unop___closed__10; static lean_object* l___regBuiltin_Lean_Parser_Term_show_formatter___closed__1; static lean_object* l_Lean_Parser_Tactic_quotSeq_parenthesizer___closed__3; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__4; static lean_object* l_Lean_Parser_Term_arrow___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_nomatch(lean_object*); static lean_object* l_Lean_Parser_Term_whereDecls_parenthesizer___closed__2; @@ -3574,7 +3615,6 @@ static lean_object* l_Lean_Parser_Term_type_parenthesizer___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Term_structInst_formatter___closed__2; static lean_object* l_Lean_Parser_Term_letMVar___closed__4; static lean_object* l_Lean_Parser_Tactic_sepByIndentSemicolon___closed__7; -static lean_object* l_Lean_Parser_Term_explicitBinder_parenthesizer___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Term_binderType_formatter(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_stateRefT___closed__9; static lean_object* l_Lean_Parser_Term_noindex_parenthesizer___closed__4; @@ -3583,7 +3623,6 @@ static lean_object* l_Lean_Parser_Term_falseVal___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Term_nomatch; static lean_object* l_Lean_Parser_Term_show___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Term_waitIfContainsMVar_formatter___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__14; static lean_object* l_Lean_Parser_Tactic_sepByIndentSemicolon___closed__10; static lean_object* l_Lean_Parser_Term_letRecDecl___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Term_noImplicitLambda_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -3598,6 +3637,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_structInst(lean_object* static lean_object* l_Lean_Parser_Term_syntheticHole_parenthesizer___closed__2; lean_object* l_id___rarg___boxed(lean_object*); static lean_object* l_Lean_Parser_Term_subst___closed__8; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__27; static lean_object* l_Lean_Parser_Term_byTactic_parenthesizer___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Term_matchAlt_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_app_parenthesizer(lean_object*); @@ -3605,7 +3645,6 @@ static lean_object* l_Lean_Parser_Tactic_tacticSeqBracketed_parenthesizer___clos static lean_object* l_Lean_Parser_Term_stateRefT_parenthesizer___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Term_waitIfTypeMVar_parenthesizer___closed__1; 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_initFn____x40_Lean_Parser_Term___hyg_5007____closed__8; static lean_object* l_Lean_Parser_Term_prop_formatter___closed__1; static lean_object* l_Lean_Parser_semicolonOrLinebreak_formatter___closed__1; lean_object* l_Lean_Parser_incQuotDepth(lean_object*); @@ -3656,6 +3695,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_doubleQuotedName_declRange__ static lean_object* l_Lean_Parser_Term_match_parenthesizer___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_forInMacro_x27(lean_object*); static lean_object* l_Lean_Parser_Term_ident_formatter___closed__2; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__31; static lean_object* l_Lean_Parser_Term_structInst___closed__10; static lean_object* l_Lean_Parser_Term_prop_parenthesizer___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Tactic_tacticSeq_parenthesizer___closed__2; @@ -3707,6 +3747,7 @@ static lean_object* l_Lean_Parser_Term_dotIdent_formatter___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_scientific_declRange(lean_object*); static lean_object* l_Lean_Parser_Term_byTactic_x27_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_structInstLVal___closed__4; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__8; static lean_object* l_Lean_Parser_Term_structInst___closed__9; lean_object* l_Lean_Parser_unicodeSymbol(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_byTactic_x27___closed__2; @@ -3732,7 +3773,6 @@ static lean_object* l_Lean_Parser_Term_defaultOrOfNonempty_parenthesizer___close lean_object* l_Lean_Parser_registerAlias(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_letRecDecl_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_tuple___closed__3; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__29; lean_object* l_Lean_Parser_checkStackTop(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_let__fun___closed__9; static lean_object* l_Lean_Parser_Term_noindex___closed__2; @@ -3741,7 +3781,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_defaultOrOfNonempty_formatte static lean_object* l_Lean_Parser_Term_suffices___closed__11; static lean_object* l_Lean_Parser_Term_doubleQuotedName_formatter___closed__7; static lean_object* l_Lean_Parser_Term_sort_parenthesizer___closed__4; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__21; LEAN_EXPORT lean_object* l_Lean_Parser_Term_strictImplicitBinder(uint8_t); static lean_object* l_Lean_Parser_Term_sort___closed__6; static lean_object* l_Lean_Parser_Term_forall_parenthesizer___closed__4; @@ -3784,7 +3823,6 @@ static lean_object* l_Lean_Parser_Term_namedPattern___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_Term_cdot_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_let__delayed_declRange___closed__2; static lean_object* l_Lean_Parser_Term_unop_parenthesizer___closed__3; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__23; static lean_object* l_Lean_Parser_Term_namedArgument___closed__8; static lean_object* l_Lean_Parser_Term_typeOf_formatter___closed__1; static lean_object* l_Lean_Parser_Term_forInMacro___closed__3; @@ -3797,14 +3835,14 @@ static lean_object* l_Lean_Parser_Term_waitIfTypeContainsMVar___closed__3; static lean_object* l_Lean_Parser_Term_match_parenthesizer___closed__10; static lean_object* l_Lean_Parser_Term_funStrictImplicitBinder_formatter___closed__6; static lean_object* l_Lean_Parser_Term_dotIdent_formatter___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__20; static lean_object* l_Lean_Parser_Term_withAnonymousAntiquot_formatter___closed__3; static lean_object* l_Lean_Parser_Term_typeAscription_parenthesizer___closed__4; +static lean_object* l_Lean_Parser_Term_stateRefT_formatter___closed__7; static lean_object* l_Lean_Parser_Term_fromTerm___closed__7; lean_object* l_Lean_Parser_many1(lean_object*); static lean_object* l_Lean_Parser_Term_tuple___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Term_noImplicitLambda_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__19; +static lean_object* l_Lean_Parser_Term_attributes_parenthesizer___closed__8; static lean_object* l_Lean_Parser_Command_docComment___closed__6; static lean_object* l_Lean_Parser_Term_dotIdent___closed__7; static lean_object* l_Lean_Parser_Term_letMVar_formatter___closed__1; @@ -3846,6 +3884,7 @@ static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_94____clo static lean_object* l___regBuiltin_Lean_Parser_Term_explicit_docString___closed__1; static lean_object* l_Lean_Parser_Term_basicFun_formatter___closed__10; static lean_object* l_Lean_Parser_Term_matchAlt_parenthesizer___closed__4; +static lean_object* l_Lean_Parser_Term_sufficesDecl___closed__14; static lean_object* l___regBuiltin_Lean_Parser_Term_letEqnsDecl_formatter___closed__1; static lean_object* l_Lean_Parser_Term_declName___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Term_syntheticHole_declRange___closed__7; @@ -3875,6 +3914,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Command_commentBody; static lean_object* l_Lean_Parser_Term_trailing__parser___closed__8; static lean_object* l_Lean_Parser_Term_trailing__parser_formatter___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Term_forall_formatter___closed__1; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__30; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_letRecDecls_formatter(lean_object*); static lean_object* l_Lean_Parser_Term_paren_parenthesizer___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Term_macroLastArg; @@ -3910,6 +3950,7 @@ static lean_object* l_Lean_Parser_Term_waitIfTypeContainsMVar_formatter___closed static lean_object* l_Lean_Parser_Term_explicit___closed__9; static lean_object* l_Lean_Parser_Term_let___closed__7; LEAN_EXPORT lean_object* l_Lean_Parser_Term_letrec_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__22; static lean_object* l_Lean_Parser_Term_tuple_formatter___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_anonymousCtor_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_type_formatter___closed__1; @@ -3964,6 +4005,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_let__delayed_formatter(lean_object*, lean_object* l_Lean_PrettyPrinter_Formatter_checkColGe_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_subst___closed__1; static lean_object* l_Lean_Parser_Term_funImplicitBinder_formatter___closed__2; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__7; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_app(lean_object*); static lean_object* l_Lean_Parser_Term_stateRefT_formatter___closed__5; static lean_object* l_Lean_Parser_Term_fromTerm___closed__3; @@ -3979,7 +4021,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_ellipsis_parenthesizer___clo static lean_object* l_Lean_Parser_Term_letRecDecl_formatter___closed__2; static lean_object* l_Lean_Parser_Term_basicFun___closed__9; static lean_object* l___regBuiltin_Lean_Parser_Term_unop_declRange___closed__7; -static lean_object* l_Lean_Parser_Term_trailing__parser___closed__11; LEAN_EXPORT lean_object* l_Lean_Parser_Term_byTactic_x27_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_ellipsis___closed__7; lean_object* l_Lean_Parser_sepBy(lean_object*, lean_object*, lean_object*, uint8_t); @@ -4006,6 +4047,7 @@ static lean_object* l_Lean_Parser_Term_borrowed___closed__4; static lean_object* l_Lean_Parser_Term_letEqnsDecl_formatter___closed__2; static lean_object* l_Lean_Parser_Term_ident___closed__1; static lean_object* l_Lean_Parser_Term_binrel_formatter___closed__5; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__25; static lean_object* l_Lean_Parser_Term_clear___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Term_subst_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_forInMacro_x27___closed__6; @@ -4055,6 +4097,7 @@ static lean_object* l_Lean_Parser_Term_typeAscription_formatter___closed__12; static lean_object* l_Lean_Parser_Term_strictImplicitRightBracket___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_Term_proj_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_inaccessible_formatter___closed__1; +static lean_object* l_Lean_Parser_Term_anonymousCtor_formatter___closed__8; static lean_object* l___regBuiltin_Lean_Parser_Term_declName_parenthesizer___closed__2; static lean_object* l_Lean_Parser_Term_pipeProj_formatter___closed__2; static lean_object* l_Lean_Parser_Term_structInstArrayRef___closed__10; @@ -4167,6 +4210,7 @@ static lean_object* l_Lean_Parser_Term_noImplicitLambda_parenthesizer___closed__ static lean_object* l___regBuiltin_Lean_Parser_Term_typeSpec_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Term_macroDollarArg___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_inaccessible_formatter(lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__33; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_dbgTrace(lean_object*); static lean_object* l_Lean_Parser_Term_type_formatter___closed__3; static lean_object* l_Lean_Parser_Term_syntheticHole___closed__2; @@ -4204,6 +4248,7 @@ static lean_object* l_Lean_Parser_Term_haveEqnsDecl___closed__2; static lean_object* l_Lean_Parser_Term_doubleQuotedName_formatter___closed__8; static lean_object* l_Lean_Parser_Term_noindex___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Term_panic_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Term_whereDecls_parenthesizer___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Term_pipeProj_docString___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_letDecl_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Term_structInstArrayRef___closed__11; @@ -4307,7 +4352,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_namedPattern_declRange___clo lean_object* l_Lean_Parser_mkAntiquot(lean_object*, lean_object*, uint8_t, uint8_t); 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_Tactic_tacticSeq___closed__3; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__34; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_tacticSeqIndentGt; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_binop__lazy(lean_object*); static lean_object* l_Lean_Parser_Tactic_seq1_formatter___closed__2; @@ -4381,6 +4425,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_attrInstance_parenthesizer__ static lean_object* l_Lean_Parser_Term_noindex___closed__4; static lean_object* l_Lean_Parser_Term_letPatDecl_parenthesizer___closed__2; lean_object* l_Lean_Parser_incQuotDepth_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Term_typeAscription_parenthesizer___closed__14; static lean_object* l_Lean_Parser_Term_ellipsis___closed__2; static lean_object* l_Lean_Parser_Term_paren_formatter___closed__6; static lean_object* l_Lean_Parser_Term_strictImplicitRightBracket_parenthesizer___closed__2; @@ -4426,10 +4471,13 @@ static lean_object* l_Lean_Parser_Term_let__tmp___closed__3; 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_doubleQuotedName___closed__12; static lean_object* l_Lean_Parser_Tactic_tacticSeqBracketed_parenthesizer___closed__3; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__45; +static lean_object* l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__10; static lean_object* l_Lean_Parser_Term_tuple_parenthesizer___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Term_sort_declRange___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Term_basicFun; static lean_object* l_Lean_Parser_Term_attrKind_parenthesizer___closed__1; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__41; static lean_object* l___regBuiltin_Lean_Parser_Term_typeOf_declRange___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Term_depArrow_declRange___closed__7; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_ensureTypeOf_declRange(lean_object*); @@ -4476,6 +4524,7 @@ static lean_object* l_Lean_Parser_Term_tuple___closed__4; static lean_object* l_Lean_Parser_Term_dynamicQuot___closed__2; static lean_object* l_Lean_Parser_Term_strictImplicitLeftBracket_parenthesizer___closed__3; static lean_object* l_Lean_Parser_Term_attributes_parenthesizer___closed__2; +static lean_object* l_Lean_Parser_Term_sufficesDecl___closed__10; LEAN_EXPORT lean_object* l_Lean_Parser_Term_letIdBinder_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_basicFun_parenthesizer___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_Term_letEqnsDecl_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -4566,6 +4615,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_binop(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_defaultOrOfNonempty_declRange___closed__3; static lean_object* l_Lean_Parser_Term_nomatch___closed__8; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_unop_formatter(lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__39; LEAN_EXPORT lean_object* l_Lean_Parser_Term_noErrorIfUnused_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_panic_declRange(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_waitIfTypeMVar_declRange___closed__2; @@ -4590,6 +4640,7 @@ static lean_object* l_Lean_Parser_Term_syntheticHole___closed__5; lean_object* l_Lean_Parser_addBuiltinParser(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_doubleQuotedName_parenthesizer___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Term_dynamicQuot_declRange___closed__4; +static lean_object* l_Lean_Parser_Term_withAnonymousAntiquot___closed__14; static lean_object* l_Lean_Parser_Term_let__fun_formatter___closed__3; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_forInMacro_x27_parenthesizer(lean_object*); extern lean_object* l_Lean_Parser_minPrec; @@ -4598,6 +4649,7 @@ static lean_object* l_Lean_Parser_Term_stateRefT___closed__3; static lean_object* l_Lean_Parser_Term_namedPattern_formatter___closed__6; static lean_object* l_Lean_Parser_Term_typeOf_formatter___closed__3; static lean_object* l___regBuiltin_Lean_Parser_Term_leading__parser_declRange___closed__1; +static lean_object* l_Lean_Parser_Term_sufficesDecl_formatter___closed__9; static lean_object* l___regBuiltin_Lean_Parser_Term_byTactic_declRange___closed__3; static lean_object* l_Lean_Parser_Term_waitIfTypeContainsMVar___closed__7; static lean_object* l_Lean_Parser_Term_noImplicitLambda_parenthesizer___closed__4; @@ -4643,7 +4695,6 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_trueVal_parenthesizer(lean_object*, static lean_object* l_Lean_Parser_Term_let___closed__11; static lean_object* l___regBuiltin_Lean_Parser_Term_typeSpec_parenthesizer___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_namedArgument_parenthesizer___closed__2; -static lean_object* l_Lean_Parser_Term_basicFun_parenthesizer___closed__11; static lean_object* l___regBuiltin_Lean_Parser_Term_binop_declRange___closed__6; static lean_object* l_Lean_Parser_Term_macroDollarArg___closed__3; static lean_object* l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Term___hyg_195____closed__12; @@ -4763,6 +4814,7 @@ static lean_object* l_Lean_Parser_Term_attrInstance___closed__9; static lean_object* l_Lean_Parser_Term_completion_formatter___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_pipeProj_declRange(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_paren_declRange___closed__6; +static lean_object* l_Lean_Parser_Term_ellipsis___closed__10; static lean_object* l_Lean_Parser_Term_typeAscription___closed__3; static lean_object* l_Lean_Parser_Term_withDeclName___closed__4; static lean_object* l_Lean_Parser_Term_dbgTrace_parenthesizer___closed__4; @@ -4785,6 +4837,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_generalizingParam_parenthesi static lean_object* l___regBuiltin_Lean_Parser_Term_clear_parenthesizer___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_local_formatter(lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_noindex_parenthesizer(lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__38; static lean_object* l_Lean_Parser_Term_funImplicitBinder_parenthesizer___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_Term_explicit_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_binop_parenthesizer___closed__1; @@ -4796,6 +4849,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_dynamicQuot; static lean_object* l___regBuiltin_Lean_Parser_Term_typeSpec_formatter___closed__1; static lean_object* l_Lean_Parser_Term_doubleQuotedName___closed__6; static lean_object* l_Lean_Parser_Term_letRecDecls___closed__4; +static lean_object* l_Lean_Parser_Term_optIdent___closed__4; static lean_object* l_Lean_Parser_Tactic_tacticSeq1Indented_parenthesizer___closed__1; static lean_object* l_Lean_Parser_Term_generalizingParam_parenthesizer___closed__6; static lean_object* l___regBuiltin_Lean_Parser_Term_noImplicitLambda_declRange___closed__6; @@ -4878,6 +4932,7 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_explicitUniv_declRange___clo static lean_object* l_Lean_Parser_Term_macroDollarArg___closed__10; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_quotedName(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_waitIfContainsMVar_declRange___closed__5; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__14; static lean_object* l_Lean_Parser_Term_attrInstance___closed__8; static lean_object* l___regBuiltin_Lean_Parser_Term_type_declRange___closed__2; static lean_object* l_Lean_Parser_Term_typeSpec___closed__6; @@ -4885,7 +4940,6 @@ static lean_object* l_Lean_Parser_Term_borrowed_formatter___closed__4; lean_object* l_Lean_PrettyPrinter_Parenthesizer_term_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Term_sort___closed__5; static lean_object* l_Lean_Parser_Tactic_tacticSeq_parenthesizer___closed__2; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__11; static lean_object* l___regBuiltin_Lean_Parser_Term_structInstArrayRef_formatter___closed__2; static lean_object* l_Lean_Parser_Tactic_quot___closed__7; static lean_object* l___regBuiltin_Lean_Parser_Term_waitIfTypeContainsMVar_parenthesizer___closed__2; @@ -4894,7 +4948,6 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_ensureTypeOf_declRange___clo static lean_object* l___regBuiltin_Lean_Parser_Term_stateRefT_formatter___closed__2; static lean_object* l_Lean_Parser_Term_letIdLhs_formatter___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_dbgTrace_parenthesizer___closed__1; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__9; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Parser_Term_noImplicitLambda(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_letrec_formatter___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_borrowed_formatter___closed__2; @@ -4986,9 +5039,9 @@ static lean_object* l_Lean_Parser_Term_structInst_parenthesizer___closed__17; static lean_object* l_Lean_Parser_Term_suffices___closed__4; static lean_object* l_Lean_Parser_Term_letPatDecl_formatter___closed__7; static lean_object* l_Lean_Parser_Term_typeAscription_formatter___closed__9; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__2; static lean_object* l_Lean_Parser_Term_noImplicitLambda_formatter___closed__2; static lean_object* l_Lean_Parser_Term_type___closed__16; +static lean_object* l_Lean_Parser_Term_match___closed__18; static lean_object* l_Lean_Parser_Term_match_formatter___closed__7; static lean_object* l_Lean_Parser_Term_binop__lazy___closed__1; static lean_object* l_Lean_Parser_Term_waitIfTypeMVar_parenthesizer___closed__7; @@ -5074,9 +5127,11 @@ static lean_object* l___regBuiltin_Lean_Parser_Term_letMVar_declRange___closed__ static lean_object* l___regBuiltin_Lean_Parser_Term_completion_formatter___closed__2; static lean_object* l___regBuiltin_Lean_Parser_Term_trailing__parser_parenthesizer___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_Term_unreachable_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098_(lean_object*); static lean_object* l___regBuiltin_Lean_Parser_Term_dbgTrace_declRange___closed__1; static lean_object* l___regBuiltin_Lean_Parser_Term_match_declRange___closed__5; static lean_object* l___regBuiltin_Lean_Parser_Term_panic_declRange___closed__1; +static lean_object* l_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___closed__9; static lean_object* l_Lean_Parser_Term_matchAlts_formatter___closed__1; static lean_object* l_Lean_Parser_Term_binderDefault___closed__4; static lean_object* l___regBuiltin_Lean_Parser_Term_scientific_declRange___closed__5; @@ -11425,6 +11480,14 @@ 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_binderIdent() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Parser_Term_syntheticHole___closed__6; +return x_1; +} +} static lean_object* _init_l_Lean_Parser_Term_sorry___closed__1() { _start: { @@ -11540,7 +11603,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_sorry_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(123u); +x_1 = lean_unsigned_to_nat(124u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11552,7 +11615,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_sorry_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(124u); +x_1 = lean_unsigned_to_nat(125u); x_2 = lean_unsigned_to_nat(9u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11580,7 +11643,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_sorry_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(123u); +x_1 = lean_unsigned_to_nat(124u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11592,7 +11655,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_sorry_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(123u); +x_1 = lean_unsigned_to_nat(124u); x_2 = lean_unsigned_to_nat(34u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11958,7 +12021,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_cdot_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(129u); +x_1 = lean_unsigned_to_nat(130u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11970,7 +12033,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_cdot_declRange___close _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(20u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11998,7 +12061,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_cdot_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(129u); +x_1 = lean_unsigned_to_nat(130u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -12010,7 +12073,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_cdot_declRange___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(129u); +x_1 = lean_unsigned_to_nat(130u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -12332,7 +12395,7 @@ static lean_object* _init_l_Lean_Parser_Term_typeAscription___closed__7() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes(" : ", 3); +x_1 = lean_mk_string_from_bytes(" :", 2); return x_1; } } @@ -12348,33 +12411,43 @@ return x_2; static lean_object* _init_l_Lean_Parser_Term_typeAscription___closed__9() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_skip; +x_2 = l_Lean_Parser_Term_typeAscription___closed__6; +x_3 = l_Lean_Parser_andthen(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Term_typeAscription___closed__10() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_typeAscription___closed__6; +x_1 = l_Lean_Parser_Term_typeAscription___closed__9; x_2 = l_Lean_Parser_optional(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_typeAscription___closed__10() { +static lean_object* _init_l_Lean_Parser_Term_typeAscription___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_typeAscription___closed__8; -x_2 = l_Lean_Parser_Term_typeAscription___closed__9; +x_2 = l_Lean_Parser_Term_typeAscription___closed__10; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_typeAscription___closed__11() { +static lean_object* _init_l_Lean_Parser_Term_typeAscription___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_typeAscription___closed__6; -x_2 = l_Lean_Parser_Term_typeAscription___closed__10; +x_2 = l_Lean_Parser_Term_typeAscription___closed__11; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_typeAscription___closed__12() { +static lean_object* _init_l_Lean_Parser_Term_typeAscription___closed__13() { _start: { lean_object* x_1; @@ -12382,17 +12455,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_withoutForbidden___lambda__1), 1, return x_1; } } -static lean_object* _init_l_Lean_Parser_Term_typeAscription___closed__13() { +static lean_object* _init_l_Lean_Parser_Term_typeAscription___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_typeAscription___closed__12; -x_2 = l_Lean_Parser_Term_typeAscription___closed__11; +x_1 = l_Lean_Parser_Term_typeAscription___closed__13; +x_2 = l_Lean_Parser_Term_typeAscription___closed__12; x_3 = l_Lean_Parser_adaptCacheableContext(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_typeAscription___closed__14() { +static lean_object* _init_l_Lean_Parser_Term_typeAscription___closed__15() { _start: { lean_object* x_1; @@ -12400,17 +12473,17 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_withoutPosition___lambda__1), 1, return x_1; } } -static lean_object* _init_l_Lean_Parser_Term_typeAscription___closed__15() { +static lean_object* _init_l_Lean_Parser_Term_typeAscription___closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_typeAscription___closed__14; -x_2 = l_Lean_Parser_Term_typeAscription___closed__13; +x_1 = l_Lean_Parser_Term_typeAscription___closed__15; +x_2 = l_Lean_Parser_Term_typeAscription___closed__14; x_3 = l_Lean_Parser_adaptCacheableContext(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_typeAscription___closed__16() { +static lean_object* _init_l_Lean_Parser_Term_typeAscription___closed__17() { _start: { lean_object* x_1; @@ -12418,62 +12491,62 @@ x_1 = lean_mk_string_from_bytes(")", 1); return x_1; } } -static lean_object* _init_l_Lean_Parser_Term_typeAscription___closed__17() { +static lean_object* _init_l_Lean_Parser_Term_typeAscription___closed__18() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_typeAscription___closed__16; +x_1 = l_Lean_Parser_Term_typeAscription___closed__17; x_2 = l_Lean_Parser_symbol(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_typeAscription___closed__18() { +static lean_object* _init_l_Lean_Parser_Term_typeAscription___closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_typeAscription___closed__15; -x_2 = l_Lean_Parser_Term_typeAscription___closed__17; +x_1 = l_Lean_Parser_Term_typeAscription___closed__16; +x_2 = l_Lean_Parser_Term_typeAscription___closed__18; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_typeAscription___closed__19() { +static lean_object* _init_l_Lean_Parser_Term_typeAscription___closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_typeAscription___closed__5; -x_2 = l_Lean_Parser_Term_typeAscription___closed__18; +x_2 = l_Lean_Parser_Term_typeAscription___closed__19; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_typeAscription___closed__20() { +static lean_object* _init_l_Lean_Parser_Term_typeAscription___closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_typeAscription___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_typeAscription___closed__19; +x_3 = l_Lean_Parser_Term_typeAscription___closed__20; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Term_typeAscription___closed__21() { +static lean_object* _init_l_Lean_Parser_Term_typeAscription___closed__22() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_typeAscription___closed__3; -x_2 = l_Lean_Parser_Term_typeAscription___closed__20; +x_2 = l_Lean_Parser_Term_typeAscription___closed__21; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_typeAscription___closed__22() { +static lean_object* _init_l_Lean_Parser_Term_typeAscription___closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_typeAscription___closed__2; -x_2 = l_Lean_Parser_Term_typeAscription___closed__21; +x_2 = l_Lean_Parser_Term_typeAscription___closed__22; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -12482,7 +12555,7 @@ static lean_object* _init_l_Lean_Parser_Term_typeAscription() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Term_typeAscription___closed__22; +x_1 = l_Lean_Parser_Term_typeAscription___closed__23; return x_1; } } @@ -12521,7 +12594,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_typeAscription_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(137u); +x_1 = lean_unsigned_to_nat(138u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -12533,8 +12606,8 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_typeAscription_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(138u); -x_2 = lean_unsigned_to_nat(97u); +x_1 = lean_unsigned_to_nat(139u); +x_2 = lean_unsigned_to_nat(109u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -12548,7 +12621,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Parser_Term_typeAscription_declRange___closed__1; x_2 = lean_unsigned_to_nat(23u); x_3 = l___regBuiltin_Lean_Parser_Term_typeAscription_declRange___closed__2; -x_4 = lean_unsigned_to_nat(97u); +x_4 = lean_unsigned_to_nat(109u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -12561,7 +12634,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_typeAscription_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(137u); +x_1 = lean_unsigned_to_nat(138u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -12573,7 +12646,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_typeAscription_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(137u); +x_1 = lean_unsigned_to_nat(138u); x_2 = lean_unsigned_to_nat(41u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -12661,105 +12734,125 @@ static lean_object* _init_l_Lean_Parser_Term_typeAscription_formatter___closed__ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Parser_termParser_formatter___rarg), 5, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_ppSpace_formatter___boxed), 5, 0); return x_1; } } static lean_object* _init_l_Lean_Parser_Term_typeAscription_formatter___closed__5() { _start: { -lean_object* x_1; lean_object* x_2; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Parser_termParser_formatter___rarg), 5, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Term_typeAscription_formatter___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; +x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__5; +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; +} +} +static lean_object* _init_l_Lean_Parser_Term_typeAscription_formatter___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__6; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_optional_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_typeAscription_formatter___closed__6() { +static lean_object* _init_l_Lean_Parser_Term_typeAscription_formatter___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__3; -x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__5; +x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__7; 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; } } -static lean_object* _init_l_Lean_Parser_Term_typeAscription_formatter___closed__7() { +static lean_object* _init_l_Lean_Parser_Term_typeAscription_formatter___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; -x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__6; +x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__5; +x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__8; 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; } } -static lean_object* _init_l_Lean_Parser_Term_typeAscription_formatter___closed__8() { +static lean_object* _init_l_Lean_Parser_Term_typeAscription_formatter___closed__10() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__7; +x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__9; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_withoutForbidden_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_typeAscription_formatter___closed__9() { +static lean_object* _init_l_Lean_Parser_Term_typeAscription_formatter___closed__11() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__8; +x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__10; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_withoutPosition_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_typeAscription_formatter___closed__10() { +static lean_object* _init_l_Lean_Parser_Term_typeAscription_formatter___closed__12() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_typeAscription___closed__16; +x_1 = l_Lean_Parser_Term_typeAscription___closed__17; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_typeAscription_formatter___closed__11() { +static lean_object* _init_l_Lean_Parser_Term_typeAscription_formatter___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__9; -x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__10; +x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__11; +x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__12; 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; } } -static lean_object* _init_l_Lean_Parser_Term_typeAscription_formatter___closed__12() { +static lean_object* _init_l_Lean_Parser_Term_typeAscription_formatter___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__2; -x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__11; +x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__13; 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; } } -static lean_object* _init_l_Lean_Parser_Term_typeAscription_formatter___closed__13() { +static lean_object* _init_l_Lean_Parser_Term_typeAscription_formatter___closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_typeAscription___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_typeAscription_formatter___closed__12; +x_3 = l_Lean_Parser_Term_typeAscription_formatter___closed__14; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -12772,7 +12865,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_typeAscription_formatter(lean_object { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Term_typeAscription_formatter___closed__1; -x_7 = l_Lean_Parser_Term_typeAscription_formatter___closed__13; +x_7 = l_Lean_Parser_Term_typeAscription_formatter___closed__15; 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; } @@ -12861,98 +12954,118 @@ return x_2; static lean_object* _init_l_Lean_Parser_Term_typeAscription_parenthesizer___closed__5() { _start: { +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Parser_ppSpace_parenthesizer___boxed), 4, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Term_typeAscription_parenthesizer___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__5; +x_2 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__3; +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; +} +} +static lean_object* _init_l_Lean_Parser_Term_typeAscription_parenthesizer___closed__7() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__3; +x_1 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__6; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_optional_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_typeAscription_parenthesizer___closed__6() { +static lean_object* _init_l_Lean_Parser_Term_typeAscription_parenthesizer___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__4; -x_2 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__5; +x_2 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__7; 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; } } -static lean_object* _init_l_Lean_Parser_Term_typeAscription_parenthesizer___closed__7() { +static lean_object* _init_l_Lean_Parser_Term_typeAscription_parenthesizer___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__3; -x_2 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__6; +x_2 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__8; 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; } } -static lean_object* _init_l_Lean_Parser_Term_typeAscription_parenthesizer___closed__8() { +static lean_object* _init_l_Lean_Parser_Term_typeAscription_parenthesizer___closed__10() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__7; +x_1 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__9; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_withoutForbidden_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_typeAscription_parenthesizer___closed__9() { +static lean_object* _init_l_Lean_Parser_Term_typeAscription_parenthesizer___closed__11() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__8; +x_1 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__10; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_withoutPosition_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_typeAscription_parenthesizer___closed__10() { +static lean_object* _init_l_Lean_Parser_Term_typeAscription_parenthesizer___closed__12() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_typeAscription___closed__16; +x_1 = l_Lean_Parser_Term_typeAscription___closed__17; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_typeAscription_parenthesizer___closed__11() { +static lean_object* _init_l_Lean_Parser_Term_typeAscription_parenthesizer___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__9; -x_2 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__10; +x_1 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__11; +x_2 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__12; 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; } } -static lean_object* _init_l_Lean_Parser_Term_typeAscription_parenthesizer___closed__12() { +static lean_object* _init_l_Lean_Parser_Term_typeAscription_parenthesizer___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__11; +x_2 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__13; 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; } } -static lean_object* _init_l_Lean_Parser_Term_typeAscription_parenthesizer___closed__13() { +static lean_object* _init_l_Lean_Parser_Term_typeAscription_parenthesizer___closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_typeAscription___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__12; +x_3 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__14; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -12965,7 +13078,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_typeAscription_parenthesizer(lean_ob { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__13; +x_7 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__15; 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; } @@ -13088,7 +13201,7 @@ static lean_object* _init_l_Lean_Parser_Term_tuple___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_typeAscription___closed__12; +x_1 = l_Lean_Parser_Term_typeAscription___closed__13; x_2 = l_Lean_Parser_Term_tuple___closed__8; x_3 = l_Lean_Parser_adaptCacheableContext(x_1, x_2); return x_3; @@ -13098,7 +13211,7 @@ static lean_object* _init_l_Lean_Parser_Term_tuple___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_typeAscription___closed__14; +x_1 = l_Lean_Parser_Term_typeAscription___closed__15; x_2 = l_Lean_Parser_Term_tuple___closed__9; x_3 = l_Lean_Parser_adaptCacheableContext(x_1, x_2); return x_3; @@ -13118,7 +13231,7 @@ static lean_object* _init_l_Lean_Parser_Term_tuple___closed__12() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_tuple___closed__11; -x_2 = l_Lean_Parser_Term_typeAscription___closed__17; +x_2 = l_Lean_Parser_Term_typeAscription___closed__18; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -13207,7 +13320,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_tuple_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(140u); +x_1 = lean_unsigned_to_nat(141u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13219,7 +13332,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_tuple_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(141u); +x_1 = lean_unsigned_to_nat(142u); x_2 = lean_unsigned_to_nat(108u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13247,7 +13360,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_tuple_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(140u); +x_1 = lean_unsigned_to_nat(141u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13259,7 +13372,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_tuple_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(140u); +x_1 = lean_unsigned_to_nat(141u); x_2 = lean_unsigned_to_nat(32u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13337,7 +13450,7 @@ static lean_object* _init_l_Lean_Parser_Term_tuple_formatter___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; uint8_t x_4; lean_object* x_5; lean_object* x_6; -x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; +x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__5; x_2 = l_Lean_Parser_Term_tuple___closed__4; x_3 = l_Lean_Parser_Term_tuple_formatter___closed__2; x_4 = 0; @@ -13366,7 +13479,7 @@ static lean_object* _init_l_Lean_Parser_Term_tuple_formatter___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; +x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__5; x_2 = l_Lean_Parser_Term_tuple_formatter___closed__4; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -13409,7 +13522,7 @@ static lean_object* _init_l_Lean_Parser_Term_tuple_formatter___closed__9() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_tuple_formatter___closed__8; -x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__10; +x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__12; 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); @@ -13589,7 +13702,7 @@ static lean_object* _init_l_Lean_Parser_Term_tuple_parenthesizer___closed__9() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_tuple_parenthesizer___closed__8; -x_2 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__10; +x_2 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__12; 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); @@ -13701,7 +13814,7 @@ static lean_object* _init_l_Lean_Parser_Term_paren___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_typeAscription___closed__12; +x_1 = l_Lean_Parser_Term_typeAscription___closed__13; x_2 = l_Lean_Parser_Term_typeAscription___closed__6; x_3 = l_Lean_Parser_adaptCacheableContext(x_1, x_2); return x_3; @@ -13711,7 +13824,7 @@ static lean_object* _init_l_Lean_Parser_Term_paren___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_typeAscription___closed__14; +x_1 = l_Lean_Parser_Term_typeAscription___closed__15; x_2 = l_Lean_Parser_Term_paren___closed__4; x_3 = l_Lean_Parser_adaptCacheableContext(x_1, x_2); return x_3; @@ -13722,7 +13835,7 @@ static lean_object* _init_l_Lean_Parser_Term_paren___closed__6() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_paren___closed__5; -x_2 = l_Lean_Parser_Term_typeAscription___closed__17; +x_2 = l_Lean_Parser_Term_typeAscription___closed__18; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -13811,7 +13924,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_paren_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(151u); +x_1 = lean_unsigned_to_nat(152u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13823,7 +13936,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_paren_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(152u); +x_1 = lean_unsigned_to_nat(153u); x_2 = lean_unsigned_to_nat(81u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13851,7 +13964,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_paren_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(151u); +x_1 = lean_unsigned_to_nat(152u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13863,7 +13976,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_paren_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(151u); +x_1 = lean_unsigned_to_nat(152u); x_2 = lean_unsigned_to_nat(32u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13931,7 +14044,7 @@ static lean_object* _init_l_Lean_Parser_Term_paren_formatter___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; +x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__5; x_2 = lean_alloc_closure((void*)(l_Lean_ppDedentIfGrouped_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; @@ -13962,7 +14075,7 @@ static lean_object* _init_l_Lean_Parser_Term_paren_formatter___closed__5() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_paren_formatter___closed__4; -x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__10; +x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__12; 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); @@ -14091,7 +14204,7 @@ static lean_object* _init_l_Lean_Parser_Term_paren_parenthesizer___closed__5() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_paren_parenthesizer___closed__4; -x_2 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__10; +x_2 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__12; 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); @@ -14231,67 +14344,87 @@ return x_5; static lean_object* _init_l_Lean_Parser_Term_anonymousCtor___closed__7() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_typeAscription___closed__13; +x_2 = l_Lean_Parser_Term_anonymousCtor___closed__6; +x_3 = l_Lean_Parser_adaptCacheableContext(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Term_anonymousCtor___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_typeAscription___closed__15; +x_2 = l_Lean_Parser_Term_anonymousCtor___closed__7; +x_3 = l_Lean_Parser_adaptCacheableContext(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Term_anonymousCtor___closed__9() { +_start: +{ lean_object* x_1; x_1 = lean_mk_string_from_bytes("⟩", 3); return x_1; } } -static lean_object* _init_l_Lean_Parser_Term_anonymousCtor___closed__8() { +static lean_object* _init_l_Lean_Parser_Term_anonymousCtor___closed__10() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_anonymousCtor___closed__7; +x_1 = l_Lean_Parser_Term_anonymousCtor___closed__9; x_2 = l_Lean_Parser_symbol(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_anonymousCtor___closed__9() { +static lean_object* _init_l_Lean_Parser_Term_anonymousCtor___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_anonymousCtor___closed__6; -x_2 = l_Lean_Parser_Term_anonymousCtor___closed__8; +x_1 = l_Lean_Parser_Term_anonymousCtor___closed__8; +x_2 = l_Lean_Parser_Term_anonymousCtor___closed__10; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_anonymousCtor___closed__10() { +static lean_object* _init_l_Lean_Parser_Term_anonymousCtor___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_anonymousCtor___closed__5; -x_2 = l_Lean_Parser_Term_anonymousCtor___closed__9; +x_2 = l_Lean_Parser_Term_anonymousCtor___closed__11; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_anonymousCtor___closed__11() { +static lean_object* _init_l_Lean_Parser_Term_anonymousCtor___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_anonymousCtor___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_anonymousCtor___closed__10; +x_3 = l_Lean_Parser_Term_anonymousCtor___closed__12; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Term_anonymousCtor___closed__12() { +static lean_object* _init_l_Lean_Parser_Term_anonymousCtor___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_anonymousCtor___closed__3; -x_2 = l_Lean_Parser_Term_anonymousCtor___closed__11; +x_2 = l_Lean_Parser_Term_anonymousCtor___closed__13; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_anonymousCtor___closed__13() { +static lean_object* _init_l_Lean_Parser_Term_anonymousCtor___closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_anonymousCtor___closed__2; -x_2 = l_Lean_Parser_Term_anonymousCtor___closed__12; +x_2 = l_Lean_Parser_Term_anonymousCtor___closed__14; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -14300,7 +14433,7 @@ static lean_object* _init_l_Lean_Parser_Term_anonymousCtor() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Term_anonymousCtor___closed__13; +x_1 = l_Lean_Parser_Term_anonymousCtor___closed__15; return x_1; } } @@ -14339,7 +14472,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_anonymousCtor_declRang _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(160u); +x_1 = lean_unsigned_to_nat(161u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -14351,8 +14484,8 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_anonymousCtor_declRang _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(161u); -x_2 = lean_unsigned_to_nat(37u); +x_1 = lean_unsigned_to_nat(162u); +x_2 = lean_unsigned_to_nat(74u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -14366,7 +14499,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Parser_Term_anonymousCtor_declRange___closed__1; x_2 = lean_unsigned_to_nat(23u); x_3 = l___regBuiltin_Lean_Parser_Term_anonymousCtor_declRange___closed__2; -x_4 = lean_unsigned_to_nat(37u); +x_4 = lean_unsigned_to_nat(74u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -14379,7 +14512,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_anonymousCtor_declRang _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(160u); +x_1 = lean_unsigned_to_nat(161u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -14391,7 +14524,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_anonymousCtor_declRang _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(160u); +x_1 = lean_unsigned_to_nat(161u); x_2 = lean_unsigned_to_nat(40u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -14469,7 +14602,7 @@ static lean_object* _init_l_Lean_Parser_Term_anonymousCtor_formatter___closed__3 _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; uint8_t x_4; lean_object* x_5; lean_object* x_6; -x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; +x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__5; x_2 = l_Lean_Parser_Term_tuple___closed__4; x_3 = l_Lean_Parser_Term_tuple_formatter___closed__2; x_4 = 0; @@ -14486,8 +14619,8 @@ static lean_object* _init_l_Lean_Parser_Term_anonymousCtor_formatter___closed__4 _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_anonymousCtor___closed__7; -x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_formatter), 6, 1); +x_1 = l_Lean_Parser_Term_anonymousCtor_formatter___closed__3; +x_2 = lean_alloc_closure((void*)(l_Lean_Parser_withoutForbidden_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } @@ -14495,34 +14628,54 @@ return x_2; static lean_object* _init_l_Lean_Parser_Term_anonymousCtor_formatter___closed__5() { _start: { +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Term_anonymousCtor_formatter___closed__4; +x_2 = lean_alloc_closure((void*)(l_Lean_Parser_withoutPosition_formatter), 6, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Term_anonymousCtor_formatter___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Term_anonymousCtor___closed__9; +x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_formatter), 6, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Term_anonymousCtor_formatter___closed__7() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_anonymousCtor_formatter___closed__3; -x_2 = l_Lean_Parser_Term_anonymousCtor_formatter___closed__4; +x_1 = l_Lean_Parser_Term_anonymousCtor_formatter___closed__5; +x_2 = l_Lean_Parser_Term_anonymousCtor_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; } } -static lean_object* _init_l_Lean_Parser_Term_anonymousCtor_formatter___closed__6() { +static lean_object* _init_l_Lean_Parser_Term_anonymousCtor_formatter___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_anonymousCtor_formatter___closed__2; -x_2 = l_Lean_Parser_Term_anonymousCtor_formatter___closed__5; +x_2 = l_Lean_Parser_Term_anonymousCtor_formatter___closed__7; 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; } } -static lean_object* _init_l_Lean_Parser_Term_anonymousCtor_formatter___closed__7() { +static lean_object* _init_l_Lean_Parser_Term_anonymousCtor_formatter___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_anonymousCtor___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_anonymousCtor_formatter___closed__6; +x_3 = l_Lean_Parser_Term_anonymousCtor_formatter___closed__8; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -14535,7 +14688,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_anonymousCtor_formatter(lean_object* { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Term_anonymousCtor_formatter___closed__1; -x_7 = l_Lean_Parser_Term_anonymousCtor_formatter___closed__7; +x_7 = l_Lean_Parser_Term_anonymousCtor_formatter___closed__9; 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; } @@ -14622,8 +14775,8 @@ static lean_object* _init_l_Lean_Parser_Term_anonymousCtor_parenthesizer___close _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_anonymousCtor___closed__7; -x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_parenthesizer), 6, 1); +x_1 = l_Lean_Parser_Term_anonymousCtor_parenthesizer___closed__3; +x_2 = lean_alloc_closure((void*)(l_Lean_Parser_withoutForbidden_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } @@ -14631,34 +14784,54 @@ return x_2; static lean_object* _init_l_Lean_Parser_Term_anonymousCtor_parenthesizer___closed__5() { _start: { +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Term_anonymousCtor_parenthesizer___closed__4; +x_2 = lean_alloc_closure((void*)(l_Lean_Parser_withoutPosition_parenthesizer), 6, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Term_anonymousCtor_parenthesizer___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Term_anonymousCtor___closed__9; +x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_parenthesizer), 6, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Term_anonymousCtor_parenthesizer___closed__7() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_anonymousCtor_parenthesizer___closed__3; -x_2 = l_Lean_Parser_Term_anonymousCtor_parenthesizer___closed__4; +x_1 = l_Lean_Parser_Term_anonymousCtor_parenthesizer___closed__5; +x_2 = l_Lean_Parser_Term_anonymousCtor_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; } } -static lean_object* _init_l_Lean_Parser_Term_anonymousCtor_parenthesizer___closed__6() { +static lean_object* _init_l_Lean_Parser_Term_anonymousCtor_parenthesizer___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_anonymousCtor_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Term_anonymousCtor_parenthesizer___closed__5; +x_2 = l_Lean_Parser_Term_anonymousCtor_parenthesizer___closed__7; 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; } } -static lean_object* _init_l_Lean_Parser_Term_anonymousCtor_parenthesizer___closed__7() { +static lean_object* _init_l_Lean_Parser_Term_anonymousCtor_parenthesizer___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_anonymousCtor___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_anonymousCtor_parenthesizer___closed__6; +x_3 = l_Lean_Parser_Term_anonymousCtor_parenthesizer___closed__8; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -14671,7 +14844,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_anonymousCtor_parenthesizer(lean_obj { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Term_anonymousCtor_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Term_anonymousCtor_parenthesizer___closed__7; +x_7 = l_Lean_Parser_Term_anonymousCtor_parenthesizer___closed__9; 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; } @@ -14712,27 +14885,44 @@ return x_6; static lean_object* _init_l_Lean_Parser_Term_optIdent___closed__1() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(" : ", 3); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Term_optIdent___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Term_optIdent___closed__1; +x_2 = l_Lean_Parser_symbol(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Term_optIdent___closed__3() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_ident; -x_2 = l_Lean_Parser_Term_typeAscription___closed__8; +x_2 = l_Lean_Parser_Term_optIdent___closed__2; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_optIdent___closed__2() { +static lean_object* _init_l_Lean_Parser_Term_optIdent___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_optIdent___closed__1; +x_1 = l_Lean_Parser_Term_optIdent___closed__3; x_2 = l_Lean_Parser_atomic(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_optIdent___closed__3() { +static lean_object* _init_l_Lean_Parser_Term_optIdent___closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_optIdent___closed__2; +x_1 = l_Lean_Parser_Term_optIdent___closed__4; x_2 = l_Lean_Parser_optional(x_1); return x_2; } @@ -14741,7 +14931,7 @@ static lean_object* _init_l_Lean_Parser_Term_optIdent() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Term_optIdent___closed__3; +x_1 = l_Lean_Parser_Term_optIdent___closed__5; return x_1; } } @@ -14897,59 +15087,116 @@ static lean_object* _init_l_Lean_Parser_Term_sufficesDecl___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_binderIdent; +x_2 = l_Lean_Parser_Term_optIdent___closed__2; +x_3 = l_Lean_Parser_andthen(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Term_sufficesDecl___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("group", 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Term_sufficesDecl___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_Term_sufficesDecl___closed__5; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Term_sufficesDecl___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_sufficesDecl___closed__6; +x_2 = l_Lean_Parser_Term_sufficesDecl___closed__4; +x_3 = l_Lean_Parser_node(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Term_sufficesDecl___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Term_sufficesDecl___closed__7; +x_2 = l_Lean_Parser_atomic(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Term_sufficesDecl___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_sufficesDecl___closed__8; +x_2 = l_Lean_Parser_hygieneInfo; +x_3 = l_Lean_Parser_orelse(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Term_sufficesDecl___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_skip; x_2 = l_Lean_Parser_Term_showRhs; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_sufficesDecl___closed__5() { +static lean_object* _init_l_Lean_Parser_Term_sufficesDecl___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_typeAscription___closed__6; -x_2 = l_Lean_Parser_Term_sufficesDecl___closed__4; +x_2 = l_Lean_Parser_Term_sufficesDecl___closed__10; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_sufficesDecl___closed__6() { +static lean_object* _init_l_Lean_Parser_Term_sufficesDecl___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_optIdent; -x_2 = l_Lean_Parser_Term_sufficesDecl___closed__5; +x_1 = l_Lean_Parser_Term_sufficesDecl___closed__9; +x_2 = l_Lean_Parser_Term_sufficesDecl___closed__11; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_sufficesDecl___closed__7() { +static lean_object* _init_l_Lean_Parser_Term_sufficesDecl___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_sufficesDecl___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_sufficesDecl___closed__6; +x_3 = l_Lean_Parser_Term_sufficesDecl___closed__12; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Term_sufficesDecl___closed__8() { +static lean_object* _init_l_Lean_Parser_Term_sufficesDecl___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_sufficesDecl___closed__3; -x_2 = l_Lean_Parser_Term_sufficesDecl___closed__7; +x_2 = l_Lean_Parser_Term_sufficesDecl___closed__13; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_sufficesDecl___closed__9() { +static lean_object* _init_l_Lean_Parser_Term_sufficesDecl___closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_sufficesDecl___closed__2; -x_2 = l_Lean_Parser_Term_sufficesDecl___closed__8; +x_2 = l_Lean_Parser_Term_sufficesDecl___closed__14; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -14958,7 +15205,7 @@ static lean_object* _init_l_Lean_Parser_Term_sufficesDecl() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Term_sufficesDecl___closed__9; +x_1 = l_Lean_Parser_Term_sufficesDecl___closed__15; return x_1; } } @@ -15105,7 +15352,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_suffices_declRange___c _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(170u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -15117,7 +15364,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_suffices_declRange___c _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(171u); x_2 = lean_unsigned_to_nat(71u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -15145,7 +15392,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_suffices_declRange___c _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(170u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -15157,7 +15404,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_suffices_declRange___c _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(170u); x_2 = lean_unsigned_to_nat(37u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -15203,35 +15450,14 @@ x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -static lean_object* _init_l_Lean_Parser_Term_optIdent_formatter___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_syntheticHole_formatter___closed__3; -x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__3; -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; -} -} -static lean_object* _init_l_Lean_Parser_Term_optIdent_formatter___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_optIdent_formatter___closed__1; -x_2 = lean_alloc_closure((void*)(l_Lean_Parser_atomic_formatter), 6, 1); -lean_closure_set(x_2, 0, x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_Parser_Term_optIdent_formatter(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_Parser_Term_binderIdent_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_Parser_Term_optIdent_formatter___closed__2; -x_7 = l_Lean_Parser_optional_formatter(x_6, x_1, x_2, x_3, x_4, x_5); -return x_7; +lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_6 = l_Lean_Parser_Term_syntheticHole_formatter___closed__3; +x_7 = l___regBuiltin_Lean_Parser_Term_hole_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_fromTerm_formatter___closed__1() { @@ -15267,7 +15493,7 @@ static lean_object* _init_l_Lean_Parser_Term_fromTerm_formatter___closed__3() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_fromTerm_formatter___closed__2; -x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; +x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__5; 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); @@ -15437,16 +15663,18 @@ return x_7; static lean_object* _init_l_Lean_Parser_Term_sufficesDecl_formatter___closed__2() { _start: { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_ppSpace_formatter___boxed), 5, 0); -return x_1; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Term_optIdent___closed__1; +x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_formatter), 6, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; } } static lean_object* _init_l_Lean_Parser_Term_sufficesDecl_formatter___closed__3() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Term_showRhs_formatter), 5, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Term_binderIdent_formatter), 5, 0); return x_1; } } @@ -15454,8 +15682,8 @@ static lean_object* _init_l_Lean_Parser_Term_sufficesDecl_formatter___closed__4( _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_sufficesDecl_formatter___closed__2; -x_2 = l_Lean_Parser_Term_sufficesDecl_formatter___closed__3; +x_1 = l_Lean_Parser_Term_sufficesDecl_formatter___closed__3; +x_2 = l_Lean_Parser_Term_sufficesDecl_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); @@ -15465,42 +15693,94 @@ return x_3; static lean_object* _init_l_Lean_Parser_Term_sufficesDecl_formatter___closed__5() { _start: { +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Term_sufficesDecl_formatter___closed__4; +x_2 = lean_alloc_closure((void*)(l_Lean_Parser_group_formatter), 6, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Term_sufficesDecl_formatter___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Term_sufficesDecl_formatter___closed__5; +x_2 = lean_alloc_closure((void*)(l_Lean_Parser_atomic_formatter), 6, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Term_sufficesDecl_formatter___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Parser_hygieneInfo_formatter), 5, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Term_sufficesDecl_formatter___closed__8() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; -x_2 = l_Lean_Parser_Term_sufficesDecl_formatter___closed__4; -x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); +x_1 = l_Lean_Parser_Term_sufficesDecl_formatter___closed__6; +x_2 = l_Lean_Parser_Term_sufficesDecl_formatter___closed__7; +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_sufficesDecl_formatter___closed__6() { +static lean_object* _init_l_Lean_Parser_Term_sufficesDecl_formatter___closed__9() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Term_optIdent_formatter), 5, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Term_showRhs_formatter), 5, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_Term_sufficesDecl_formatter___closed__7() { +static lean_object* _init_l_Lean_Parser_Term_sufficesDecl_formatter___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_sufficesDecl_formatter___closed__6; -x_2 = l_Lean_Parser_Term_sufficesDecl_formatter___closed__5; +x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; +x_2 = l_Lean_Parser_Term_sufficesDecl_formatter___closed__9; 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; } } -static lean_object* _init_l_Lean_Parser_Term_sufficesDecl_formatter___closed__8() { +static lean_object* _init_l_Lean_Parser_Term_sufficesDecl_formatter___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__5; +x_2 = l_Lean_Parser_Term_sufficesDecl_formatter___closed__10; +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; +} +} +static lean_object* _init_l_Lean_Parser_Term_sufficesDecl_formatter___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_sufficesDecl_formatter___closed__8; +x_2 = l_Lean_Parser_Term_sufficesDecl_formatter___closed__11; +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; +} +} +static lean_object* _init_l_Lean_Parser_Term_sufficesDecl_formatter___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_sufficesDecl___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_sufficesDecl_formatter___closed__7; +x_3 = l_Lean_Parser_Term_sufficesDecl_formatter___closed__12; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -15513,7 +15793,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_sufficesDecl_formatter(lean_object* { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Term_sufficesDecl_formatter___closed__1; -x_7 = l_Lean_Parser_Term_sufficesDecl_formatter___closed__8; +x_7 = l_Lean_Parser_Term_sufficesDecl_formatter___closed__13; 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; } @@ -15677,7 +15957,7 @@ static lean_object* _init_l_Lean_Parser_Term_suffices_formatter___closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; +x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__5; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_Term_optSemicolon_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; @@ -15752,35 +16032,14 @@ 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_optIdent_parenthesizer___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_syntheticHole_parenthesizer___closed__3; -x_2 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__4; -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; -} -} -static lean_object* _init_l_Lean_Parser_Term_optIdent_parenthesizer___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_optIdent_parenthesizer___closed__1; -x_2 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_atomic_parenthesizer), 6, 1); -lean_closure_set(x_2, 0, x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_Parser_Term_optIdent_parenthesizer(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_Parser_Term_binderIdent_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; -x_6 = l_Lean_Parser_Term_optIdent_parenthesizer___closed__2; -x_7 = l_Lean_Parser_optional_parenthesizer(x_6, x_1, x_2, x_3, x_4, x_5); -return x_7; +lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_6 = l_Lean_Parser_Term_syntheticHole_parenthesizer___closed__3; +x_7 = l___regBuiltin_Lean_Parser_Term_hole_parenthesizer___closed__2; +x_8 = l_Lean_PrettyPrinter_Parenthesizer_orelse_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_fromTerm_parenthesizer___closed__1() { @@ -15986,16 +16245,18 @@ return x_7; static lean_object* _init_l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__2() { _start: { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Parser_ppSpace_parenthesizer___boxed), 4, 0); -return x_1; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Term_optIdent___closed__1; +x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_parenthesizer), 6, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; } } static lean_object* _init_l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__3() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Term_showRhs_parenthesizer), 5, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Term_binderIdent_parenthesizer), 5, 0); return x_1; } } @@ -16003,8 +16264,8 @@ static lean_object* _init_l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__3; +x_1 = l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__3; +x_2 = l_Lean_Parser_Term_sufficesDecl_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); @@ -16014,42 +16275,94 @@ return x_3; static lean_object* _init_l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__5() { _start: { +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__4; +x_2 = lean_alloc_closure((void*)(l_Lean_Parser_group_parenthesizer), 6, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__5; +x_2 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_atomic_parenthesizer), 6, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Parser_hygieneInfo_parenthesizer), 5, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__8() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__3; -x_2 = l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__4; -x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); +x_1 = l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__6; +x_2 = l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__7; +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_sufficesDecl_parenthesizer___closed__6() { +static lean_object* _init_l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__9() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Term_optIdent_parenthesizer), 5, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Term_showRhs_parenthesizer), 5, 0); return x_1; } } -static lean_object* _init_l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__7() { +static lean_object* _init_l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__6; -x_2 = l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__5; +x_1 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__5; +x_2 = l_Lean_Parser_Term_sufficesDecl_parenthesizer___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); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__8() { +static lean_object* _init_l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__3; +x_2 = l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__10; +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; +} +} +static lean_object* _init_l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__8; +x_2 = l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__11; +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; +} +} +static lean_object* _init_l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_sufficesDecl___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__7; +x_3 = l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__12; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -16062,7 +16375,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_sufficesDecl_parenthesizer(lean_obje { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__8; +x_7 = l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__13; 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; } @@ -16352,7 +16665,7 @@ static lean_object* _init_l_Lean_Parser_Term_show___closed__6() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_show___closed__5; -x_2 = l_Lean_Parser_Term_sufficesDecl___closed__5; +x_2 = l_Lean_Parser_Term_sufficesDecl___closed__11; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -16413,7 +16726,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_show_declRange___close _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(172u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -16425,7 +16738,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_show_declRange___close _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(172u); x_2 = lean_unsigned_to_nat(108u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -16453,7 +16766,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_show_declRange___close _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(172u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -16465,7 +16778,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_show_declRange___close _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(172u); x_2 = lean_unsigned_to_nat(33u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -16544,7 +16857,7 @@ static lean_object* _init_l_Lean_Parser_Term_show_formatter___closed__3() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_show_formatter___closed__2; -x_2 = l_Lean_Parser_Term_sufficesDecl_formatter___closed__5; +x_2 = l_Lean_Parser_Term_sufficesDecl_formatter___closed__11; 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); @@ -16641,7 +16954,7 @@ static lean_object* _init_l_Lean_Parser_Term_show_parenthesizer___closed__3() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_show_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__5; +x_2 = l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__11; 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); @@ -16758,7 +17071,7 @@ static lean_object* _init_l_Lean_Parser_Term_structInstArrayRef___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_typeAscription___closed__14; +x_1 = l_Lean_Parser_Term_typeAscription___closed__15; x_2 = l_Lean_Parser_Term_typeAscription___closed__6; x_3 = l_Lean_Parser_adaptCacheableContext(x_1, x_2); return x_3; @@ -16915,87 +17228,69 @@ return x_3; static lean_object* _init_l_Lean_Parser_Term_structInstLVal___closed__8() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("group", 5); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Term_structInstLVal___closed__9() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Parser_Term_structInstLVal___closed__8; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Parser_Term_structInstLVal___closed__10() { -_start: -{ lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_structInstLVal___closed__9; +x_1 = l_Lean_Parser_Term_sufficesDecl___closed__6; x_2 = l_Lean_Parser_Term_structInstLVal___closed__7; x_3 = l_Lean_Parser_node(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_structInstLVal___closed__11() { +static lean_object* _init_l_Lean_Parser_Term_structInstLVal___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_structInstLVal___closed__10; +x_1 = l_Lean_Parser_Term_structInstLVal___closed__8; x_2 = l_Lean_Parser_Term_structInstArrayRef; x_3 = l_Lean_Parser_orelse(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_structInstLVal___closed__12() { +static lean_object* _init_l_Lean_Parser_Term_structInstLVal___closed__10() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_structInstLVal___closed__11; +x_1 = l_Lean_Parser_Term_structInstLVal___closed__9; x_2 = l_Lean_Parser_many(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_structInstLVal___closed__13() { +static lean_object* _init_l_Lean_Parser_Term_structInstLVal___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_structInstLVal___closed__5; -x_2 = l_Lean_Parser_Term_structInstLVal___closed__12; +x_2 = l_Lean_Parser_Term_structInstLVal___closed__10; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_structInstLVal___closed__14() { +static lean_object* _init_l_Lean_Parser_Term_structInstLVal___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_structInstLVal___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_structInstLVal___closed__13; +x_3 = l_Lean_Parser_Term_structInstLVal___closed__11; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Term_structInstLVal___closed__15() { +static lean_object* _init_l_Lean_Parser_Term_structInstLVal___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_structInstLVal___closed__3; -x_2 = l_Lean_Parser_Term_structInstLVal___closed__14; +x_2 = l_Lean_Parser_Term_structInstLVal___closed__12; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_structInstLVal___closed__16() { +static lean_object* _init_l_Lean_Parser_Term_structInstLVal___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_structInstLVal___closed__2; -x_2 = l_Lean_Parser_Term_structInstLVal___closed__15; +x_2 = l_Lean_Parser_Term_structInstLVal___closed__13; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -17004,7 +17299,7 @@ static lean_object* _init_l_Lean_Parser_Term_structInstLVal() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Term_structInstLVal___closed__16; +x_1 = l_Lean_Parser_Term_structInstLVal___closed__14; return x_1; } } @@ -17297,7 +17592,7 @@ static lean_object* _init_l_Lean_Parser_Term_optEllipsis___closed__4() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("..", 2); +x_1 = lean_mk_string_from_bytes(" ..", 3); return x_1; } } @@ -17518,7 +17813,7 @@ static lean_object* _init_l_Lean_Parser_Term_structInst___closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_typeAscription___closed__8; +x_1 = l_Lean_Parser_Term_optIdent___closed__2; x_2 = l_Lean_Parser_Term_typeAscription___closed__6; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; @@ -17567,7 +17862,7 @@ static lean_object* _init_l_Lean_Parser_Term_structInst___closed__22() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_typeAscription___closed__14; +x_1 = l_Lean_Parser_Term_typeAscription___closed__15; x_2 = l_Lean_Parser_Term_structInst___closed__21; x_3 = l_Lean_Parser_adaptCacheableContext(x_1, x_2); return x_3; @@ -17684,7 +17979,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_structInst_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(193u); +x_1 = lean_unsigned_to_nat(194u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -17696,7 +17991,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_structInst_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(197u); +x_1 = lean_unsigned_to_nat(198u); x_2 = lean_unsigned_to_nat(46u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -17724,7 +18019,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_structInst_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(193u); +x_1 = lean_unsigned_to_nat(194u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -17736,7 +18031,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_structInst_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(193u); +x_1 = lean_unsigned_to_nat(194u); x_2 = lean_unsigned_to_nat(37u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -17955,7 +18250,7 @@ static lean_object* _init_l_Lean_Parser_Term_structInstArrayRef_formatter___clos _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; +x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__5; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_withoutPosition_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; @@ -18260,7 +18555,7 @@ static lean_object* _init_l_Lean_Parser_Term_structInstField_formatter___closed_ { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_structInstField_formatter___closed__2; -x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; +x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__5; 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); @@ -18552,8 +18847,8 @@ static lean_object* _init_l_Lean_Parser_Term_structInst_formatter___closed__9() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__3; -x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; +x_1 = l_Lean_Parser_Term_sufficesDecl_formatter___closed__2; +x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__5; 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); @@ -19459,7 +19754,7 @@ static lean_object* _init_l_Lean_Parser_Term_structInst_parenthesizer___closed__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__4; +x_1 = l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__2; x_2 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__3; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -19650,7 +19945,7 @@ static lean_object* _init_l_Lean_Parser_Term_typeSpec___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_typeAscription___closed__8; +x_1 = l_Lean_Parser_Term_optIdent___closed__2; x_2 = l_Lean_Parser_Term_typeAscription___closed__6; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; @@ -19854,7 +20149,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_explicit_declRange___c _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(205u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -19866,7 +20161,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_explicit_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(205u); +x_1 = lean_unsigned_to_nat(206u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -19894,7 +20189,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_explicit_declRange___c _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(205u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -19906,7 +20201,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_explicit_declRange___c _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(205u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -19985,7 +20280,7 @@ static lean_object* _init_l_Lean_Parser_Term_explicit_formatter___closed__3() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_explicit_formatter___closed__2; -x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; +x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__5; 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); @@ -20210,7 +20505,7 @@ static lean_object* _init_l_Lean_Parser_Term_inaccessible___closed__6() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_structInstArrayRef___closed__6; -x_2 = l_Lean_Parser_Term_typeAscription___closed__17; +x_2 = l_Lean_Parser_Term_typeAscription___closed__18; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -20299,7 +20594,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_inaccessible_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(210u); +x_1 = lean_unsigned_to_nat(211u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -20311,7 +20606,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_inaccessible_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(211u); +x_1 = lean_unsigned_to_nat(212u); x_2 = lean_unsigned_to_nat(43u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -20339,7 +20634,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_inaccessible_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(210u); +x_1 = lean_unsigned_to_nat(211u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -20351,7 +20646,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_inaccessible_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(210u); +x_1 = lean_unsigned_to_nat(211u); x_2 = lean_unsigned_to_nat(39u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -20430,7 +20725,7 @@ static lean_object* _init_l_Lean_Parser_Term_inaccessible_formatter___closed__3( { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_structInstArrayRef_formatter___closed__2; -x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__10; +x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__12; 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); @@ -20539,7 +20834,7 @@ static lean_object* _init_l_Lean_Parser_Term_inaccessible_parenthesizer___closed { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_structInstArrayRef_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__10; +x_2 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__12; 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); @@ -20615,14 +20910,6 @@ 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_binderIdent() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_Parser_Term_syntheticHole___closed__6; -return x_1; -} -} static lean_object* _init_l_Lean_Parser_Term_binderType___closed__1() { _start: { @@ -21063,9 +21350,9 @@ x_3 = l_Lean_Parser_Term_explicitBinder___closed__6; x_4 = l_Lean_Parser_andthen(x_2, x_3); x_5 = l_Lean_Parser_Term_explicitBinder___closed__4; x_6 = l_Lean_Parser_andthen(x_5, x_4); -x_7 = l_Lean_Parser_Term_typeAscription___closed__14; +x_7 = l_Lean_Parser_Term_typeAscription___closed__15; x_8 = l_Lean_Parser_adaptCacheableContext(x_7, x_6); -x_9 = l_Lean_Parser_Term_typeAscription___closed__17; +x_9 = l_Lean_Parser_Term_typeAscription___closed__18; x_10 = l_Lean_Parser_andthen(x_8, x_9); x_11 = l_Lean_Parser_Term_typeAscription___closed__5; x_12 = l_Lean_Parser_andthen(x_11, x_10); @@ -21126,7 +21413,7 @@ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_obj x_2 = l_Lean_Parser_Term_binderType(x_1); x_3 = l_Lean_Parser_Term_explicitBinder___closed__4; x_4 = l_Lean_Parser_andthen(x_3, x_2); -x_5 = l_Lean_Parser_Term_typeAscription___closed__14; +x_5 = l_Lean_Parser_Term_typeAscription___closed__15; x_6 = l_Lean_Parser_adaptCacheableContext(x_5, x_4); x_7 = l_Lean_Parser_Tactic_tacticSeqBracketed___closed__8; x_8 = l_Lean_Parser_andthen(x_6, x_7); @@ -21163,7 +21450,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_Parser_Term_structInstLVal___closed__9; +x_1 = l_Lean_Parser_Term_sufficesDecl___closed__6; x_2 = l_Lean_Parser_Term_strictImplicitLeftBracket___closed__1; x_3 = l_Lean_Parser_node(x_1, x_2); return x_3; @@ -21226,7 +21513,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_Parser_Term_structInstLVal___closed__9; +x_1 = l_Lean_Parser_Term_sufficesDecl___closed__6; x_2 = l_Lean_Parser_Term_strictImplicitRightBracket___closed__1; x_3 = l_Lean_Parser_node(x_1, x_2); return x_3; @@ -21383,7 +21670,7 @@ static lean_object* _init_l_Lean_Parser_Term_instBinder___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_typeAscription___closed__14; +x_1 = l_Lean_Parser_Term_typeAscription___closed__15; x_2 = l_Lean_Parser_Term_instBinder___closed__4; x_3 = l_Lean_Parser_adaptCacheableContext(x_1, x_2); return x_3; @@ -21746,16 +22033,6 @@ x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Term_binderIdent_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_syntheticHole_formatter___closed__3; -x_7 = l___regBuiltin_Lean_Parser_Term_hole_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; -} -} LEAN_EXPORT lean_object* l_Lean_Parser_Term_binderType_formatter(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { @@ -22070,22 +22347,14 @@ return x_7; static lean_object* _init_l_Lean_Parser_Term_explicitBinder_formatter___closed__2() { _start: { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Term_binderIdent_formatter), 5, 0); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Term_explicitBinder_formatter___closed__3() { -_start: -{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_explicitBinder_formatter___closed__2; +x_1 = l_Lean_Parser_Term_sufficesDecl_formatter___closed__3; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_many1_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_explicitBinder_formatter___closed__4() { +static lean_object* _init_l_Lean_Parser_Term_explicitBinder_formatter___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -22097,11 +22366,11 @@ lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_explicitBinder_formatter___closed__5() { +static lean_object* _init_l_Lean_Parser_Term_explicitBinder_formatter___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_explicitBinder_formatter___closed__4; +x_1 = l_Lean_Parser_Term_explicitBinder_formatter___closed__3; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_optional_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; @@ -22114,17 +22383,17 @@ lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_ob x_7 = lean_box(x_1); x_8 = lean_alloc_closure((void*)(l_Lean_Parser_Term_binderType_formatter___boxed), 6, 1); lean_closure_set(x_8, 0, x_7); -x_9 = l_Lean_Parser_Term_explicitBinder_formatter___closed__5; +x_9 = l_Lean_Parser_Term_explicitBinder_formatter___closed__4; x_10 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_10, 0, x_8); lean_closure_set(x_10, 1, x_9); -x_11 = l_Lean_Parser_Term_explicitBinder_formatter___closed__3; +x_11 = l_Lean_Parser_Term_explicitBinder_formatter___closed__2; x_12 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_12, 0, x_11); lean_closure_set(x_12, 1, x_10); x_13 = lean_alloc_closure((void*)(l_Lean_Parser_withoutPosition_formatter), 6, 1); lean_closure_set(x_13, 0, x_12); -x_14 = l_Lean_Parser_Term_typeAscription_formatter___closed__10; +x_14 = l_Lean_Parser_Term_typeAscription_formatter___closed__12; x_15 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_15, 0, x_13); lean_closure_set(x_15, 1, x_14); @@ -22303,7 +22572,7 @@ x_9 = l_Lean_Parser_Term_strictImplicitBinder_formatter___closed__2; x_10 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_10, 0, x_8); lean_closure_set(x_10, 1, x_9); -x_11 = l_Lean_Parser_Term_explicitBinder_formatter___closed__3; +x_11 = l_Lean_Parser_Term_explicitBinder_formatter___closed__2; x_12 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_12, 0, x_11); lean_closure_set(x_12, 1, x_10); @@ -22360,7 +22629,7 @@ lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_ob x_7 = lean_box(x_1); x_8 = lean_alloc_closure((void*)(l_Lean_Parser_Term_binderType_formatter___boxed), 6, 1); lean_closure_set(x_8, 0, x_7); -x_9 = l_Lean_Parser_Term_explicitBinder_formatter___closed__3; +x_9 = l_Lean_Parser_Term_explicitBinder_formatter___closed__2; x_10 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_10, 0, x_9); lean_closure_set(x_10, 1, x_8); @@ -22398,6 +22667,37 @@ x_8 = l_Lean_Parser_Term_implicitBinder_formatter(x_7, x_2, x_3, x_4, x_5, x_6); return x_8; } } +static lean_object* _init_l_Lean_Parser_Term_optIdent_formatter___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_syntheticHole_formatter___closed__3; +x_2 = l_Lean_Parser_Term_sufficesDecl_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; +} +} +static lean_object* _init_l_Lean_Parser_Term_optIdent_formatter___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Term_optIdent_formatter___closed__1; +x_2 = lean_alloc_closure((void*)(l_Lean_Parser_atomic_formatter), 6, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Parser_Term_optIdent_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_Parser_Term_optIdent_formatter___closed__2; +x_7 = l_Lean_Parser_optional_formatter(x_6, x_1, x_2, x_3, x_4, x_5); +return x_7; +} +} static lean_object* _init_l_Lean_Parser_Term_instBinder_formatter___closed__1() { _start: { @@ -22419,30 +22719,38 @@ return x_7; static lean_object* _init_l_Lean_Parser_Term_instBinder_formatter___closed__2() { _start: { +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Term_optIdent_formatter), 5, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Term_instBinder_formatter___closed__3() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_sufficesDecl_formatter___closed__6; -x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; +x_1 = l_Lean_Parser_Term_instBinder_formatter___closed__2; +x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__5; 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; } } -static lean_object* _init_l_Lean_Parser_Term_instBinder_formatter___closed__3() { +static lean_object* _init_l_Lean_Parser_Term_instBinder_formatter___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_instBinder_formatter___closed__2; +x_1 = l_Lean_Parser_Term_instBinder_formatter___closed__3; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_withoutPosition_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_instBinder_formatter___closed__4() { +static lean_object* _init_l_Lean_Parser_Term_instBinder_formatter___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_instBinder_formatter___closed__3; +x_1 = l_Lean_Parser_Term_instBinder_formatter___closed__4; x_2 = l_Lean_Parser_Term_structInstArrayRef_formatter___closed__3; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -22450,25 +22758,25 @@ lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_instBinder_formatter___closed__5() { +static lean_object* _init_l_Lean_Parser_Term_instBinder_formatter___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_structInstFieldAbbrev_formatter___closed__3; -x_2 = l_Lean_Parser_Term_instBinder_formatter___closed__4; +x_2 = l_Lean_Parser_Term_instBinder_formatter___closed__5; 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; } } -static lean_object* _init_l_Lean_Parser_Term_instBinder_formatter___closed__6() { +static lean_object* _init_l_Lean_Parser_Term_instBinder_formatter___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_instBinder___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_instBinder_formatter___closed__5; +x_3 = l_Lean_Parser_Term_instBinder_formatter___closed__6; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -22476,23 +22784,23 @@ lean_closure_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Term_instBinder_formatter___closed__7() { +static lean_object* _init_l_Lean_Parser_Term_instBinder_formatter___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_instBinder_formatter___closed__1; -x_2 = l_Lean_Parser_Term_instBinder_formatter___closed__6; +x_2 = l_Lean_Parser_Term_instBinder_formatter___closed__7; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_withAntiquot_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_instBinder_formatter___closed__8() { +static lean_object* _init_l_Lean_Parser_Term_instBinder_formatter___closed__9() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_instBinder_formatter___closed__7; +x_1 = l_Lean_Parser_Term_instBinder_formatter___closed__8; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_withCache_formatter___rarg), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; @@ -22502,7 +22810,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_instBinder_formatter(lean_object* x_ _start: { lean_object* x_6; lean_object* x_7; -x_6 = l_Lean_Parser_Term_instBinder_formatter___closed__8; +x_6 = l_Lean_Parser_Term_instBinder_formatter___closed__9; x_7 = l_Lean_Parser_ppGroup_formatter(x_6, x_1, x_2, x_3, x_4, x_5); return x_7; } @@ -22641,7 +22949,7 @@ static lean_object* _init_l_Lean_Parser_Term_depArrow_formatter___closed__4() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_depArrow_formatter___closed__3; -x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; +x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__5; 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); @@ -22717,21 +23025,11 @@ 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_binderIdent_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_syntheticHole_parenthesizer___closed__3; -x_7 = l___regBuiltin_Lean_Parser_Term_hole_parenthesizer___closed__2; -x_8 = l_Lean_PrettyPrinter_Parenthesizer_orelse_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_binderType_parenthesizer___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__4; +x_1 = l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__2; x_2 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__3; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -22978,22 +23276,14 @@ return x_7; static lean_object* _init_l_Lean_Parser_Term_explicitBinder_parenthesizer___closed__2() { _start: { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Term_binderIdent_parenthesizer), 5, 0); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Term_explicitBinder_parenthesizer___closed__3() { -_start: -{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_explicitBinder_parenthesizer___closed__2; +x_1 = l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__3; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_many1_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_explicitBinder_parenthesizer___closed__4() { +static lean_object* _init_l_Lean_Parser_Term_explicitBinder_parenthesizer___closed__3() { _start: { lean_object* x_1; @@ -23001,23 +23291,23 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Term_binderDefault_parenthesizer) return x_1; } } -static lean_object* _init_l_Lean_Parser_Term_explicitBinder_parenthesizer___closed__5() { +static lean_object* _init_l_Lean_Parser_Term_explicitBinder_parenthesizer___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_Lean_Parser_Term_binderTactic_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Term_explicitBinder_parenthesizer___closed__4; +x_2 = l_Lean_Parser_Term_explicitBinder_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_explicitBinder_parenthesizer___closed__6() { +static lean_object* _init_l_Lean_Parser_Term_explicitBinder_parenthesizer___closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_explicitBinder_parenthesizer___closed__5; +x_1 = l_Lean_Parser_Term_explicitBinder_parenthesizer___closed__4; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_optional_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; @@ -23030,17 +23320,17 @@ lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_ob x_7 = lean_box(x_1); x_8 = lean_alloc_closure((void*)(l_Lean_Parser_Term_binderType_parenthesizer___boxed), 6, 1); lean_closure_set(x_8, 0, x_7); -x_9 = l_Lean_Parser_Term_explicitBinder_parenthesizer___closed__6; +x_9 = l_Lean_Parser_Term_explicitBinder_parenthesizer___closed__5; x_10 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_10, 0, x_8); lean_closure_set(x_10, 1, x_9); -x_11 = l_Lean_Parser_Term_explicitBinder_parenthesizer___closed__3; +x_11 = l_Lean_Parser_Term_explicitBinder_parenthesizer___closed__2; x_12 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_12, 0, x_11); lean_closure_set(x_12, 1, x_10); x_13 = lean_alloc_closure((void*)(l_Lean_Parser_withoutPosition_parenthesizer), 6, 1); lean_closure_set(x_13, 0, x_12); -x_14 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__10; +x_14 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__12; x_15 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_15, 0, x_13); lean_closure_set(x_15, 1, x_14); @@ -23216,7 +23506,7 @@ x_9 = l_Lean_Parser_Term_strictImplicitBinder_parenthesizer___closed__2; x_10 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_10, 0, x_8); lean_closure_set(x_10, 1, x_9); -x_11 = l_Lean_Parser_Term_explicitBinder_parenthesizer___closed__3; +x_11 = l_Lean_Parser_Term_explicitBinder_parenthesizer___closed__2; x_12 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_12, 0, x_11); lean_closure_set(x_12, 1, x_10); @@ -23270,7 +23560,7 @@ lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_ob x_7 = lean_box(x_1); x_8 = lean_alloc_closure((void*)(l_Lean_Parser_Term_binderType_parenthesizer___boxed), 6, 1); lean_closure_set(x_8, 0, x_7); -x_9 = l_Lean_Parser_Term_explicitBinder_parenthesizer___closed__3; +x_9 = l_Lean_Parser_Term_explicitBinder_parenthesizer___closed__2; x_10 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_10, 0, x_9); lean_closure_set(x_10, 1, x_8); @@ -23305,6 +23595,37 @@ x_8 = l_Lean_Parser_Term_implicitBinder_parenthesizer(x_7, x_2, x_3, x_4, x_5, x return x_8; } } +static lean_object* _init_l_Lean_Parser_Term_optIdent_parenthesizer___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_syntheticHole_parenthesizer___closed__3; +x_2 = l_Lean_Parser_Term_sufficesDecl_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; +} +} +static lean_object* _init_l_Lean_Parser_Term_optIdent_parenthesizer___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Term_optIdent_parenthesizer___closed__1; +x_2 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_atomic_parenthesizer), 6, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Parser_Term_optIdent_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; +x_6 = l_Lean_Parser_Term_optIdent_parenthesizer___closed__2; +x_7 = l_Lean_Parser_optional_parenthesizer(x_6, x_1, x_2, x_3, x_4, x_5); +return x_7; +} +} static lean_object* _init_l_Lean_Parser_Term_instBinder_parenthesizer___closed__1() { _start: { @@ -23326,8 +23647,16 @@ return x_7; static lean_object* _init_l_Lean_Parser_Term_instBinder_parenthesizer___closed__2() { _start: { +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Parser_Term_optIdent_parenthesizer), 5, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Term_instBinder_parenthesizer___closed__3() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__6; +x_1 = l_Lean_Parser_Term_instBinder_parenthesizer___closed__2; x_2 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__3; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -23335,21 +23664,21 @@ lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_instBinder_parenthesizer___closed__3() { +static lean_object* _init_l_Lean_Parser_Term_instBinder_parenthesizer___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_instBinder_parenthesizer___closed__2; +x_1 = l_Lean_Parser_Term_instBinder_parenthesizer___closed__3; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_withoutPosition_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_instBinder_parenthesizer___closed__4() { +static lean_object* _init_l_Lean_Parser_Term_instBinder_parenthesizer___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_instBinder_parenthesizer___closed__3; +x_1 = l_Lean_Parser_Term_instBinder_parenthesizer___closed__4; x_2 = l_Lean_Parser_Term_structInstArrayRef_parenthesizer___closed__3; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -23357,25 +23686,25 @@ lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_instBinder_parenthesizer___closed__5() { +static lean_object* _init_l_Lean_Parser_Term_instBinder_parenthesizer___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_structInstFieldAbbrev_parenthesizer___closed__3; -x_2 = l_Lean_Parser_Term_instBinder_parenthesizer___closed__4; +x_2 = l_Lean_Parser_Term_instBinder_parenthesizer___closed__5; 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; } } -static lean_object* _init_l_Lean_Parser_Term_instBinder_parenthesizer___closed__6() { +static lean_object* _init_l_Lean_Parser_Term_instBinder_parenthesizer___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_instBinder___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_instBinder_parenthesizer___closed__5; +x_3 = l_Lean_Parser_Term_instBinder_parenthesizer___closed__6; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -23388,7 +23717,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_instBinder_parenthesizer(lean_object { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Term_instBinder_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Term_instBinder_parenthesizer___closed__6; +x_7 = l_Lean_Parser_Term_instBinder_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; } @@ -24014,7 +24343,7 @@ static lean_object* _init_l_Lean_Parser_Term_forall_formatter___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_explicitBinder_formatter___closed__2; +x_1 = l_Lean_Parser_Term_sufficesDecl_formatter___closed__3; 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); @@ -24026,7 +24355,7 @@ static lean_object* _init_l_Lean_Parser_Term_forall_formatter___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_sufficesDecl_formatter___closed__2; +x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; x_2 = l_Lean_Parser_Term_forall_formatter___closed__4; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -24049,7 +24378,7 @@ static lean_object* _init_l_Lean_Parser_Term_forall_formatter___closed__7() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_tuple_formatter___closed__2; -x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; +x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__5; 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); @@ -24286,7 +24615,7 @@ static lean_object* _init_l_Lean_Parser_Term_forall_parenthesizer___closed__4() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_explicitBinder_parenthesizer___closed__2; +x_1 = l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__3; 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); @@ -24298,7 +24627,7 @@ static lean_object* _init_l_Lean_Parser_Term_forall_parenthesizer___closed__5() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__2; +x_1 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__5; x_2 = l_Lean_Parser_Term_forall_parenthesizer___closed__4; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -24481,7 +24810,7 @@ static lean_object* _init_l_Lean_Parser_Term_matchAlt___closed__6() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("|", 1); +x_1 = lean_mk_string_from_bytes(" | ", 3); return x_1; } } @@ -24664,7 +24993,7 @@ static lean_object* _init_l_Lean_Parser_Term_matchDiscr___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_optIdent___closed__3; +x_1 = l_Lean_Parser_Term_optIdent___closed__5; x_2 = l_Lean_Parser_Term_typeAscription___closed__6; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; @@ -24970,7 +25299,7 @@ static lean_object* _init_l_Lean_Parser_Term_generalizingParam___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_typeAscription___closed__17; +x_1 = l_Lean_Parser_Term_typeAscription___closed__18; x_2 = l_Lean_Parser_skip; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; @@ -25257,89 +25586,106 @@ return x_5; static lean_object* _init_l_Lean_Parser_Term_match___closed__9() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(" with", 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Term_match___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Term_match___closed__9; +x_2 = l_Lean_Parser_symbol(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Term_match___closed__11() { +_start: +{ lean_object* x_1; lean_object* x_2; x_1 = l_Lean_Parser_Term_typeAscription___closed__6; x_2 = l_Lean_Parser_Term_matchAlts(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_match___closed__10() { +static lean_object* _init_l_Lean_Parser_Term_match___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_structInst___closed__7; -x_2 = l_Lean_Parser_Term_match___closed__9; +x_1 = l_Lean_Parser_Term_match___closed__10; +x_2 = l_Lean_Parser_Term_match___closed__11; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_match___closed__11() { +static lean_object* _init_l_Lean_Parser_Term_match___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_match___closed__8; -x_2 = l_Lean_Parser_Term_match___closed__10; +x_2 = l_Lean_Parser_Term_match___closed__12; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_match___closed__12() { +static lean_object* _init_l_Lean_Parser_Term_match___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_match___closed__7; -x_2 = l_Lean_Parser_Term_match___closed__11; +x_2 = l_Lean_Parser_Term_match___closed__13; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_match___closed__13() { +static lean_object* _init_l_Lean_Parser_Term_match___closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_match___closed__6; -x_2 = l_Lean_Parser_Term_match___closed__12; +x_2 = l_Lean_Parser_Term_match___closed__14; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_match___closed__14() { +static lean_object* _init_l_Lean_Parser_Term_match___closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_match___closed__5; -x_2 = l_Lean_Parser_Term_match___closed__13; +x_2 = l_Lean_Parser_Term_match___closed__15; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_match___closed__15() { +static lean_object* _init_l_Lean_Parser_Term_match___closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_match___closed__2; x_2 = l_Lean_Parser_leadPrec; -x_3 = l_Lean_Parser_Term_match___closed__14; +x_3 = l_Lean_Parser_Term_match___closed__16; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Term_match___closed__16() { +static lean_object* _init_l_Lean_Parser_Term_match___closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_match___closed__3; -x_2 = l_Lean_Parser_Term_match___closed__15; +x_2 = l_Lean_Parser_Term_match___closed__17; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_match___closed__17() { +static lean_object* _init_l_Lean_Parser_Term_match___closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_match___closed__2; -x_2 = l_Lean_Parser_Term_match___closed__16; +x_2 = l_Lean_Parser_Term_match___closed__18; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -25348,7 +25694,7 @@ static lean_object* _init_l_Lean_Parser_Term_match() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Term_match___closed__17; +x_1 = l_Lean_Parser_Term_match___closed__19; return x_1; } } @@ -25400,7 +25746,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_match_declRange___clos { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_unsigned_to_nat(345u); -x_2 = lean_unsigned_to_nat(32u); +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); @@ -25414,7 +25760,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Parser_Term_match_declRange___closed__1; x_2 = lean_unsigned_to_nat(23u); x_3 = l___regBuiltin_Lean_Parser_Term_match_declRange___closed__2; -x_4 = lean_unsigned_to_nat(32u); +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); @@ -25730,8 +26076,8 @@ static lean_object* _init_l_Lean_Parser_Term_generalizingParam_formatter___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__10; -x_2 = l_Lean_Parser_Term_sufficesDecl_formatter___closed__2; +x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__12; +x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; 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); @@ -26010,7 +26356,7 @@ static lean_object* _init_l_Lean_Parser_Term_matchDiscr_formatter___closed__3() { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_matchDiscr_formatter___closed__2; -x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; +x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__5; 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); @@ -26298,8 +26644,8 @@ static lean_object* _init_l_Lean_Parser_Term_match_formatter___closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; -x_2 = lean_alloc_closure((void*)(l_Lean_Parser_Term_matchAlts_formatter), 6, 1); +x_1 = l_Lean_Parser_Term_match___closed__9; +x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } @@ -26308,8 +26654,8 @@ static lean_object* _init_l_Lean_Parser_Term_match_formatter___closed__7() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_match_formatter___closed__6; -x_2 = lean_alloc_closure((void*)(l_Lean_ppDedent_formatter), 6, 1); +x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__5; +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; } @@ -26317,20 +26663,18 @@ return x_2; static lean_object* _init_l_Lean_Parser_Term_match_formatter___closed__8() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_structInst_formatter___closed__3; -x_2 = l_Lean_Parser_Term_match_formatter___closed__7; -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_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Term_match_formatter___closed__7; +x_2 = lean_alloc_closure((void*)(l_Lean_ppDedent_formatter), 6, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; } } static lean_object* _init_l_Lean_Parser_Term_match_formatter___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_match_formatter___closed__5; +x_1 = l_Lean_Parser_Term_match_formatter___closed__6; x_2 = l_Lean_Parser_Term_match_formatter___closed__8; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -26342,7 +26686,7 @@ static lean_object* _init_l_Lean_Parser_Term_match_formatter___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_match_formatter___closed__4; +x_1 = l_Lean_Parser_Term_match_formatter___closed__5; x_2 = l_Lean_Parser_Term_match_formatter___closed__9; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -26354,7 +26698,7 @@ static lean_object* _init_l_Lean_Parser_Term_match_formatter___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_match_formatter___closed__3; +x_1 = l_Lean_Parser_Term_match_formatter___closed__4; x_2 = l_Lean_Parser_Term_match_formatter___closed__10; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -26366,7 +26710,7 @@ static lean_object* _init_l_Lean_Parser_Term_match_formatter___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_match_formatter___closed__2; +x_1 = l_Lean_Parser_Term_match_formatter___closed__3; x_2 = l_Lean_Parser_Term_match_formatter___closed__11; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -26377,10 +26721,22 @@ return x_3; static lean_object* _init_l_Lean_Parser_Term_match_formatter___closed__13() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_match_formatter___closed__2; +x_2 = l_Lean_Parser_Term_match_formatter___closed__12; +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; +} +} +static lean_object* _init_l_Lean_Parser_Term_match_formatter___closed__14() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_match___closed__2; x_2 = l_Lean_Parser_leadPrec; -x_3 = l_Lean_Parser_Term_match_formatter___closed__12; +x_3 = l_Lean_Parser_Term_match_formatter___closed__13; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -26393,7 +26749,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_match_formatter(lean_object* x_1, le { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Term_match_formatter___closed__1; -x_7 = l_Lean_Parser_Term_match_formatter___closed__13; +x_7 = l_Lean_Parser_Term_match_formatter___closed__14; 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; } @@ -26676,8 +27032,8 @@ static lean_object* _init_l_Lean_Parser_Term_generalizingParam_parenthesizer___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__10; -x_2 = l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__2; +x_1 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__12; +x_2 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__5; 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); @@ -27244,89 +27600,99 @@ static lean_object* _init_l_Lean_Parser_Term_match_parenthesizer___closed__6() { _start: { lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Term_match___closed__9; +x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_parenthesizer), 6, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Term_match_parenthesizer___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; x_1 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__3; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_Term_matchAlts_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_match_parenthesizer___closed__7() { +static lean_object* _init_l_Lean_Parser_Term_match_parenthesizer___closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_match_parenthesizer___closed__6; +x_1 = l_Lean_Parser_Term_match_parenthesizer___closed__7; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_ppDedent_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_match_parenthesizer___closed__8() { +static lean_object* _init_l_Lean_Parser_Term_match_parenthesizer___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_structInst_parenthesizer___closed__3; -x_2 = l_Lean_Parser_Term_match_parenthesizer___closed__7; +x_1 = l_Lean_Parser_Term_match_parenthesizer___closed__6; +x_2 = l_Lean_Parser_Term_match_parenthesizer___closed__8; 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; } } -static lean_object* _init_l_Lean_Parser_Term_match_parenthesizer___closed__9() { +static lean_object* _init_l_Lean_Parser_Term_match_parenthesizer___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_match_parenthesizer___closed__5; -x_2 = l_Lean_Parser_Term_match_parenthesizer___closed__8; +x_2 = l_Lean_Parser_Term_match_parenthesizer___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); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_match_parenthesizer___closed__10() { +static lean_object* _init_l_Lean_Parser_Term_match_parenthesizer___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_match_parenthesizer___closed__4; -x_2 = l_Lean_Parser_Term_match_parenthesizer___closed__9; +x_2 = l_Lean_Parser_Term_match_parenthesizer___closed__10; 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; } } -static lean_object* _init_l_Lean_Parser_Term_match_parenthesizer___closed__11() { +static lean_object* _init_l_Lean_Parser_Term_match_parenthesizer___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_match_parenthesizer___closed__3; -x_2 = l_Lean_Parser_Term_match_parenthesizer___closed__10; +x_2 = l_Lean_Parser_Term_match_parenthesizer___closed__11; 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; } } -static lean_object* _init_l_Lean_Parser_Term_match_parenthesizer___closed__12() { +static lean_object* _init_l_Lean_Parser_Term_match_parenthesizer___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_match_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Term_match_parenthesizer___closed__11; +x_2 = l_Lean_Parser_Term_match_parenthesizer___closed__12; 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; } } -static lean_object* _init_l_Lean_Parser_Term_match_parenthesizer___closed__13() { +static lean_object* _init_l_Lean_Parser_Term_match_parenthesizer___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_match___closed__2; x_2 = l_Lean_Parser_leadPrec; -x_3 = l_Lean_Parser_Term_match_parenthesizer___closed__12; +x_3 = l_Lean_Parser_Term_match_parenthesizer___closed__13; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -27339,7 +27705,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_match_parenthesizer(lean_object* x_1 { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Term_match_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Term_match_parenthesizer___closed__13; +x_7 = l_Lean_Parser_Term_match_parenthesizer___closed__14; 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; } @@ -27641,7 +28007,7 @@ static lean_object* _init_l_Lean_Parser_Term_nomatch_formatter___closed__3() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_nomatch_formatter___closed__2; -x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; +x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__5; 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); @@ -27806,7 +28172,7 @@ static lean_object* _init_l_Lean_Parser_Term_funImplicitBinder___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_typeAscription___closed__8; +x_1 = l_Lean_Parser_Term_optIdent___closed__2; x_2 = l_Lean_Parser_Tactic_tacticSeqBracketed___closed__8; x_3 = l_Lean_Parser_orelse(x_1, x_2); return x_3; @@ -27891,7 +28257,7 @@ static lean_object* _init_l_Lean_Parser_Term_funStrictImplicitBinder___closed__1 _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_typeAscription___closed__8; +x_1 = l_Lean_Parser_Term_optIdent___closed__2; x_2 = l_Lean_Parser_Term_strictImplicitRightBracket; x_3 = l_Lean_Parser_orelse(x_1, x_2); return x_3; @@ -28141,49 +28507,39 @@ static lean_object* _init_l_Lean_Parser_Term_basicFun___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_skip; -x_2 = l_Lean_Parser_Term_typeAscription___closed__6; -x_3 = l_Lean_Parser_andthen(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Parser_Term_basicFun___closed__12() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_basicFun___closed__10; -x_2 = l_Lean_Parser_Term_basicFun___closed__11; +x_2 = l_Lean_Parser_Term_typeAscription___closed__9; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_basicFun___closed__13() { +static lean_object* _init_l_Lean_Parser_Term_basicFun___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_basicFun___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_basicFun___closed__12; +x_3 = l_Lean_Parser_Term_basicFun___closed__11; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Term_basicFun___closed__14() { +static lean_object* _init_l_Lean_Parser_Term_basicFun___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_basicFun___closed__3; -x_2 = l_Lean_Parser_Term_basicFun___closed__13; +x_2 = l_Lean_Parser_Term_basicFun___closed__12; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_basicFun___closed__15() { +static lean_object* _init_l_Lean_Parser_Term_basicFun___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_basicFun___closed__2; -x_2 = l_Lean_Parser_Term_basicFun___closed__14; +x_2 = l_Lean_Parser_Term_basicFun___closed__13; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -28192,7 +28548,7 @@ static lean_object* _init_l_Lean_Parser_Term_basicFun() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Term_basicFun___closed__15; +x_1 = l_Lean_Parser_Term_basicFun___closed__14; return x_1; } } @@ -28251,7 +28607,7 @@ static lean_object* _init_l_Lean_Parser_Term_fun___closed__6() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_basicFun; -x_2 = l_Lean_Parser_Term_match___closed__9; +x_2 = l_Lean_Parser_Term_match___closed__11; x_3 = l_Lean_Parser_orelse(x_1, x_2); return x_3; } @@ -28434,7 +28790,7 @@ static lean_object* _init_l_Lean_Parser_Term_funStrictImplicitBinder_formatter__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__3; +x_1 = l_Lean_Parser_Term_sufficesDecl_formatter___closed__2; x_2 = l_Lean_Parser_Term_strictImplicitBinder_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); @@ -28446,7 +28802,7 @@ static lean_object* _init_l_Lean_Parser_Term_funStrictImplicitBinder_formatter__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_explicitBinder_formatter___closed__3; +x_1 = l_Lean_Parser_Term_explicitBinder_formatter___closed__2; x_2 = l_Lean_Parser_Term_funStrictImplicitBinder_formatter___closed__1; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -28511,7 +28867,7 @@ static lean_object* _init_l_Lean_Parser_Term_funImplicitBinder_formatter___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__3; +x_1 = l_Lean_Parser_Term_sufficesDecl_formatter___closed__2; x_2 = l_Lean_Parser_Tactic_tacticSeqBracketed_formatter___closed__5; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_orelse_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -28523,7 +28879,7 @@ static lean_object* _init_l_Lean_Parser_Term_funImplicitBinder_formatter___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_explicitBinder_formatter___closed__3; +x_1 = l_Lean_Parser_Term_explicitBinder_formatter___closed__2; x_2 = l_Lean_Parser_Term_funImplicitBinder_formatter___closed__1; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -28618,7 +28974,7 @@ static lean_object* _init_l_Lean_Parser_Term_funBinder_formatter___closed__2() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_Lean_Parser_Term_instBinder_formatter___closed__2; -x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; +x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__5; 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); @@ -28704,7 +29060,7 @@ static lean_object* _init_l_Lean_Parser_Term_basicFun_formatter___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_sufficesDecl_formatter___closed__2; +x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; x_2 = l_Lean_Parser_Term_basicFun_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); @@ -28772,33 +29128,21 @@ static lean_object* _init_l_Lean_Parser_Term_basicFun_formatter___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_sufficesDecl_formatter___closed__2; -x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; -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; -} -} -static lean_object* _init_l_Lean_Parser_Term_basicFun_formatter___closed__10() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_basicFun_formatter___closed__8; -x_2 = l_Lean_Parser_Term_basicFun_formatter___closed__9; +x_2 = l_Lean_Parser_Term_typeAscription_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; } } -static lean_object* _init_l_Lean_Parser_Term_basicFun_formatter___closed__11() { +static lean_object* _init_l_Lean_Parser_Term_basicFun_formatter___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_basicFun___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_basicFun_formatter___closed__10; +x_3 = l_Lean_Parser_Term_basicFun_formatter___closed__9; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -28811,7 +29155,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_basicFun_formatter(lean_object* x_1, { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Term_basicFun_formatter___closed__1; -x_7 = l_Lean_Parser_Term_basicFun_formatter___closed__11; +x_7 = l_Lean_Parser_Term_basicFun_formatter___closed__10; 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; } @@ -28884,7 +29228,7 @@ static lean_object* _init_l_Lean_Parser_Term_fun_formatter___closed__3() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_Lean_Parser_Term_basicFun_formatter___closed__2; -x_2 = l_Lean_Parser_Term_match_formatter___closed__6; +x_2 = l_Lean_Parser_Term_match_formatter___closed__7; 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); @@ -28976,7 +29320,7 @@ static lean_object* _init_l_Lean_Parser_Term_funStrictImplicitBinder_parenthesiz _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__4; +x_1 = l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__2; x_2 = l_Lean_Parser_Term_strictImplicitBinder_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); @@ -28988,7 +29332,7 @@ static lean_object* _init_l_Lean_Parser_Term_funStrictImplicitBinder_parenthesiz _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_explicitBinder_parenthesizer___closed__3; +x_1 = l_Lean_Parser_Term_explicitBinder_parenthesizer___closed__2; x_2 = l_Lean_Parser_Term_funStrictImplicitBinder_parenthesizer___closed__1; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -29053,7 +29397,7 @@ static lean_object* _init_l_Lean_Parser_Term_funImplicitBinder_parenthesizer___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__4; +x_1 = l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__2; x_2 = l_Lean_Parser_Tactic_tacticSeqBracketed_parenthesizer___closed__5; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_orelse_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -29065,7 +29409,7 @@ static lean_object* _init_l_Lean_Parser_Term_funImplicitBinder_parenthesizer___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_explicitBinder_parenthesizer___closed__3; +x_1 = l_Lean_Parser_Term_explicitBinder_parenthesizer___closed__2; x_2 = l_Lean_Parser_Term_funImplicitBinder_parenthesizer___closed__1; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -29246,7 +29590,7 @@ static lean_object* _init_l_Lean_Parser_Term_basicFun_parenthesizer___closed__3( _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__2; +x_1 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__5; x_2 = l_Lean_Parser_Term_basicFun_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); @@ -29314,33 +29658,21 @@ static lean_object* _init_l_Lean_Parser_Term_basicFun_parenthesizer___closed__9( _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__3; -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; -} -} -static lean_object* _init_l_Lean_Parser_Term_basicFun_parenthesizer___closed__10() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_basicFun_parenthesizer___closed__8; -x_2 = l_Lean_Parser_Term_basicFun_parenthesizer___closed__9; +x_2 = l_Lean_Parser_Term_typeAscription_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; } } -static lean_object* _init_l_Lean_Parser_Term_basicFun_parenthesizer___closed__11() { +static lean_object* _init_l_Lean_Parser_Term_basicFun_parenthesizer___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_basicFun___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_basicFun_parenthesizer___closed__10; +x_3 = l_Lean_Parser_Term_basicFun_parenthesizer___closed__9; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -29353,7 +29685,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_basicFun_parenthesizer(lean_object* { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Term_basicFun_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Term_basicFun_parenthesizer___closed__11; +x_7 = l_Lean_Parser_Term_basicFun_parenthesizer___closed__10; 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; } @@ -29426,7 +29758,7 @@ static lean_object* _init_l_Lean_Parser_Term_fun_parenthesizer___closed__3() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_Lean_Parser_Term_basicFun_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Term_match_parenthesizer___closed__6; +x_2 = l_Lean_Parser_Term_match_parenthesizer___closed__7; 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); @@ -29602,6 +29934,23 @@ return x_5; static lean_object* _init_l_Lean_Parser_Term_withAnonymousAntiquot___closed__4() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(" (", 2); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Term_withAnonymousAntiquot___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Term_withAnonymousAntiquot___closed__4; +x_2 = l_Lean_Parser_symbol(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Term_withAnonymousAntiquot___closed__6() { +_start: +{ lean_object* x_1; uint8_t x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_withAnonymousAntiquot___closed__1; x_2 = 0; @@ -29609,72 +29958,82 @@ x_3 = l_Lean_Parser_nonReservedSymbol(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_withAnonymousAntiquot___closed__5() { +static lean_object* _init_l_Lean_Parser_Term_withAnonymousAntiquot___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_withAnonymousAntiquot___closed__4; +x_1 = l_Lean_Parser_Term_withAnonymousAntiquot___closed__6; x_2 = l_Lean_Parser_Term_structInstField___closed__5; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_withAnonymousAntiquot___closed__6() { +static lean_object* _init_l_Lean_Parser_Term_withAnonymousAntiquot___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_typeAscription___closed__5; -x_2 = l_Lean_Parser_Term_withAnonymousAntiquot___closed__5; +x_1 = l_Lean_Parser_Term_withAnonymousAntiquot___closed__5; +x_2 = l_Lean_Parser_Term_withAnonymousAntiquot___closed__7; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_withAnonymousAntiquot___closed__7() { +static lean_object* _init_l_Lean_Parser_Term_withAnonymousAntiquot___closed__9() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_withAnonymousAntiquot___closed__6; +x_1 = l_Lean_Parser_Term_withAnonymousAntiquot___closed__8; x_2 = l_Lean_Parser_atomic(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_withAnonymousAntiquot___closed__8() { +static lean_object* _init_l_Lean_Parser_Term_withAnonymousAntiquot___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_withAnonymousAntiquot___closed__7; -x_2 = l_Lean_Parser_Term_generalizingParam___closed__10; +x_1 = l_Lean_Parser_Term_generalizingParam___closed__8; +x_2 = l_Lean_Parser_Term_typeAscription___closed__18; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_withAnonymousAntiquot___closed__9() { +static lean_object* _init_l_Lean_Parser_Term_withAnonymousAntiquot___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_withAnonymousAntiquot___closed__9; +x_2 = l_Lean_Parser_Term_withAnonymousAntiquot___closed__10; +x_3 = l_Lean_Parser_andthen(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Term_withAnonymousAntiquot___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_withAnonymousAntiquot___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_withAnonymousAntiquot___closed__8; +x_3 = l_Lean_Parser_Term_withAnonymousAntiquot___closed__11; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Term_withAnonymousAntiquot___closed__10() { +static lean_object* _init_l_Lean_Parser_Term_withAnonymousAntiquot___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_withAnonymousAntiquot___closed__3; -x_2 = l_Lean_Parser_Term_withAnonymousAntiquot___closed__9; +x_2 = l_Lean_Parser_Term_withAnonymousAntiquot___closed__12; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_withAnonymousAntiquot___closed__11() { +static lean_object* _init_l_Lean_Parser_Term_withAnonymousAntiquot___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_withAnonymousAntiquot___closed__2; -x_2 = l_Lean_Parser_Term_withAnonymousAntiquot___closed__10; +x_2 = l_Lean_Parser_Term_withAnonymousAntiquot___closed__13; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -29683,7 +30042,7 @@ static lean_object* _init_l_Lean_Parser_Term_withAnonymousAntiquot() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Term_withAnonymousAntiquot___closed__11; +x_1 = l_Lean_Parser_Term_withAnonymousAntiquot___closed__14; return x_1; } } @@ -29722,21 +30081,13 @@ return x_5; static lean_object* _init_l_Lean_Parser_Term_leading__parser___closed__4() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("leading_parser ", 15); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Term_leading__parser___closed__5() { -_start: -{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_leading__parser___closed__4; +x_1 = l_Lean_Parser_Term_leading__parser___closed__1; x_2 = l_Lean_Parser_symbol(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_leading__parser___closed__6() { +static lean_object* _init_l_Lean_Parser_Term_leading__parser___closed__5() { _start: { lean_object* x_1; lean_object* x_2; @@ -29745,63 +30096,63 @@ x_2 = l_Lean_Parser_optional(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_leading__parser___closed__7() { +static lean_object* _init_l_Lean_Parser_Term_leading__parser___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_leading__parser___closed__6; -x_2 = l_Lean_Parser_Term_typeAscription___closed__6; +x_1 = l_Lean_Parser_Term_leading__parser___closed__5; +x_2 = l_Lean_Parser_Term_typeAscription___closed__9; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_leading__parser___closed__8() { +static lean_object* _init_l_Lean_Parser_Term_leading__parser___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_optExprPrecedence; -x_2 = l_Lean_Parser_Term_leading__parser___closed__7; +x_2 = l_Lean_Parser_Term_leading__parser___closed__6; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_leading__parser___closed__9() { +static lean_object* _init_l_Lean_Parser_Term_leading__parser___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_leading__parser___closed__5; -x_2 = l_Lean_Parser_Term_leading__parser___closed__8; +x_1 = l_Lean_Parser_Term_leading__parser___closed__4; +x_2 = l_Lean_Parser_Term_leading__parser___closed__7; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_leading__parser___closed__10() { +static lean_object* _init_l_Lean_Parser_Term_leading__parser___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_leading__parser___closed__2; x_2 = l_Lean_Parser_leadPrec; -x_3 = l_Lean_Parser_Term_leading__parser___closed__9; +x_3 = l_Lean_Parser_Term_leading__parser___closed__8; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Term_leading__parser___closed__11() { +static lean_object* _init_l_Lean_Parser_Term_leading__parser___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_leading__parser___closed__3; -x_2 = l_Lean_Parser_Term_leading__parser___closed__10; +x_2 = l_Lean_Parser_Term_leading__parser___closed__9; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_leading__parser___closed__12() { +static lean_object* _init_l_Lean_Parser_Term_leading__parser___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_leading__parser___closed__2; -x_2 = l_Lean_Parser_Term_leading__parser___closed__11; +x_2 = l_Lean_Parser_Term_leading__parser___closed__10; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -29810,7 +30161,7 @@ static lean_object* _init_l_Lean_Parser_Term_leading__parser() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Term_leading__parser___closed__12; +x_1 = l_Lean_Parser_Term_leading__parser___closed__11; return x_1; } } @@ -29844,7 +30195,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_leading__parser_declRa { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_unsigned_to_nat(375u); -x_2 = lean_unsigned_to_nat(88u); +x_2 = lean_unsigned_to_nat(98u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -29858,7 +30209,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Parser_Term_leading__parser_declRange___closed__1; x_2 = lean_unsigned_to_nat(23u); x_3 = l___regBuiltin_Lean_Parser_Term_leading__parser_declRange___closed__2; -x_4 = lean_unsigned_to_nat(88u); +x_4 = lean_unsigned_to_nat(98u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -29954,7 +30305,7 @@ static lean_object* _init_l_Lean_Parser_Term_optExprPrecedence_formatter___close { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_optExprPrecedence_formatter___closed__2; -x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; +x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__5; 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); @@ -29991,6 +30342,16 @@ return x_7; static lean_object* _init_l_Lean_Parser_Term_withAnonymousAntiquot_formatter___closed__2() { _start: { +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Term_withAnonymousAntiquot___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; +} +} +static lean_object* _init_l_Lean_Parser_Term_withAnonymousAntiquot_formatter___closed__3() { +_start: +{ lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_withAnonymousAntiquot___closed__1; x_2 = 0; @@ -30001,11 +30362,11 @@ lean_closure_set(x_4, 1, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Term_withAnonymousAntiquot_formatter___closed__3() { +static lean_object* _init_l_Lean_Parser_Term_withAnonymousAntiquot_formatter___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_withAnonymousAntiquot_formatter___closed__2; +x_1 = l_Lean_Parser_Term_withAnonymousAntiquot_formatter___closed__3; x_2 = l_Lean_Parser_Term_structInstField_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); @@ -30013,47 +30374,59 @@ lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_withAnonymousAntiquot_formatter___closed__4() { +static lean_object* _init_l_Lean_Parser_Term_withAnonymousAntiquot_formatter___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__2; -x_2 = l_Lean_Parser_Term_withAnonymousAntiquot_formatter___closed__3; +x_1 = l_Lean_Parser_Term_withAnonymousAntiquot_formatter___closed__2; +x_2 = l_Lean_Parser_Term_withAnonymousAntiquot_formatter___closed__4; 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; } } -static lean_object* _init_l_Lean_Parser_Term_withAnonymousAntiquot_formatter___closed__5() { +static lean_object* _init_l_Lean_Parser_Term_withAnonymousAntiquot_formatter___closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_withAnonymousAntiquot_formatter___closed__4; +x_1 = l_Lean_Parser_Term_withAnonymousAntiquot_formatter___closed__5; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_atomic_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_withAnonymousAntiquot_formatter___closed__6() { +static lean_object* _init_l_Lean_Parser_Term_withAnonymousAntiquot_formatter___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_withAnonymousAntiquot_formatter___closed__5; -x_2 = l_Lean_Parser_Term_generalizingParam_formatter___closed__7; +x_1 = l_Lean_Parser_Term_generalizingParam_formatter___closed__5; +x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__12; 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; } } -static lean_object* _init_l_Lean_Parser_Term_withAnonymousAntiquot_formatter___closed__7() { +static lean_object* _init_l_Lean_Parser_Term_withAnonymousAntiquot_formatter___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_withAnonymousAntiquot_formatter___closed__6; +x_2 = l_Lean_Parser_Term_withAnonymousAntiquot_formatter___closed__7; +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; +} +} +static lean_object* _init_l_Lean_Parser_Term_withAnonymousAntiquot_formatter___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_withAnonymousAntiquot___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_withAnonymousAntiquot_formatter___closed__6; +x_3 = l_Lean_Parser_Term_withAnonymousAntiquot_formatter___closed__8; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -30066,7 +30439,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_withAnonymousAntiquot_formatter(lean { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Term_withAnonymousAntiquot_formatter___closed__1; -x_7 = l_Lean_Parser_Term_withAnonymousAntiquot_formatter___closed__7; +x_7 = l_Lean_Parser_Term_withAnonymousAntiquot_formatter___closed__9; 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; } @@ -30126,7 +30499,7 @@ static lean_object* _init_l_Lean_Parser_Term_leading__parser_formatter___closed_ _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_leading__parser___closed__4; +x_1 = l_Lean_Parser_Term_leading__parser___closed__1; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; @@ -30147,7 +30520,7 @@ static lean_object* _init_l_Lean_Parser_Term_leading__parser_formatter___closed_ { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_leading__parser_formatter___closed__3; -x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; +x_2 = l_Lean_Parser_Term_typeAscription_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); @@ -30305,6 +30678,16 @@ return x_7; static lean_object* _init_l_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___closed__2() { _start: { +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Term_withAnonymousAntiquot___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; +} +} +static lean_object* _init_l_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___closed__3() { +_start: +{ lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_withAnonymousAntiquot___closed__1; x_2 = 0; @@ -30315,11 +30698,11 @@ lean_closure_set(x_4, 1, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___closed__3() { +static lean_object* _init_l_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___closed__2; +x_1 = l_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___closed__3; x_2 = l_Lean_Parser_Term_structInstField_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); @@ -30327,47 +30710,59 @@ lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___closed__4() { +static lean_object* _init_l_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___closed__3; +x_1 = l_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___closed__2; +x_2 = l_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___closed__4; 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; } } -static lean_object* _init_l_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___closed__5() { +static lean_object* _init_l_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___closed__4; +x_1 = l_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___closed__5; x_2 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_atomic_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___closed__6() { +static lean_object* _init_l_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___closed__5; -x_2 = l_Lean_Parser_Term_generalizingParam_parenthesizer___closed__7; +x_1 = l_Lean_Parser_Term_generalizingParam_parenthesizer___closed__5; +x_2 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__12; 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; } } -static lean_object* _init_l_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___closed__7() { +static lean_object* _init_l_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___closed__6; +x_2 = l_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___closed__7; +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; +} +} +static lean_object* _init_l_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_withAnonymousAntiquot___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___closed__6; +x_3 = l_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___closed__8; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -30380,7 +30775,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer( { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___closed__7; +x_7 = l_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___closed__9; 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; } @@ -30440,7 +30835,7 @@ static lean_object* _init_l_Lean_Parser_Term_leading__parser_parenthesizer___clo _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_leading__parser___closed__4; +x_1 = l_Lean_Parser_Term_leading__parser___closed__1; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; @@ -30461,7 +30856,7 @@ static lean_object* _init_l_Lean_Parser_Term_leading__parser_parenthesizer___clo { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_leading__parser_parenthesizer___closed__3; -x_2 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__3; +x_2 = l_Lean_Parser_Term_typeAscription_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); @@ -30592,77 +30987,69 @@ return x_5; static lean_object* _init_l_Lean_Parser_Term_trailing__parser___closed__4() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("trailing_parser ", 16); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Term_trailing__parser___closed__5() { -_start: -{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_trailing__parser___closed__4; +x_1 = l_Lean_Parser_Term_trailing__parser___closed__1; x_2 = l_Lean_Parser_symbol(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_trailing__parser___closed__6() { +static lean_object* _init_l_Lean_Parser_Term_trailing__parser___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_optExprPrecedence; -x_2 = l_Lean_Parser_Term_typeAscription___closed__6; +x_2 = l_Lean_Parser_Term_typeAscription___closed__9; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_trailing__parser___closed__7() { +static lean_object* _init_l_Lean_Parser_Term_trailing__parser___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_optExprPrecedence; -x_2 = l_Lean_Parser_Term_trailing__parser___closed__6; +x_2 = l_Lean_Parser_Term_trailing__parser___closed__5; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_trailing__parser___closed__8() { +static lean_object* _init_l_Lean_Parser_Term_trailing__parser___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_trailing__parser___closed__5; -x_2 = l_Lean_Parser_Term_trailing__parser___closed__7; +x_1 = l_Lean_Parser_Term_trailing__parser___closed__4; +x_2 = l_Lean_Parser_Term_trailing__parser___closed__6; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_trailing__parser___closed__9() { +static lean_object* _init_l_Lean_Parser_Term_trailing__parser___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_trailing__parser___closed__2; x_2 = l_Lean_Parser_leadPrec; -x_3 = l_Lean_Parser_Term_trailing__parser___closed__8; +x_3 = l_Lean_Parser_Term_trailing__parser___closed__7; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Term_trailing__parser___closed__10() { +static lean_object* _init_l_Lean_Parser_Term_trailing__parser___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_trailing__parser___closed__3; -x_2 = l_Lean_Parser_Term_trailing__parser___closed__9; +x_2 = l_Lean_Parser_Term_trailing__parser___closed__8; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_trailing__parser___closed__11() { +static lean_object* _init_l_Lean_Parser_Term_trailing__parser___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_trailing__parser___closed__2; -x_2 = l_Lean_Parser_Term_trailing__parser___closed__10; +x_2 = l_Lean_Parser_Term_trailing__parser___closed__9; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -30671,7 +31058,7 @@ static lean_object* _init_l_Lean_Parser_Term_trailing__parser() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Term_trailing__parser___closed__11; +x_1 = l_Lean_Parser_Term_trailing__parser___closed__10; return x_1; } } @@ -30705,7 +31092,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_trailing__parser_declR { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_unsigned_to_nat(377u); -x_2 = lean_unsigned_to_nat(76u); +x_2 = lean_unsigned_to_nat(86u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -30719,7 +31106,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Parser_Term_trailing__parser_declRange___closed__1; x_2 = lean_unsigned_to_nat(23u); x_3 = l___regBuiltin_Lean_Parser_Term_trailing__parser_declRange___closed__2; -x_4 = lean_unsigned_to_nat(76u); +x_4 = lean_unsigned_to_nat(86u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -30812,7 +31199,7 @@ static lean_object* _init_l_Lean_Parser_Term_trailing__parser_formatter___closed _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_trailing__parser___closed__4; +x_1 = l_Lean_Parser_Term_trailing__parser___closed__1; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; @@ -30823,7 +31210,7 @@ static lean_object* _init_l_Lean_Parser_Term_trailing__parser_formatter___closed { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_leading__parser_formatter___closed__5; -x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; +x_2 = l_Lean_Parser_Term_typeAscription_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); @@ -30933,7 +31320,7 @@ static lean_object* _init_l_Lean_Parser_Term_trailing__parser_parenthesizer___cl _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_trailing__parser___closed__4; +x_1 = l_Lean_Parser_Term_trailing__parser___closed__1; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; @@ -30944,7 +31331,7 @@ static lean_object* _init_l_Lean_Parser_Term_trailing__parser_parenthesizer___cl { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_leading__parser_parenthesizer___closed__5; -x_2 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__3; +x_2 = l_Lean_Parser_Term_typeAscription_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); @@ -31287,7 +31674,7 @@ static lean_object* _init_l_Lean_Parser_Term_borrowed_formatter___closed__3() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_borrowed_formatter___closed__2; -x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; +x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__5; 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); @@ -32531,7 +32918,7 @@ static lean_object* _init_l_Lean_Parser_Term_letIdLhs___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_ident; +x_1 = l_Lean_Parser_Term_binderIdent; x_2 = l_Lean_Parser_Term_letIdLhs___closed__8; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; @@ -32799,7 +33186,7 @@ static lean_object* _init_l_Lean_Parser_Term_letEqnsDecl___closed__4() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_structInstField___closed__5; -x_2 = l_Lean_Parser_Term_match___closed__9; +x_2 = l_Lean_Parser_Term_match___closed__11; x_3 = l_Lean_Parser_orelse(x_1, x_2); return x_3; } @@ -33292,7 +33679,7 @@ static lean_object* _init_l_Lean_Parser_Term_letIdLhs_formatter___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_sufficesDecl_formatter___closed__2; +x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; x_2 = l_Lean_Parser_Term_letIdLhs_formatter___closed__3; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -33338,7 +33725,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_letIdLhs_formatter(lean_object* x_1, _start: { lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_6 = l_Lean_Parser_Term_syntheticHole_formatter___closed__3; +x_6 = l_Lean_Parser_Term_sufficesDecl_formatter___closed__3; x_7 = l_Lean_Parser_Term_letIdLhs_formatter___closed__7; 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; @@ -33396,7 +33783,7 @@ static lean_object* _init_l_Lean_Parser_Term_letIdDecl_formatter___closed__5() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_letIdDecl_formatter___closed__4; -x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; +x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__5; 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); @@ -33505,7 +33892,7 @@ static lean_object* _init_l_Lean_Parser_Term_letPatDecl_formatter___closed__4() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; +x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__5; x_2 = l_Lean_Parser_Term_letPatDecl_formatter___closed__3; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -33528,7 +33915,7 @@ static lean_object* _init_l_Lean_Parser_Term_letPatDecl_formatter___closed__6() { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_letPatDecl_formatter___closed__5; -x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; +x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__5; 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); @@ -33614,7 +34001,7 @@ static lean_object* _init_l_Lean_Parser_Term_letEqnsDecl_formatter___closed__2() { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_structInstField_formatter___closed__2; -x_2 = l_Lean_Parser_Term_match_formatter___closed__6; +x_2 = l_Lean_Parser_Term_match_formatter___closed__7; 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); @@ -34003,7 +34390,7 @@ static lean_object* _init_l_Lean_Parser_Term_letIdLhs_parenthesizer___closed__4( _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__2; +x_1 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__5; x_2 = l_Lean_Parser_Term_letIdLhs_parenthesizer___closed__3; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -34049,7 +34436,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_letIdLhs_parenthesizer(lean_object* _start: { lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_6 = l_Lean_Parser_Term_syntheticHole_parenthesizer___closed__3; +x_6 = l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__3; x_7 = l_Lean_Parser_Term_letIdLhs_parenthesizer___closed__7; x_8 = l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer(x_6, x_7, x_1, x_2, x_3, x_4, x_5); return x_8; @@ -34325,7 +34712,7 @@ static lean_object* _init_l_Lean_Parser_Term_letEqnsDecl_parenthesizer___closed_ { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_structInstField_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Term_match_parenthesizer___closed__6; +x_2 = l_Lean_Parser_Term_match_parenthesizer___closed__7; 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); @@ -34706,7 +35093,7 @@ static lean_object* _init_l_Lean_Parser_Term_let__fun___closed__6() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("let_λ", 6); +x_1 = lean_mk_string_from_bytes("let_λ ", 7); return x_1; } } @@ -34845,7 +35232,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_let__fun_declRange___c { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_unsigned_to_nat(456u); -x_2 = lean_unsigned_to_nat(86u); +x_2 = lean_unsigned_to_nat(87u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -34859,7 +35246,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Parser_Term_let__fun_declRange___closed__1; x_2 = lean_unsigned_to_nat(23u); x_3 = l___regBuiltin_Lean_Parser_Term_let__fun_declRange___closed__2; -x_4 = lean_unsigned_to_nat(86u); +x_4 = lean_unsigned_to_nat(87u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -36192,8 +36579,8 @@ static lean_object* _init_l_Lean_Parser_Term_haveIdLhs___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_ident; -x_2 = l_Lean_Parser_Term_letIdLhs___closed__6; +x_1 = l_Lean_Parser_skip; +x_2 = l_Lean_Parser_Term_binderIdent; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -36201,10 +36588,11 @@ return x_3; static lean_object* _init_l_Lean_Parser_Term_haveIdLhs___closed__2() { _start: { -lean_object* x_1; lean_object* x_2; +lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_haveIdLhs___closed__1; -x_2 = l_Lean_Parser_optional(x_1); -return x_2; +x_2 = l_Lean_Parser_hygieneInfo; +x_3 = l_Lean_Parser_orelse(x_1, x_2); +return x_3; } } static lean_object* _init_l_Lean_Parser_Term_haveIdLhs___closed__3() { @@ -36212,7 +36600,7 @@ static lean_object* _init_l_Lean_Parser_Term_haveIdLhs___closed__3() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_haveIdLhs___closed__2; -x_2 = l_Lean_Parser_Term_optType; +x_2 = l_Lean_Parser_Term_letIdLhs___closed__7; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -36360,7 +36748,7 @@ static lean_object* _init_l_Lean_Parser_Term_haveEqnsDecl___closed__4() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_haveIdLhs; -x_2 = l_Lean_Parser_Term_match___closed__9; +x_2 = l_Lean_Parser_Term_match___closed__11; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -36439,49 +36827,59 @@ static lean_object* _init_l_Lean_Parser_Term_haveDecl___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_letPatDecl; +x_1 = l_Lean_Parser_skip; +x_2 = l_Lean_Parser_Term_letPatDecl; +x_3 = l_Lean_Parser_andthen(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Term_haveDecl___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_haveDecl___closed__4; x_2 = l_Lean_Parser_Term_haveEqnsDecl; x_3 = l_Lean_Parser_orelse(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_haveDecl___closed__5() { +static lean_object* _init_l_Lean_Parser_Term_haveDecl___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_haveIdDecl; -x_2 = l_Lean_Parser_Term_haveDecl___closed__4; +x_2 = l_Lean_Parser_Term_haveDecl___closed__5; x_3 = l_Lean_Parser_orelse(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_haveDecl___closed__6() { +static lean_object* _init_l_Lean_Parser_Term_haveDecl___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_haveDecl___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_haveDecl___closed__5; +x_3 = l_Lean_Parser_Term_haveDecl___closed__6; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Term_haveDecl___closed__7() { +static lean_object* _init_l_Lean_Parser_Term_haveDecl___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_haveDecl___closed__3; -x_2 = l_Lean_Parser_Term_haveDecl___closed__6; +x_2 = l_Lean_Parser_Term_haveDecl___closed__7; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_haveDecl___closed__8() { +static lean_object* _init_l_Lean_Parser_Term_haveDecl___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_haveDecl___closed__2; -x_2 = l_Lean_Parser_Term_haveDecl___closed__7; +x_2 = l_Lean_Parser_Term_haveDecl___closed__8; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -36490,7 +36888,7 @@ static lean_object* _init_l_Lean_Parser_Term_haveDecl() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Term_haveDecl___closed__8; +x_1 = l_Lean_Parser_Term_haveDecl___closed__9; return x_1; } } @@ -36529,76 +36927,68 @@ return x_5; static lean_object* _init_l_Lean_Parser_Term_have___closed__4() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("have ", 5); -return x_1; -} -} -static lean_object* _init_l_Lean_Parser_Term_have___closed__5() { -_start: -{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_have___closed__4; +x_1 = l_Lean_Parser_Term_have___closed__1; x_2 = l_Lean_Parser_symbol(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_have___closed__6() { +static lean_object* _init_l_Lean_Parser_Term_have___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_have___closed__5; +x_1 = l_Lean_Parser_Term_have___closed__4; x_2 = l_Lean_Parser_Term_haveDecl; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_have___closed__7() { +static lean_object* _init_l_Lean_Parser_Term_have___closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_have___closed__6; +x_1 = l_Lean_Parser_Term_have___closed__5; x_2 = l_Lean_Parser_withPosition(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_have___closed__8() { +static lean_object* _init_l_Lean_Parser_Term_have___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_have___closed__7; +x_1 = l_Lean_Parser_Term_have___closed__6; x_2 = l_Lean_Parser_Term_suffices___closed__8; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_have___closed__9() { +static lean_object* _init_l_Lean_Parser_Term_have___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_have___closed__2; x_2 = l_Lean_Parser_leadPrec; -x_3 = l_Lean_Parser_Term_have___closed__8; +x_3 = l_Lean_Parser_Term_have___closed__7; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Term_have___closed__10() { +static lean_object* _init_l_Lean_Parser_Term_have___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_have___closed__3; -x_2 = l_Lean_Parser_Term_have___closed__9; +x_2 = l_Lean_Parser_Term_have___closed__8; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_have___closed__11() { +static lean_object* _init_l_Lean_Parser_Term_have___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_have___closed__2; -x_2 = l_Lean_Parser_Term_have___closed__10; +x_2 = l_Lean_Parser_Term_have___closed__9; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -36607,7 +36997,7 @@ static lean_object* _init_l_Lean_Parser_Term_have() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Term_have___closed__11; +x_1 = l_Lean_Parser_Term_have___closed__10; return x_1; } } @@ -36641,7 +37031,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_have_declRange___close { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_unsigned_to_nat(478u); -x_2 = lean_unsigned_to_nat(63u); +x_2 = lean_unsigned_to_nat(62u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -36655,7 +37045,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Parser_Term_have_declRange___closed__1; x_2 = lean_unsigned_to_nat(23u); x_3 = l___regBuiltin_Lean_Parser_Term_have_declRange___closed__2; -x_4 = lean_unsigned_to_nat(63u); +x_4 = lean_unsigned_to_nat(62u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -36730,8 +37120,8 @@ static lean_object* _init_l_Lean_Parser_Term_haveIdLhs_formatter___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_syntheticHole_formatter___closed__3; -x_2 = l_Lean_Parser_Term_letIdLhs_formatter___closed__5; +x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; +x_2 = l_Lean_Parser_Term_sufficesDecl_formatter___closed__3; 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); @@ -36741,11 +37131,13 @@ return x_3; static lean_object* _init_l_Lean_Parser_Term_haveIdLhs_formatter___closed__2() { _start: { -lean_object* x_1; lean_object* x_2; +lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_haveIdLhs_formatter___closed__1; -x_2 = lean_alloc_closure((void*)(l_Lean_Parser_optional_formatter), 6, 1); -lean_closure_set(x_2, 0, x_1); -return x_2; +x_2 = l_Lean_Parser_Term_sufficesDecl_formatter___closed__7; +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_haveIdLhs_formatter(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { @@ -36753,7 +37145,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_haveIdLhs_formatter(lean_object* x_1 { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Term_haveIdLhs_formatter___closed__2; -x_7 = l_Lean_Parser_Term_forall_formatter___closed__8; +x_7 = l_Lean_Parser_Term_letIdLhs_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); return x_8; } @@ -36810,7 +37202,7 @@ static lean_object* _init_l_Lean_Parser_Term_haveIdDecl_formatter___closed__5() { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_haveIdDecl_formatter___closed__4; -x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; +x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__5; 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); @@ -36896,7 +37288,7 @@ static lean_object* _init_l_Lean_Parser_Term_haveEqnsDecl_formatter___closed__2( { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_haveIdDecl_formatter___closed__2; -x_2 = l_Lean_Parser_Term_match_formatter___closed__6; +x_2 = l_Lean_Parser_Term_match_formatter___closed__7; 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); @@ -36981,7 +37373,19 @@ static lean_object* _init_l_Lean_Parser_Term_haveDecl_formatter___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Parser_Term_letPatDecl_formatter___closed__2; +x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; +x_2 = l___regBuiltin_Lean_Parser_Term_letPatDecl_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; +} +} +static lean_object* _init_l_Lean_Parser_Term_haveDecl_formatter___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_haveDecl_formatter___closed__2; x_2 = l___regBuiltin_Lean_Parser_Term_haveEqnsDecl_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); @@ -36989,25 +37393,25 @@ lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_haveDecl_formatter___closed__3() { +static lean_object* _init_l_Lean_Parser_Term_haveDecl_formatter___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_Lean_Parser_Term_haveIdDecl_formatter___closed__2; -x_2 = l_Lean_Parser_Term_haveDecl_formatter___closed__2; +x_2 = l_Lean_Parser_Term_haveDecl_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_haveDecl_formatter___closed__4() { +static lean_object* _init_l_Lean_Parser_Term_haveDecl_formatter___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_haveDecl___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_haveDecl_formatter___closed__3; +x_3 = l_Lean_Parser_Term_haveDecl_formatter___closed__4; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -37020,7 +37424,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_haveDecl_formatter(lean_object* x_1, { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Term_haveDecl_formatter___closed__1; -x_7 = l_Lean_Parser_Term_haveDecl_formatter___closed__4; +x_7 = l_Lean_Parser_Term_haveDecl_formatter___closed__5; 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; } @@ -37080,7 +37484,7 @@ static lean_object* _init_l_Lean_Parser_Term_have_formatter___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_have___closed__4; +x_1 = l_Lean_Parser_Term_have___closed__1; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; @@ -37181,8 +37585,8 @@ static lean_object* _init_l_Lean_Parser_Term_haveIdLhs_parenthesizer___closed__1 _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_syntheticHole_parenthesizer___closed__3; -x_2 = l_Lean_Parser_Term_letIdLhs_parenthesizer___closed__5; +x_1 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__5; +x_2 = l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__3; 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); @@ -37192,11 +37596,13 @@ return x_3; static lean_object* _init_l_Lean_Parser_Term_haveIdLhs_parenthesizer___closed__2() { _start: { -lean_object* x_1; lean_object* x_2; +lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_haveIdLhs_parenthesizer___closed__1; -x_2 = lean_alloc_closure((void*)(l_Lean_Parser_optional_parenthesizer), 6, 1); -lean_closure_set(x_2, 0, x_1); -return x_2; +x_2 = l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__7; +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_haveIdLhs_parenthesizer(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { @@ -37204,7 +37610,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_haveIdLhs_parenthesizer(lean_object* { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Term_haveIdLhs_parenthesizer___closed__2; -x_7 = l_Lean_Parser_Term_forall_parenthesizer___closed__8; +x_7 = l_Lean_Parser_Term_letIdLhs_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); return x_8; } @@ -37347,7 +37753,7 @@ static lean_object* _init_l_Lean_Parser_Term_haveEqnsDecl_parenthesizer___closed { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_haveIdDecl_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Term_match_parenthesizer___closed__6; +x_2 = l_Lean_Parser_Term_match_parenthesizer___closed__7; 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); @@ -37432,7 +37838,19 @@ static lean_object* _init_l_Lean_Parser_Term_haveDecl_parenthesizer___closed__2( _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___regBuiltin_Lean_Parser_Term_letPatDecl_parenthesizer___closed__2; +x_1 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__5; +x_2 = l___regBuiltin_Lean_Parser_Term_letPatDecl_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; +} +} +static lean_object* _init_l_Lean_Parser_Term_haveDecl_parenthesizer___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_haveDecl_parenthesizer___closed__2; x_2 = l___regBuiltin_Lean_Parser_Term_haveEqnsDecl_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); @@ -37440,25 +37858,25 @@ lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_haveDecl_parenthesizer___closed__3() { +static lean_object* _init_l_Lean_Parser_Term_haveDecl_parenthesizer___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_Lean_Parser_Term_haveIdDecl_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Term_haveDecl_parenthesizer___closed__2; +x_2 = l_Lean_Parser_Term_haveDecl_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_haveDecl_parenthesizer___closed__4() { +static lean_object* _init_l_Lean_Parser_Term_haveDecl_parenthesizer___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_haveDecl___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_haveDecl_parenthesizer___closed__3; +x_3 = l_Lean_Parser_Term_haveDecl_parenthesizer___closed__4; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -37471,7 +37889,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_haveDecl_parenthesizer(lean_object* { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Term_haveDecl_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Term_haveDecl_parenthesizer___closed__4; +x_7 = l_Lean_Parser_Term_haveDecl_parenthesizer___closed__5; 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; } @@ -37531,7 +37949,7 @@ static lean_object* _init_l_Lean_Parser_Term_have_parenthesizer___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_have___closed__4; +x_1 = l_Lean_Parser_Term_have___closed__1; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; @@ -38068,7 +38486,7 @@ static lean_object* _init_l_Lean_Parser_Term_attributes___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_typeAscription___closed__14; +x_1 = l_Lean_Parser_Term_typeAscription___closed__15; x_2 = l_Lean_Parser_Term_attributes___closed__6; x_3 = l_Lean_Parser_adaptCacheableContext(x_1, x_2); return x_3; @@ -38077,50 +38495,67 @@ return x_3; static lean_object* _init_l_Lean_Parser_Term_attributes___closed__8() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("] ", 2); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Term_attributes___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Term_attributes___closed__8; +x_2 = l_Lean_Parser_symbol(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Term_attributes___closed__10() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_attributes___closed__7; -x_2 = l_Lean_Parser_Term_structInstArrayRef___closed__8; +x_2 = l_Lean_Parser_Term_attributes___closed__9; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_attributes___closed__9() { +static lean_object* _init_l_Lean_Parser_Term_attributes___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_attributes___closed__5; -x_2 = l_Lean_Parser_Term_attributes___closed__8; +x_2 = l_Lean_Parser_Term_attributes___closed__10; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_attributes___closed__10() { +static lean_object* _init_l_Lean_Parser_Term_attributes___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_attributes___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_attributes___closed__9; +x_3 = l_Lean_Parser_Term_attributes___closed__11; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Term_attributes___closed__11() { +static lean_object* _init_l_Lean_Parser_Term_attributes___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_attributes___closed__3; -x_2 = l_Lean_Parser_Term_attributes___closed__10; +x_2 = l_Lean_Parser_Term_attributes___closed__12; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_attributes___closed__12() { +static lean_object* _init_l_Lean_Parser_Term_attributes___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_attributes___closed__2; -x_2 = l_Lean_Parser_Term_attributes___closed__11; +x_2 = l_Lean_Parser_Term_attributes___closed__13; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -38129,7 +38564,7 @@ static lean_object* _init_l_Lean_Parser_Term_attributes() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Term_attributes___closed__12; +x_1 = l_Lean_Parser_Term_attributes___closed__14; return x_1; } } @@ -38389,7 +38824,7 @@ static lean_object* _init_l_Lean_Parser_Term_letrec___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_structInstLVal___closed__9; +x_1 = l_Lean_Parser_Term_sufficesDecl___closed__6; x_2 = l_Lean_Parser_Term_letrec___closed__6; x_3 = l_Lean_Parser_node(x_1, x_2); return x_3; @@ -38630,7 +39065,7 @@ static lean_object* _init_l_Lean_Parser_Command_docComment_formatter___closed__5 _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_sufficesDecl_formatter___closed__2; +x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; x_2 = l_Lean_Parser_Command_docComment_formatter___closed__4; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -39158,34 +39593,44 @@ return x_2; static lean_object* _init_l_Lean_Parser_Term_attributes_formatter___closed__5() { _start: { +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Term_attributes___closed__8; +x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_formatter), 6, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Term_attributes_formatter___closed__6() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_attributes_formatter___closed__4; -x_2 = l_Lean_Parser_Term_structInstArrayRef_formatter___closed__3; +x_2 = l_Lean_Parser_Term_attributes_formatter___closed__5; 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; } } -static lean_object* _init_l_Lean_Parser_Term_attributes_formatter___closed__6() { +static lean_object* _init_l_Lean_Parser_Term_attributes_formatter___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_attributes_formatter___closed__2; -x_2 = l_Lean_Parser_Term_attributes_formatter___closed__5; +x_2 = l_Lean_Parser_Term_attributes_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; } } -static lean_object* _init_l_Lean_Parser_Term_attributes_formatter___closed__7() { +static lean_object* _init_l_Lean_Parser_Term_attributes_formatter___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_attributes___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_attributes_formatter___closed__6; +x_3 = l_Lean_Parser_Term_attributes_formatter___closed__7; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -39198,7 +39643,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_attributes_formatter(lean_object* x_ { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Term_attributes_formatter___closed__1; -x_7 = l_Lean_Parser_Term_attributes_formatter___closed__7; +x_7 = l_Lean_Parser_Term_attributes_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; } @@ -39643,7 +40088,7 @@ static lean_object* _init_l_Lean_Parser_Command_docComment_parenthesizer___close _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__2; +x_1 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__5; x_2 = l_Lean_Parser_Command_docComment_parenthesizer___closed__4; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -40152,34 +40597,44 @@ return x_2; static lean_object* _init_l_Lean_Parser_Term_attributes_parenthesizer___closed__5() { _start: { +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Term_attributes___closed__8; +x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_parenthesizer), 6, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Term_attributes_parenthesizer___closed__6() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_attributes_parenthesizer___closed__4; -x_2 = l_Lean_Parser_Term_structInstArrayRef_parenthesizer___closed__3; +x_2 = l_Lean_Parser_Term_attributes_parenthesizer___closed__5; 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; } } -static lean_object* _init_l_Lean_Parser_Term_attributes_parenthesizer___closed__6() { +static lean_object* _init_l_Lean_Parser_Term_attributes_parenthesizer___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_attributes_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Term_attributes_parenthesizer___closed__5; +x_2 = l_Lean_Parser_Term_attributes_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; } } -static lean_object* _init_l_Lean_Parser_Term_attributes_parenthesizer___closed__7() { +static lean_object* _init_l_Lean_Parser_Term_attributes_parenthesizer___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_attributes___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_attributes_parenthesizer___closed__6; +x_3 = l_Lean_Parser_Term_attributes_parenthesizer___closed__7; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -40192,7 +40647,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_attributes_parenthesizer(lean_object { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Term_attributes_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Term_attributes_parenthesizer___closed__7; +x_7 = l_Lean_Parser_Term_attributes_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; } @@ -40626,22 +41081,32 @@ return x_7; static lean_object* _init_l_Lean_Parser_Term_whereDecls_formatter___closed__4() { _start: { +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Tactic_tacticSeqBracketed_formatter___closed__6; +x_2 = lean_alloc_closure((void*)(l_Lean_ppDedent_formatter), 6, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Term_whereDecls_formatter___closed__5() { +_start: +{ lean_object* x_1; x_1 = lean_mk_string_from_bytes("where", 5); return x_1; } } -static lean_object* _init_l_Lean_Parser_Term_whereDecls_formatter___closed__5() { +static lean_object* _init_l_Lean_Parser_Term_whereDecls_formatter___closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_whereDecls_formatter___closed__4; +x_1 = l_Lean_Parser_Term_whereDecls_formatter___closed__5; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_whereDecls_formatter___closed__6() { +static lean_object* _init_l_Lean_Parser_Term_whereDecls_formatter___closed__7() { _start: { lean_object* x_1; lean_object* x_2; @@ -40651,11 +41116,11 @@ lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_whereDecls_formatter___closed__7() { +static lean_object* _init_l_Lean_Parser_Term_whereDecls_formatter___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Parser_Term_whereDecls_formatter___closed__6; +x_1 = l_Lean_Parser_Term_whereDecls_formatter___closed__7; x_2 = l_Lean_Parser_Tactic_sepByIndentSemicolon_formatter___closed__1; x_3 = l_Lean_Parser_Tactic_sepByIndentSemicolon_formatter___closed__2; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_sepBy1Indent_formatter___boxed), 8, 3); @@ -40665,25 +41130,37 @@ lean_closure_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Term_whereDecls_formatter___closed__8() { +static lean_object* _init_l_Lean_Parser_Term_whereDecls_formatter___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_whereDecls_formatter___closed__5; -x_2 = l_Lean_Parser_Term_whereDecls_formatter___closed__7; +x_1 = l_Lean_Parser_Term_whereDecls_formatter___closed__6; +x_2 = l_Lean_Parser_Term_whereDecls_formatter___closed__8; 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; } } -static lean_object* _init_l_Lean_Parser_Term_whereDecls_formatter___closed__9() { +static lean_object* _init_l_Lean_Parser_Term_whereDecls_formatter___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_whereDecls_formatter___closed__4; +x_2 = l_Lean_Parser_Term_whereDecls_formatter___closed__9; +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; +} +} +static lean_object* _init_l_Lean_Parser_Term_whereDecls_formatter___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_whereDecls_formatter___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_whereDecls_formatter___closed__8; +x_3 = l_Lean_Parser_Term_whereDecls_formatter___closed__10; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -40696,7 +41173,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_whereDecls_formatter(lean_object* x_ { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Term_whereDecls_formatter___closed__3; -x_7 = l_Lean_Parser_Term_whereDecls_formatter___closed__9; +x_7 = l_Lean_Parser_Term_whereDecls_formatter___closed__11; 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; } @@ -40756,8 +41233,8 @@ static lean_object* _init_l_Lean_Parser_Term_whereDecls_parenthesizer___closed__ _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_whereDecls_formatter___closed__4; -x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_parenthesizer), 6, 1); +x_1 = l_Lean_Parser_Tactic_tacticSeqBracketed_parenthesizer___closed__6; +x_2 = lean_alloc_closure((void*)(l_Lean_Parser_ppDedent_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } @@ -40766,17 +41243,27 @@ static lean_object* _init_l_Lean_Parser_Term_whereDecls_parenthesizer___closed__ _start: { lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Term_whereDecls_formatter___closed__5; +x_2 = lean_alloc_closure((void*)(l_Lean_Parser_symbol_parenthesizer), 6, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Term_whereDecls_parenthesizer___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; x_1 = l___regBuiltin_Lean_Parser_Term_letRecDecl_parenthesizer___closed__2; x_2 = lean_alloc_closure((void*)(l_Lean_Parser_ppGroup_parenthesizer), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_whereDecls_parenthesizer___closed__4() { +static lean_object* _init_l_Lean_Parser_Term_whereDecls_parenthesizer___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; uint8_t x_4; lean_object* x_5; lean_object* x_6; -x_1 = l_Lean_Parser_Term_whereDecls_parenthesizer___closed__3; +x_1 = l_Lean_Parser_Term_whereDecls_parenthesizer___closed__4; x_2 = l_Lean_Parser_Tactic_sepByIndentSemicolon_formatter___closed__1; x_3 = l_Lean_Parser_Tactic_sepByIndentSemicolon_parenthesizer___closed__1; x_4 = 1; @@ -40789,25 +41276,37 @@ lean_closure_set(x_6, 3, x_5); return x_6; } } -static lean_object* _init_l_Lean_Parser_Term_whereDecls_parenthesizer___closed__5() { +static lean_object* _init_l_Lean_Parser_Term_whereDecls_parenthesizer___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_whereDecls_parenthesizer___closed__3; +x_2 = l_Lean_Parser_Term_whereDecls_parenthesizer___closed__5; +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; +} +} +static lean_object* _init_l_Lean_Parser_Term_whereDecls_parenthesizer___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_whereDecls_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Term_whereDecls_parenthesizer___closed__4; +x_2 = l_Lean_Parser_Term_whereDecls_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; } } -static lean_object* _init_l_Lean_Parser_Term_whereDecls_parenthesizer___closed__6() { +static lean_object* _init_l_Lean_Parser_Term_whereDecls_parenthesizer___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_whereDecls_formatter___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_whereDecls_parenthesizer___closed__5; +x_3 = l_Lean_Parser_Term_whereDecls_parenthesizer___closed__7; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -40820,7 +41319,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_whereDecls_parenthesizer(lean_object { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Term_whereDecls_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Term_whereDecls_parenthesizer___closed__6; +x_7 = l_Lean_Parser_Term_whereDecls_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; } @@ -40874,7 +41373,7 @@ static lean_object* _init_l_Lean_Parser_Term_whereDecls___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_whereDecls_formatter___closed__4; +x_1 = l_Lean_Parser_Term_whereDecls_formatter___closed__5; x_2 = l_Lean_Parser_symbol(x_1); return x_2; } @@ -40934,30 +41433,40 @@ return x_3; static lean_object* _init_l_Lean_Parser_Term_whereDecls___closed__8() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_skip; +x_2 = l_Lean_Parser_Term_whereDecls___closed__7; +x_3 = l_Lean_Parser_andthen(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Term_whereDecls___closed__9() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_whereDecls_formatter___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_whereDecls___closed__7; +x_3 = l_Lean_Parser_Term_whereDecls___closed__8; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Term_whereDecls___closed__9() { +static lean_object* _init_l_Lean_Parser_Term_whereDecls___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_whereDecls___closed__1; -x_2 = l_Lean_Parser_Term_whereDecls___closed__8; +x_2 = l_Lean_Parser_Term_whereDecls___closed__9; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_whereDecls___closed__10() { +static lean_object* _init_l_Lean_Parser_Term_whereDecls___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_whereDecls_formatter___closed__2; -x_2 = l_Lean_Parser_Term_whereDecls___closed__9; +x_2 = l_Lean_Parser_Term_whereDecls___closed__10; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -40966,7 +41475,7 @@ static lean_object* _init_l_Lean_Parser_Term_whereDecls() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Term_whereDecls___closed__10; +x_1 = l_Lean_Parser_Term_whereDecls___closed__11; return x_1; } } @@ -41022,7 +41531,7 @@ static lean_object* _init_l_Lean_Parser_Term_matchAltsWhereDecls_formatter___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_match_formatter___closed__6; +x_1 = l_Lean_Parser_Term_match_formatter___closed__7; x_2 = l_Lean_Parser_Term_matchAltsWhereDecls_formatter___closed__4; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -41119,7 +41628,7 @@ static lean_object* _init_l_Lean_Parser_Term_matchAltsWhereDecls_parenthesizer__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_match_parenthesizer___closed__6; +x_1 = l_Lean_Parser_Term_match_parenthesizer___closed__7; x_2 = l_Lean_Parser_Term_matchAltsWhereDecls_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); @@ -41209,7 +41718,7 @@ static lean_object* _init_l_Lean_Parser_Term_matchAltsWhereDecls___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_match___closed__9; +x_1 = l_Lean_Parser_Term_match___closed__11; x_2 = l_Lean_Parser_Term_matchAltsWhereDecls___closed__2; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; @@ -41500,7 +42009,7 @@ static lean_object* _init_l_Lean_Parser_Term_noindex_formatter___closed__3() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_noindex_formatter___closed__2; -x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; +x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__5; 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); @@ -41713,17 +42222,18 @@ return x_2; static lean_object* _init_l_Lean_Parser_Term_binrel___closed__6() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_explicit___closed__6; -x_2 = l_Lean_Parser_andthen(x_1, x_1); -return x_2; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_skip; +x_2 = l_Lean_Parser_Term_explicit___closed__6; +x_3 = l_Lean_Parser_andthen(x_1, x_2); +return x_3; } } static lean_object* _init_l_Lean_Parser_Term_binrel___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_skip; +x_1 = l_Lean_Parser_Term_explicit___closed__6; x_2 = l_Lean_Parser_Term_binrel___closed__6; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; @@ -41733,7 +42243,7 @@ static lean_object* _init_l_Lean_Parser_Term_binrel___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_ident; +x_1 = l_Lean_Parser_skip; x_2 = l_Lean_Parser_Term_binrel___closed__7; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; @@ -41743,7 +42253,7 @@ static lean_object* _init_l_Lean_Parser_Term_binrel___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_binrel___closed__5; +x_1 = l_Lean_Parser_Term_ident; x_2 = l_Lean_Parser_Term_binrel___closed__8; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; @@ -41752,30 +42262,40 @@ return x_3; static lean_object* _init_l_Lean_Parser_Term_binrel___closed__10() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_binrel___closed__5; +x_2 = l_Lean_Parser_Term_binrel___closed__9; +x_3 = l_Lean_Parser_andthen(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Term_binrel___closed__11() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_binrel___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_binrel___closed__9; +x_3 = l_Lean_Parser_Term_binrel___closed__10; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Term_binrel___closed__11() { +static lean_object* _init_l_Lean_Parser_Term_binrel___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_binrel___closed__3; -x_2 = l_Lean_Parser_Term_binrel___closed__10; +x_2 = l_Lean_Parser_Term_binrel___closed__11; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_binrel___closed__12() { +static lean_object* _init_l_Lean_Parser_Term_binrel___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_binrel___closed__2; -x_2 = l_Lean_Parser_Term_binrel___closed__11; +x_2 = l_Lean_Parser_Term_binrel___closed__12; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -41784,7 +42304,7 @@ static lean_object* _init_l_Lean_Parser_Term_binrel() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Term_binrel___closed__12; +x_1 = l_Lean_Parser_Term_binrel___closed__13; return x_1; } } @@ -41818,7 +42338,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_binrel_declRange___clo { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_unsigned_to_nat(508u); -x_2 = lean_unsigned_to_nat(76u); +x_2 = lean_unsigned_to_nat(87u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -41832,7 +42352,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Parser_Term_binrel_declRange___closed__1; x_2 = lean_unsigned_to_nat(23u); x_3 = l___regBuiltin_Lean_Parser_Term_binrel_declRange___closed__2; -x_4 = lean_unsigned_to_nat(76u); +x_4 = lean_unsigned_to_nat(87u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -41934,19 +42454,20 @@ return x_2; static lean_object* _init_l_Lean_Parser_Term_binrel_formatter___closed__3() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; -x_2 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); -lean_closure_set(x_2, 0, x_1); -lean_closure_set(x_2, 1, x_1); -return x_2; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__5; +x_2 = l_Lean_Parser_Term_typeAscription_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; } } static lean_object* _init_l_Lean_Parser_Term_binrel_formatter___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_sufficesDecl_formatter___closed__2; +x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; x_2 = l_Lean_Parser_Term_binrel_formatter___closed__3; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -42066,19 +42587,20 @@ return x_2; static lean_object* _init_l_Lean_Parser_Term_binrel_parenthesizer___closed__3() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_explicit_parenthesizer___closed__3; -x_2 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); -lean_closure_set(x_2, 0, x_1); -lean_closure_set(x_2, 1, x_1); -return x_2; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__5; +x_2 = l_Lean_Parser_Term_explicit_parenthesizer___closed__3; +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; } } static lean_object* _init_l_Lean_Parser_Term_binrel_parenthesizer___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__2; +x_1 = l_Lean_Parser_Term_explicit_parenthesizer___closed__3; x_2 = l_Lean_Parser_Term_binrel_parenthesizer___closed__3; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -42090,7 +42612,7 @@ static lean_object* _init_l_Lean_Parser_Term_binrel_parenthesizer___closed__5() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_syntheticHole_parenthesizer___closed__3; +x_1 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__5; x_2 = l_Lean_Parser_Term_binrel_parenthesizer___closed__4; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -42102,7 +42624,7 @@ static lean_object* _init_l_Lean_Parser_Term_binrel_parenthesizer___closed__6() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_binrel_parenthesizer___closed__2; +x_1 = l_Lean_Parser_Term_syntheticHole_parenthesizer___closed__3; x_2 = l_Lean_Parser_Term_binrel_parenthesizer___closed__5; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -42113,10 +42635,22 @@ return x_3; static lean_object* _init_l_Lean_Parser_Term_binrel_parenthesizer___closed__7() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_binrel_parenthesizer___closed__2; +x_2 = l_Lean_Parser_Term_binrel_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; +} +} +static lean_object* _init_l_Lean_Parser_Term_binrel_parenthesizer___closed__8() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_binrel___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_binrel_parenthesizer___closed__6; +x_3 = l_Lean_Parser_Term_binrel_parenthesizer___closed__7; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -42129,7 +42663,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_binrel_parenthesizer(lean_object* x_ { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Term_binrel_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Term_binrel_parenthesizer___closed__7; +x_7 = l_Lean_Parser_Term_binrel_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; } @@ -42221,7 +42755,7 @@ static lean_object* _init_l_Lean_Parser_Term_binrel__no__prop___closed__6() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_binrel__no__prop___closed__5; -x_2 = l_Lean_Parser_Term_binrel___closed__8; +x_2 = l_Lean_Parser_Term_binrel___closed__9; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -42313,7 +42847,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_binrel__no__prop_declR { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_unsigned_to_nat(511u); -x_2 = lean_unsigned_to_nat(84u); +x_2 = lean_unsigned_to_nat(95u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -42327,7 +42861,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Parser_Term_binrel__no__prop_declRange___closed__1; x_2 = lean_unsigned_to_nat(23u); x_3 = l___regBuiltin_Lean_Parser_Term_binrel__no__prop_declRange___closed__2; -x_4 = lean_unsigned_to_nat(84u); +x_4 = lean_unsigned_to_nat(95u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -42528,7 +43062,7 @@ static lean_object* _init_l_Lean_Parser_Term_binrel__no__prop_parenthesizer___cl { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_binrel__no__prop_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Term_binrel_parenthesizer___closed__5; +x_2 = l_Lean_Parser_Term_binrel_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); @@ -42646,7 +43180,7 @@ static lean_object* _init_l_Lean_Parser_Term_binop___closed__6() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_binop___closed__5; -x_2 = l_Lean_Parser_Term_binrel___closed__8; +x_2 = l_Lean_Parser_Term_binrel___closed__9; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -42720,7 +43254,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_binop_declRange___clos { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_unsigned_to_nat(513u); -x_2 = lean_unsigned_to_nat(75u); +x_2 = lean_unsigned_to_nat(86u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -42734,7 +43268,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Parser_Term_binop_declRange___closed__1; x_2 = lean_unsigned_to_nat(23u); x_3 = l___regBuiltin_Lean_Parser_Term_binop_declRange___closed__2; -x_4 = lean_unsigned_to_nat(75u); +x_4 = lean_unsigned_to_nat(86u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -42935,7 +43469,7 @@ static lean_object* _init_l_Lean_Parser_Term_binop_parenthesizer___closed__3() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_binop_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Term_binrel_parenthesizer___closed__5; +x_2 = l_Lean_Parser_Term_binrel_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); @@ -43053,7 +43587,7 @@ static lean_object* _init_l_Lean_Parser_Term_binop__lazy___closed__6() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_binop__lazy___closed__5; -x_2 = l_Lean_Parser_Term_binrel___closed__8; +x_2 = l_Lean_Parser_Term_binrel___closed__9; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -43127,7 +43661,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_binop__lazy_declRange_ { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_unsigned_to_nat(515u); -x_2 = lean_unsigned_to_nat(80u); +x_2 = lean_unsigned_to_nat(91u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -43141,7 +43675,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Parser_Term_binop__lazy_declRange___closed__1; x_2 = lean_unsigned_to_nat(23u); x_3 = l___regBuiltin_Lean_Parser_Term_binop__lazy_declRange___closed__2; -x_4 = lean_unsigned_to_nat(80u); +x_4 = lean_unsigned_to_nat(91u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -43342,7 +43876,7 @@ static lean_object* _init_l_Lean_Parser_Term_binop__lazy_parenthesizer___closed_ { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_binop__lazy_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Term_binrel_parenthesizer___closed__5; +x_2 = l_Lean_Parser_Term_binrel_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); @@ -43459,8 +43993,8 @@ static lean_object* _init_l_Lean_Parser_Term_unop___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_skip; -x_2 = l_Lean_Parser_Term_explicit___closed__6; +x_1 = l_Lean_Parser_Term_ident; +x_2 = l_Lean_Parser_Term_binrel___closed__6; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -43469,7 +44003,7 @@ static lean_object* _init_l_Lean_Parser_Term_unop___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_ident; +x_1 = l_Lean_Parser_Term_unop___closed__5; x_2 = l_Lean_Parser_Term_unop___closed__6; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; @@ -43478,40 +44012,30 @@ return x_3; static lean_object* _init_l_Lean_Parser_Term_unop___closed__8() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_unop___closed__5; -x_2 = l_Lean_Parser_Term_unop___closed__7; -x_3 = l_Lean_Parser_andthen(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Parser_Term_unop___closed__9() { -_start: -{ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_unop___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_unop___closed__8; +x_3 = l_Lean_Parser_Term_unop___closed__7; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Term_unop___closed__10() { +static lean_object* _init_l_Lean_Parser_Term_unop___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_unop___closed__3; -x_2 = l_Lean_Parser_Term_unop___closed__9; +x_2 = l_Lean_Parser_Term_unop___closed__8; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_unop___closed__11() { +static lean_object* _init_l_Lean_Parser_Term_unop___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_unop___closed__2; -x_2 = l_Lean_Parser_Term_unop___closed__10; +x_2 = l_Lean_Parser_Term_unop___closed__9; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -43520,7 +44044,7 @@ static lean_object* _init_l_Lean_Parser_Term_unop() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Term_unop___closed__11; +x_1 = l_Lean_Parser_Term_unop___closed__10; return x_1; } } @@ -43672,7 +44196,7 @@ static lean_object* _init_l_Lean_Parser_Term_unop_formatter___closed__3() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_syntheticHole_formatter___closed__3; -x_2 = l_Lean_Parser_Term_basicFun_formatter___closed__9; +x_2 = l_Lean_Parser_Term_typeAscription_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); @@ -43780,8 +44304,8 @@ static lean_object* _init_l_Lean_Parser_Term_unop_parenthesizer___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Term_explicit_parenthesizer___closed__3; +x_1 = l_Lean_Parser_Term_syntheticHole_parenthesizer___closed__3; +x_2 = l_Lean_Parser_Term_binrel_parenthesizer___closed__3; 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); @@ -43792,7 +44316,7 @@ static lean_object* _init_l_Lean_Parser_Term_unop_parenthesizer___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_syntheticHole_parenthesizer___closed__3; +x_1 = l_Lean_Parser_Term_unop_parenthesizer___closed__2; x_2 = l_Lean_Parser_Term_unop_parenthesizer___closed__3; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -43803,22 +44327,10 @@ return x_3; static lean_object* _init_l_Lean_Parser_Term_unop_parenthesizer___closed__5() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_unop_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Term_unop_parenthesizer___closed__4; -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; -} -} -static lean_object* _init_l_Lean_Parser_Term_unop_parenthesizer___closed__6() { -_start: -{ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_unop___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_unop_parenthesizer___closed__5; +x_3 = l_Lean_Parser_Term_unop_parenthesizer___closed__4; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -43831,7 +44343,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_unop_parenthesizer(lean_object* x_1, { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Term_unop_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Term_unop_parenthesizer___closed__6; +x_7 = l_Lean_Parser_Term_unop_parenthesizer___closed__5; 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; } @@ -43923,7 +44435,7 @@ static lean_object* _init_l_Lean_Parser_Term_forInMacro___closed__6() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_explicit___closed__6; -x_2 = l_Lean_Parser_Term_binrel___closed__6; +x_2 = l_Lean_Parser_Term_binrel___closed__7; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -44007,7 +44519,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_forInMacro_declRange__ { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_unsigned_to_nat(520u); -x_2 = lean_unsigned_to_nat(78u); +x_2 = lean_unsigned_to_nat(89u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -44021,7 +44533,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Parser_Term_forInMacro_declRange___closed__1; x_2 = lean_unsigned_to_nat(23u); x_3 = l___regBuiltin_Lean_Parser_Term_forInMacro_declRange___closed__2; -x_4 = lean_unsigned_to_nat(78u); +x_4 = lean_unsigned_to_nat(89u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -44124,7 +44636,7 @@ static lean_object* _init_l_Lean_Parser_Term_forInMacro_formatter___closed__3() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; +x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__5; x_2 = l_Lean_Parser_Term_binrel_formatter___closed__3; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -44234,7 +44746,7 @@ static lean_object* _init_l_Lean_Parser_Term_forInMacro_parenthesizer___closed__ { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_explicit_parenthesizer___closed__3; -x_2 = l_Lean_Parser_Term_binrel_parenthesizer___closed__3; +x_2 = l_Lean_Parser_Term_binrel_parenthesizer___closed__4; 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); @@ -44438,7 +44950,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_forInMacro_x27_declRan { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_unsigned_to_nat(522u); -x_2 = lean_unsigned_to_nat(79u); +x_2 = lean_unsigned_to_nat(90u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -44452,7 +44964,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Parser_Term_forInMacro_x27_declRange___closed__1; x_2 = lean_unsigned_to_nat(23u); x_3 = l___regBuiltin_Lean_Parser_Term_forInMacro_x27_declRange___closed__2; -x_4 = lean_unsigned_to_nat(79u); +x_4 = lean_unsigned_to_nat(90u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -45171,7 +45683,7 @@ static lean_object* _init_l_Lean_Parser_Term_withDeclName___closed__7() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_ident; -x_2 = l_Lean_Parser_Term_typeAscription___closed__6; +x_2 = l_Lean_Parser_Term_typeAscription___closed__9; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -45283,7 +45795,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_withDeclName_declRange { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_unsigned_to_nat(534u); -x_2 = lean_unsigned_to_nat(59u); +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); @@ -45297,7 +45809,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Parser_Term_withDeclName_declRange___closed__1; x_2 = lean_unsigned_to_nat(23u); x_3 = l___regBuiltin_Lean_Parser_Term_withDeclName_declRange___closed__2; -x_4 = lean_unsigned_to_nat(59u); +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); @@ -45410,8 +45922,8 @@ static lean_object* _init_l_Lean_Parser_Term_withDeclName_formatter___closed__4( _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_syntheticHole_formatter___closed__3; -x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; +x_1 = l_Lean_Parser_Term_withDeclName_formatter___closed__3; +x_2 = l_Lean_Parser_Term_unop_formatter___closed__3; 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); @@ -45422,7 +45934,7 @@ static lean_object* _init_l_Lean_Parser_Term_withDeclName_formatter___closed__5( _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_withDeclName_formatter___closed__3; +x_1 = l_Lean_Parser_Term_withDeclName_formatter___closed__2; x_2 = l_Lean_Parser_Term_withDeclName_formatter___closed__4; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -45433,22 +45945,10 @@ return x_3; static lean_object* _init_l_Lean_Parser_Term_withDeclName_formatter___closed__6() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_withDeclName_formatter___closed__2; -x_2 = l_Lean_Parser_Term_withDeclName_formatter___closed__5; -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; -} -} -static lean_object* _init_l_Lean_Parser_Term_withDeclName_formatter___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_withDeclName___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_withDeclName_formatter___closed__6; +x_3 = l_Lean_Parser_Term_withDeclName_formatter___closed__5; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -45461,7 +45961,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_withDeclName_formatter(lean_object* { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Term_withDeclName_formatter___closed__1; -x_7 = l_Lean_Parser_Term_withDeclName_formatter___closed__7; +x_7 = l_Lean_Parser_Term_withDeclName_formatter___closed__6; 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; } @@ -45542,7 +46042,7 @@ static lean_object* _init_l_Lean_Parser_Term_withDeclName_parenthesizer___closed { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_syntheticHole_parenthesizer___closed__3; -x_2 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__3; +x_2 = l_Lean_Parser_Term_typeAscription_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); @@ -45876,7 +46376,7 @@ static lean_object* _init_l_Lean_Parser_Term_typeOf_formatter___closed__3() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_typeOf_formatter___closed__2; -x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; +x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__5; 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); @@ -46091,7 +46591,7 @@ static lean_object* _init_l_Lean_Parser_Term_ensureTypeOf___closed__6() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_strLit; -x_2 = l_Lean_Parser_Term_typeAscription___closed__6; +x_2 = l_Lean_Parser_Term_typeAscription___closed__9; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -46185,7 +46685,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_ensureTypeOf_declRange { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_unsigned_to_nat(538u); -x_2 = lean_unsigned_to_nat(66u); +x_2 = lean_unsigned_to_nat(77u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -46199,7 +46699,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Parser_Term_ensureTypeOf_declRange___closed__1; x_2 = lean_unsigned_to_nat(23u); x_3 = l___regBuiltin_Lean_Parser_Term_ensureTypeOf_declRange___closed__2; -x_4 = lean_unsigned_to_nat(66u); +x_4 = lean_unsigned_to_nat(77u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -46303,7 +46803,7 @@ static lean_object* _init_l_Lean_Parser_Term_ensureTypeOf_formatter___closed__3( { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_str_formatter___closed__1; -x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; +x_2 = l_Lean_Parser_Term_typeAscription_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); @@ -46314,7 +46814,7 @@ static lean_object* _init_l_Lean_Parser_Term_ensureTypeOf_formatter___closed__4( _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; +x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__5; x_2 = l_Lean_Parser_Term_ensureTypeOf_formatter___closed__3; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -46424,7 +46924,7 @@ static lean_object* _init_l_Lean_Parser_Term_ensureTypeOf_parenthesizer___closed { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_str_parenthesizer___closed__1; -x_2 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__3; +x_2 = l_Lean_Parser_Term_typeAscription_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); @@ -46566,7 +47066,7 @@ static lean_object* _init_l_Lean_Parser_Term_ensureExpectedType___closed__6() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_strLit; -x_2 = l_Lean_Parser_Term_explicit___closed__6; +x_2 = l_Lean_Parser_Term_binrel___closed__6; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -46650,7 +47150,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_ensureExpectedType_dec { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_unsigned_to_nat(540u); -x_2 = lean_unsigned_to_nat(58u); +x_2 = lean_unsigned_to_nat(69u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -46664,7 +47164,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Parser_Term_ensureExpectedType_declRange___closed__1; x_2 = lean_unsigned_to_nat(23u); x_3 = l___regBuiltin_Lean_Parser_Term_ensureExpectedType_declRange___closed__2; -x_4 = lean_unsigned_to_nat(58u); +x_4 = lean_unsigned_to_nat(69u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -46865,7 +47365,7 @@ static lean_object* _init_l_Lean_Parser_Term_ensureExpectedType_parenthesizer___ { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_str_parenthesizer___closed__1; -x_2 = l_Lean_Parser_Term_explicit_parenthesizer___closed__3; +x_2 = l_Lean_Parser_Term_binrel_parenthesizer___closed__3; 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); @@ -47187,7 +47687,7 @@ static lean_object* _init_l_Lean_Parser_Term_noImplicitLambda_formatter___closed { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_noImplicitLambda_formatter___closed__2; -x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; +x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__5; 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); @@ -47402,7 +47902,7 @@ static lean_object* _init_l_Lean_Parser_Term_clear___closed__6() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_semicolonOrLinebreak; -x_2 = l_Lean_Parser_Term_typeAscription___closed__6; +x_2 = l_Lean_Parser_Term_typeAscription___closed__9; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -47514,7 +48014,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_clear_declRange___clos { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_unsigned_to_nat(549u); -x_2 = lean_unsigned_to_nat(58u); +x_2 = lean_unsigned_to_nat(77u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -47528,7 +48028,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Parser_Term_clear_declRange___closed__1; x_2 = lean_unsigned_to_nat(23u); x_3 = l___regBuiltin_Lean_Parser_Term_clear_declRange___closed__2; -x_4 = lean_unsigned_to_nat(58u); +x_4 = lean_unsigned_to_nat(77u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -47631,8 +48131,8 @@ static lean_object* _init_l_Lean_Parser_Term_clear_formatter___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_optSemicolon_formatter___closed__1; -x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; +x_1 = l_Lean_Parser_Term_whereDecls_formatter___closed__4; +x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__5; 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); @@ -47643,7 +48143,7 @@ static lean_object* _init_l_Lean_Parser_Term_clear_formatter___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_syntheticHole_formatter___closed__3; +x_1 = l_Lean_Parser_Term_optSemicolon_formatter___closed__1; x_2 = l_Lean_Parser_Term_clear_formatter___closed__3; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -47655,7 +48155,7 @@ static lean_object* _init_l_Lean_Parser_Term_clear_formatter___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_clear_formatter___closed__2; +x_1 = l_Lean_Parser_Term_syntheticHole_formatter___closed__3; x_2 = l_Lean_Parser_Term_clear_formatter___closed__4; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -47666,10 +48166,22 @@ return x_3; static lean_object* _init_l_Lean_Parser_Term_clear_formatter___closed__6() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_clear_formatter___closed__2; +x_2 = l_Lean_Parser_Term_clear_formatter___closed__5; +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; +} +} +static lean_object* _init_l_Lean_Parser_Term_clear_formatter___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_clear___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_clear_formatter___closed__5; +x_3 = l_Lean_Parser_Term_clear_formatter___closed__6; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -47682,7 +48194,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_clear_formatter(lean_object* x_1, le { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Term_clear_formatter___closed__1; -x_7 = l_Lean_Parser_Term_clear_formatter___closed__6; +x_7 = l_Lean_Parser_Term_clear_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; } @@ -47752,7 +48264,7 @@ static lean_object* _init_l_Lean_Parser_Term_clear_parenthesizer___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_optSemicolon_parenthesizer___closed__1; +x_1 = l_Lean_Parser_Term_whereDecls_parenthesizer___closed__2; x_2 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__3; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -47764,7 +48276,7 @@ static lean_object* _init_l_Lean_Parser_Term_clear_parenthesizer___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_syntheticHole_parenthesizer___closed__3; +x_1 = l_Lean_Parser_Term_optSemicolon_parenthesizer___closed__1; x_2 = l_Lean_Parser_Term_clear_parenthesizer___closed__3; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -47776,7 +48288,7 @@ static lean_object* _init_l_Lean_Parser_Term_clear_parenthesizer___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_clear_parenthesizer___closed__2; +x_1 = l_Lean_Parser_Term_syntheticHole_parenthesizer___closed__3; x_2 = l_Lean_Parser_Term_clear_parenthesizer___closed__4; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -47787,10 +48299,22 @@ return x_3; static lean_object* _init_l_Lean_Parser_Term_clear_parenthesizer___closed__6() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_clear_parenthesizer___closed__2; +x_2 = l_Lean_Parser_Term_clear_parenthesizer___closed__5; +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; +} +} +static lean_object* _init_l_Lean_Parser_Term_clear_parenthesizer___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_clear___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_clear_parenthesizer___closed__5; +x_3 = l_Lean_Parser_Term_clear_parenthesizer___closed__6; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -47803,7 +48327,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_clear_parenthesizer(lean_object* x_1 { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Term_clear_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Term_clear_parenthesizer___closed__6; +x_7 = l_Lean_Parser_Term_clear_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; } @@ -48137,7 +48661,7 @@ static lean_object* _init_l_Lean_Parser_Term_letMVar_formatter___closed__3() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Tactic_sepByIndentSemicolon_formatter___closed__2; -x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; +x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__5; 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); @@ -48148,7 +48672,7 @@ static lean_object* _init_l_Lean_Parser_Term_letMVar_formatter___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; +x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__5; x_2 = l_Lean_Parser_Term_letMVar_formatter___closed__3; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -50238,7 +50762,7 @@ static lean_object* _init_l_Lean_Parser_Term_noErrorIfUnused___closed__4() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("no_error_if_unused%", 19); +x_1 = lean_mk_string_from_bytes("no_error_if_unused% ", 20); return x_1; } } @@ -50348,7 +50872,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_noErrorIfUnused_declRa { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_unsigned_to_nat(568u); -x_2 = lean_unsigned_to_nat(37u); +x_2 = lean_unsigned_to_nat(38u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -50362,7 +50886,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Parser_Term_noErrorIfUnused_declRange___closed__1; x_2 = lean_unsigned_to_nat(23u); x_3 = l___regBuiltin_Lean_Parser_Term_noErrorIfUnused_declRange___closed__2; -x_4 = lean_unsigned_to_nat(37u); +x_4 = lean_unsigned_to_nat(38u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -50466,7 +50990,7 @@ static lean_object* _init_l_Lean_Parser_Term_noErrorIfUnused_formatter___closed_ { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_noErrorIfUnused_formatter___closed__2; -x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; +x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__5; 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); @@ -50771,57 +51295,74 @@ static lean_object* _init_l_Lean_Parser_Term_ellipsis___closed__4() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("`.` immediately after `..`", 26); +x_1 = lean_mk_string_from_bytes("..", 2); return x_1; } } static lean_object* _init_l_Lean_Parser_Term_ellipsis___closed__5() { _start: { +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Term_ellipsis___closed__4; +x_2 = l_Lean_Parser_symbol(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_Term_ellipsis___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("`.` immediately after `..`", 26); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_Term_ellipsis___closed__7() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_cdot___closed__7; -x_2 = l_Lean_Parser_Term_ellipsis___closed__4; +x_2 = l_Lean_Parser_Term_ellipsis___closed__6; x_3 = l_Lean_Parser_notFollowedBy(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_ellipsis___closed__6() { +static lean_object* _init_l_Lean_Parser_Term_ellipsis___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_optEllipsis___closed__5; -x_2 = l_Lean_Parser_Term_ellipsis___closed__5; +x_1 = l_Lean_Parser_Term_ellipsis___closed__5; +x_2 = l_Lean_Parser_Term_ellipsis___closed__7; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_ellipsis___closed__7() { +static lean_object* _init_l_Lean_Parser_Term_ellipsis___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_ellipsis___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_ellipsis___closed__6; +x_3 = l_Lean_Parser_Term_ellipsis___closed__8; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Term_ellipsis___closed__8() { +static lean_object* _init_l_Lean_Parser_Term_ellipsis___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_ellipsis___closed__3; -x_2 = l_Lean_Parser_Term_ellipsis___closed__7; +x_2 = l_Lean_Parser_Term_ellipsis___closed__9; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_ellipsis___closed__9() { +static lean_object* _init_l_Lean_Parser_Term_ellipsis___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_ellipsis___closed__2; -x_2 = l_Lean_Parser_Term_ellipsis___closed__8; +x_2 = l_Lean_Parser_Term_ellipsis___closed__10; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -50830,7 +51371,7 @@ static lean_object* _init_l_Lean_Parser_Term_ellipsis() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Term_ellipsis___closed__9; +x_1 = l_Lean_Parser_Term_ellipsis___closed__11; return x_1; } } @@ -51230,31 +51771,41 @@ static lean_object* _init_l_Lean_Parser_Term_ellipsis_formatter___closed__2() { _start: { lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Term_ellipsis___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; +} +} +static lean_object* _init_l_Lean_Parser_Term_ellipsis_formatter___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; x_1 = l_Lean_Parser_Term_cdot_formatter___closed__3; x_2 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_notFollowedBy_formatter___boxed), 5, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_ellipsis_formatter___closed__3() { +static lean_object* _init_l_Lean_Parser_Term_ellipsis_formatter___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_optEllipsis_formatter___closed__2; -x_2 = l_Lean_Parser_Term_ellipsis_formatter___closed__2; +x_1 = l_Lean_Parser_Term_ellipsis_formatter___closed__2; +x_2 = l_Lean_Parser_Term_ellipsis_formatter___closed__3; 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; } } -static lean_object* _init_l_Lean_Parser_Term_ellipsis_formatter___closed__4() { +static lean_object* _init_l_Lean_Parser_Term_ellipsis_formatter___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_ellipsis___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_ellipsis_formatter___closed__3; +x_3 = l_Lean_Parser_Term_ellipsis_formatter___closed__4; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -51267,7 +51818,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_ellipsis_formatter(lean_object* x_1, { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Term_ellipsis_formatter___closed__1; -x_7 = l_Lean_Parser_Term_ellipsis_formatter___closed__4; +x_7 = l_Lean_Parser_Term_ellipsis_formatter___closed__5; 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; } @@ -51310,7 +51861,7 @@ static lean_object* _init_l_Lean_Parser_Term_argument_formatter___closed__1() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_Lean_Parser_Term_ellipsis_formatter___closed__2; -x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; +x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__5; 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); @@ -51555,31 +52106,41 @@ static lean_object* _init_l_Lean_Parser_Term_ellipsis_parenthesizer___closed__2( _start: { lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_Term_ellipsis___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; +} +} +static lean_object* _init_l_Lean_Parser_Term_ellipsis_parenthesizer___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; x_1 = l_Lean_Parser_Term_cdot_parenthesizer___closed__3; x_2 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_notFollowedBy_parenthesizer___boxed), 5, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Term_ellipsis_parenthesizer___closed__3() { +static lean_object* _init_l_Lean_Parser_Term_ellipsis_parenthesizer___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_optEllipsis_parenthesizer___closed__2; -x_2 = l_Lean_Parser_Term_ellipsis_parenthesizer___closed__2; +x_1 = l_Lean_Parser_Term_ellipsis_parenthesizer___closed__2; +x_2 = l_Lean_Parser_Term_ellipsis_parenthesizer___closed__3; 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; } } -static lean_object* _init_l_Lean_Parser_Term_ellipsis_parenthesizer___closed__4() { +static lean_object* _init_l_Lean_Parser_Term_ellipsis_parenthesizer___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_ellipsis___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_ellipsis_parenthesizer___closed__3; +x_3 = l_Lean_Parser_Term_ellipsis_parenthesizer___closed__4; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -51592,7 +52153,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_ellipsis_parenthesizer(lean_object* { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Term_ellipsis_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Term_ellipsis_parenthesizer___closed__4; +x_7 = l_Lean_Parser_Term_ellipsis_parenthesizer___closed__5; 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; } @@ -53667,7 +54228,7 @@ static lean_object* _init_l_Lean_Parser_Term_namedPattern_formatter___closed__4( { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_namedPattern_formatter___closed__3; -x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; +x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__5; 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); @@ -54791,7 +55352,7 @@ static lean_object* _init_l_Lean_Parser_Term_subst_formatter___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; uint8_t x_4; lean_object* x_5; lean_object* x_6; -x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; +x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__5; x_2 = l_Lean_Parser_Term_subst___closed__3; x_3 = l_Lean_Parser_Term_subst_formatter___closed__1; x_4 = 0; @@ -55243,7 +55804,7 @@ static lean_object* _init_l_Lean_Parser_Term_panic_formatter___closed__3() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_panic_formatter___closed__2; -x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; +x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__5; 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); @@ -56096,7 +56657,7 @@ static lean_object* _init_l_Lean_Parser_Term_dbgTrace_formatter___closed__3() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; +x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__5; x_2 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_interpolatedStr_formatter), 6, 1); lean_closure_set(x_2, 0, x_1); return x_2; @@ -56107,7 +56668,7 @@ static lean_object* _init_l_Lean_Parser_Term_dbgTrace_formatter___closed__4() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_dbgTrace_formatter___closed__3; -x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; +x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__5; 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); @@ -56629,7 +57190,7 @@ static lean_object* _init_l_Lean_Parser_Term_assert_formatter___closed__3() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_assert_formatter___closed__2; -x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; +x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__5; 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); @@ -57003,7 +57564,7 @@ static lean_object* _init_l_Lean_Parser_Term_stateRefT___closed__4() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("StateRefT", 9); +x_1 = lean_mk_string_from_bytes("StateRefT ", 10); return x_1; } } @@ -57020,7 +57581,7 @@ static lean_object* _init_l_Lean_Parser_Term_stateRefT___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_macroArg; +x_1 = l_Lean_Parser_skip; x_2 = l_Lean_Parser_Term_macroLastArg; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; @@ -57030,7 +57591,7 @@ static lean_object* _init_l_Lean_Parser_Term_stateRefT___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_stateRefT___closed__5; +x_1 = l_Lean_Parser_Term_macroArg; x_2 = l_Lean_Parser_Term_stateRefT___closed__6; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; @@ -57039,30 +57600,40 @@ return x_3; static lean_object* _init_l_Lean_Parser_Term_stateRefT___closed__8() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_stateRefT___closed__5; +x_2 = l_Lean_Parser_Term_stateRefT___closed__7; +x_3 = l_Lean_Parser_andthen(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_Term_stateRefT___closed__9() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Parser_Term_stateRefT___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_stateRefT___closed__7; +x_3 = l_Lean_Parser_Term_stateRefT___closed__8; x_4 = l_Lean_Parser_leadingNode(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Parser_Term_stateRefT___closed__9() { +static lean_object* _init_l_Lean_Parser_Term_stateRefT___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_stateRefT___closed__3; -x_2 = l_Lean_Parser_Term_stateRefT___closed__8; +x_2 = l_Lean_Parser_Term_stateRefT___closed__9; x_3 = l_Lean_Parser_withAntiquot(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Term_stateRefT___closed__10() { +static lean_object* _init_l_Lean_Parser_Term_stateRefT___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_stateRefT___closed__2; -x_2 = l_Lean_Parser_Term_stateRefT___closed__9; +x_2 = l_Lean_Parser_Term_stateRefT___closed__10; x_3 = l_Lean_Parser_withCache(x_1, x_2); return x_3; } @@ -57071,7 +57642,7 @@ static lean_object* _init_l_Lean_Parser_Term_stateRefT() { _start: { lean_object* x_1; -x_1 = l_Lean_Parser_Term_stateRefT___closed__10; +x_1 = l_Lean_Parser_Term_stateRefT___closed__11; return x_1; } } @@ -57105,7 +57676,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_stateRefT_declRange___ { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_unsigned_to_nat(671u); -x_2 = lean_unsigned_to_nat(41u); +x_2 = lean_unsigned_to_nat(53u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -57119,7 +57690,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Parser_Term_stateRefT_declRange___closed__1; x_2 = lean_unsigned_to_nat(23u); x_3 = l___regBuiltin_Lean_Parser_Term_stateRefT_declRange___closed__2; -x_4 = lean_unsigned_to_nat(41u); +x_4 = lean_unsigned_to_nat(53u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -57231,7 +57802,7 @@ static lean_object* _init_l_Lean_Parser_Term_macroDollarArg_formatter___closed__ { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_macroDollarArg_formatter___closed__2; -x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; +x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__5; 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); @@ -57353,7 +57924,7 @@ static lean_object* _init_l_Lean_Parser_Term_stateRefT_formatter___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_macroLastArg_formatter___closed__1; +x_1 = l_Lean_Parser_Term_typeAscription_formatter___closed__4; x_2 = l_Lean_Parser_Term_stateRefT_formatter___closed__3; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -57365,7 +57936,7 @@ static lean_object* _init_l_Lean_Parser_Term_stateRefT_formatter___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_stateRefT_formatter___closed__2; +x_1 = l_Lean_Parser_Term_macroLastArg_formatter___closed__1; x_2 = l_Lean_Parser_Term_stateRefT_formatter___closed__4; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -57376,10 +57947,22 @@ return x_3; static lean_object* _init_l_Lean_Parser_Term_stateRefT_formatter___closed__6() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_stateRefT_formatter___closed__2; +x_2 = l_Lean_Parser_Term_stateRefT_formatter___closed__5; +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; +} +} +static lean_object* _init_l_Lean_Parser_Term_stateRefT_formatter___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_stateRefT___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_stateRefT_formatter___closed__5; +x_3 = l_Lean_Parser_Term_stateRefT_formatter___closed__6; x_4 = lean_alloc_closure((void*)(l_Lean_Parser_leadingNode_formatter___boxed), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -57392,7 +57975,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_stateRefT_formatter(lean_object* x_1 { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Term_stateRefT_formatter___closed__1; -x_7 = l_Lean_Parser_Term_stateRefT_formatter___closed__6; +x_7 = l_Lean_Parser_Term_stateRefT_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; } @@ -57604,7 +58187,7 @@ static lean_object* _init_l_Lean_Parser_Term_stateRefT_parenthesizer___closed__4 _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_macroLastArg_parenthesizer___closed__1; +x_1 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__5; x_2 = l_Lean_Parser_Term_stateRefT_parenthesizer___closed__3; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -57616,7 +58199,7 @@ static lean_object* _init_l_Lean_Parser_Term_stateRefT_parenthesizer___closed__5 _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_stateRefT_parenthesizer___closed__2; +x_1 = l_Lean_Parser_Term_macroLastArg_parenthesizer___closed__1; x_2 = l_Lean_Parser_Term_stateRefT_parenthesizer___closed__4; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -57627,10 +58210,22 @@ return x_3; static lean_object* _init_l_Lean_Parser_Term_stateRefT_parenthesizer___closed__6() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Parser_Term_stateRefT_parenthesizer___closed__2; +x_2 = l_Lean_Parser_Term_stateRefT_parenthesizer___closed__5; +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; +} +} +static lean_object* _init_l_Lean_Parser_Term_stateRefT_parenthesizer___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_stateRefT___closed__2; x_2 = lean_unsigned_to_nat(1024u); -x_3 = l_Lean_Parser_Term_stateRefT_parenthesizer___closed__5; +x_3 = l_Lean_Parser_Term_stateRefT_parenthesizer___closed__6; x_4 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer), 8, 3); lean_closure_set(x_4, 0, x_1); lean_closure_set(x_4, 1, x_2); @@ -57643,7 +58238,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Term_stateRefT_parenthesizer(lean_object* { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_Parser_Term_stateRefT_parenthesizer___closed__1; -x_7 = l_Lean_Parser_Term_stateRefT_parenthesizer___closed__6; +x_7 = l_Lean_Parser_Term_stateRefT_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; } @@ -57754,7 +58349,7 @@ static lean_object* _init_l_Lean_Parser_Term_dynamicQuot___closed__8() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_dynamicQuot___closed__7; -x_2 = l_Lean_Parser_Term_typeAscription___closed__17; +x_2 = l_Lean_Parser_Term_typeAscription___closed__18; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -57763,7 +58358,7 @@ static lean_object* _init_l_Lean_Parser_Term_dynamicQuot___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_matchAlt___closed__7; +x_1 = l_Lean_Parser_Term_matchAlt___closed__5; x_2 = l_Lean_Parser_Term_dynamicQuot___closed__8; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; @@ -57824,7 +58419,7 @@ static lean_object* _init_l_Lean_Parser_Term_dynamicQuot___closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_typeAscription___closed__14; +x_1 = l_Lean_Parser_Term_typeAscription___closed__15; x_2 = l_Lean_Parser_Term_dynamicQuot___closed__14; x_3 = l_Lean_Parser_adaptCacheableContext(x_1, x_2); return x_3; @@ -57868,7 +58463,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Term_dynamicQuot_declRange_ { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_unsigned_to_nat(674u); -x_2 = lean_unsigned_to_nat(63u); +x_2 = lean_unsigned_to_nat(64u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -57882,7 +58477,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Parser_Term_dynamicQuot_declRange___closed__1; x_2 = lean_unsigned_to_nat(23u); x_3 = l___regBuiltin_Lean_Parser_Term_dynamicQuot_declRange___closed__2; -x_4 = lean_unsigned_to_nat(63u); +x_4 = lean_unsigned_to_nat(64u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -58008,7 +58603,7 @@ static lean_object* _init_l_Lean_Parser_Term_dynamicQuot_formatter___closed__5() { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_dynamicQuot_formatter___closed__4; -x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__10; +x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__12; 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); @@ -58019,7 +58614,7 @@ static lean_object* _init_l_Lean_Parser_Term_dynamicQuot_formatter___closed__6() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_matchAlt_formatter___closed__3; +x_1 = l_Lean_Parser_Term_matchAlt_formatter___closed__2; x_2 = l_Lean_Parser_Term_dynamicQuot_formatter___closed__5; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_andthen_formatter), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -58163,7 +58758,7 @@ static lean_object* _init_l_Lean_Parser_Term_dynamicQuot_parenthesizer___closed_ { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Term_dynamicQuot_parenthesizer___closed__4; -x_2 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__10; +x_2 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__12; 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); @@ -58174,7 +58769,7 @@ static lean_object* _init_l_Lean_Parser_Term_dynamicQuot_parenthesizer___closed_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_matchAlt_parenthesizer___closed__3; +x_1 = l_Lean_Parser_Term_matchAlt_parenthesizer___closed__2; x_2 = l_Lean_Parser_Term_dynamicQuot_parenthesizer___closed__5; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer), 7, 2); lean_closure_set(x_3, 0, x_1); @@ -58703,7 +59298,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_quot___closed__4() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("`(tactic|", 9); +x_1 = lean_mk_string_from_bytes("`(tactic| ", 10); return x_1; } } @@ -58729,7 +59324,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_quot___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_typeAscription___closed__14; +x_1 = l_Lean_Parser_Term_typeAscription___closed__15; x_2 = l_Lean_Parser_Tactic_quot___closed__6; x_3 = l_Lean_Parser_adaptCacheableContext(x_1, x_2); return x_3; @@ -58740,7 +59335,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_quot___closed__8() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Tactic_quot___closed__7; -x_2 = l_Lean_Parser_Term_typeAscription___closed__17; +x_2 = l_Lean_Parser_Term_typeAscription___closed__18; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -58824,7 +59419,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Tactic_quot_declRange___clo { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_unsigned_to_nat(682u); -x_2 = lean_unsigned_to_nat(67u); +x_2 = lean_unsigned_to_nat(68u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -58838,7 +59433,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Parser_Tactic_quot_declRange___closed__1; x_2 = lean_unsigned_to_nat(33u); x_3 = l___regBuiltin_Lean_Parser_Tactic_quot_declRange___closed__2; -x_4 = lean_unsigned_to_nat(67u); +x_4 = lean_unsigned_to_nat(68u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -58962,7 +59557,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_quot_formatter___closed__5() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Tactic_quot_formatter___closed__4; -x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__10; +x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__12; 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); @@ -59091,7 +59686,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_quot_parenthesizer___closed__5() { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Tactic_quot_parenthesizer___closed__4; -x_2 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__10; +x_2 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__12; 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); @@ -59212,7 +59807,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_quotSeq___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Term_typeAscription___closed__14; +x_1 = l_Lean_Parser_Term_typeAscription___closed__15; x_2 = l_Lean_Parser_Tactic_quotSeq___closed__4; x_3 = l_Lean_Parser_adaptCacheableContext(x_1, x_2); return x_3; @@ -59223,7 +59818,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_quotSeq___closed__6() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Tactic_quotSeq___closed__5; -x_2 = l_Lean_Parser_Term_typeAscription___closed__17; +x_2 = l_Lean_Parser_Term_typeAscription___closed__18; x_3 = l_Lean_Parser_andthen(x_1, x_2); return x_3; } @@ -59307,7 +59902,7 @@ static lean_object* _init_l___regBuiltin_Lean_Parser_Tactic_quotSeq_declRange___ { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_unsigned_to_nat(684u); -x_2 = lean_unsigned_to_nat(66u); +x_2 = lean_unsigned_to_nat(67u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -59321,7 +59916,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Parser_Tactic_quotSeq_declRange___closed__1; x_2 = lean_unsigned_to_nat(23u); x_3 = l___regBuiltin_Lean_Parser_Tactic_quotSeq_declRange___closed__2; -x_4 = lean_unsigned_to_nat(66u); +x_4 = lean_unsigned_to_nat(67u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -59505,7 +60100,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_quotSeq_formatter___closed__4() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Tactic_quotSeq_formatter___closed__3; -x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__10; +x_2 = l_Lean_Parser_Term_typeAscription_formatter___closed__12; 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); @@ -59694,7 +60289,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_quotSeq_parenthesizer___closed__4 { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_Tactic_quotSeq_parenthesizer___closed__3; -x_2 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__10; +x_2 = l_Lean_Parser_Term_typeAscription_parenthesizer___closed__12; 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); @@ -59770,7 +60365,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_5007____closed__1() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -59780,7 +60375,7 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__2() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__2() { _start: { lean_object* x_1; lean_object* x_2; @@ -59790,7 +60385,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_5007____closed__3() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__3() { _start: { lean_object* x_1; lean_object* x_2; @@ -59800,7 +60395,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_5007____closed__4() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__4() { _start: { lean_object* x_1; lean_object* x_2; @@ -59810,7 +60405,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_5007____closed__5() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__5() { _start: { lean_object* x_1; lean_object* x_2; @@ -59820,7 +60415,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_5007____closed__6() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -59830,7 +60425,7 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__7() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__7() { _start: { lean_object* x_1; lean_object* x_2; @@ -59840,7 +60435,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_5007____closed__8() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__8() { _start: { lean_object* x_1; lean_object* x_2; @@ -59850,7 +60445,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_5007____closed__9() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__9() { _start: { lean_object* x_1; lean_object* x_2; @@ -59860,7 +60455,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_5007____closed__10() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__10() { _start: { lean_object* x_1; lean_object* x_2; @@ -59870,7 +60465,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_5007____closed__11() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -59880,7 +60475,7 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__12() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__12() { _start: { lean_object* x_1; lean_object* x_2; @@ -59890,7 +60485,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_5007____closed__13() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__13() { _start: { lean_object* x_1; lean_object* x_2; @@ -59900,7 +60495,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_5007____closed__14() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__14() { _start: { lean_object* x_1; lean_object* x_2; @@ -59910,7 +60505,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_5007____closed__15() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__15() { _start: { lean_object* x_1; lean_object* x_2; @@ -59920,7 +60515,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_5007____closed__16() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -59930,7 +60525,7 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__17() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__17() { _start: { lean_object* x_1; lean_object* x_2; @@ -59940,7 +60535,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_5007____closed__18() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__18() { _start: { lean_object* x_1; lean_object* x_2; @@ -59950,7 +60545,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_5007____closed__19() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__19() { _start: { lean_object* x_1; lean_object* x_2; @@ -59960,7 +60555,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_5007____closed__20() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__20() { _start: { lean_object* x_1; lean_object* x_2; @@ -59970,7 +60565,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_5007____closed__21() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -59980,7 +60575,7 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__22() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__22() { _start: { lean_object* x_1; lean_object* x_2; @@ -59990,7 +60585,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_5007____closed__23() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__23() { _start: { lean_object* x_1; lean_object* x_2; @@ -60000,7 +60595,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_5007____closed__24() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__24() { _start: { lean_object* x_1; lean_object* x_2; @@ -60010,7 +60605,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_5007____closed__25() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__25() { _start: { lean_object* x_1; lean_object* x_2; @@ -60020,7 +60615,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_5007____closed__26() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__26() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -60030,7 +60625,7 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__27() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__27() { _start: { lean_object* x_1; lean_object* x_2; @@ -60040,7 +60635,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_5007____closed__28() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__28() { _start: { lean_object* x_1; lean_object* x_2; @@ -60050,7 +60645,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_5007____closed__29() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__29() { _start: { lean_object* x_1; lean_object* x_2; @@ -60060,7 +60655,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_5007____closed__30() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__30() { _start: { lean_object* x_1; lean_object* x_2; @@ -60070,7 +60665,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_5007____closed__31() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__31() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -60080,7 +60675,7 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__32() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__32() { _start: { lean_object* x_1; lean_object* x_2; @@ -60090,7 +60685,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_5007____closed__33() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__33() { _start: { lean_object* x_1; lean_object* x_2; @@ -60100,7 +60695,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_5007____closed__34() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__34() { _start: { lean_object* x_1; lean_object* x_2; @@ -60110,7 +60705,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_5007____closed__35() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__35() { _start: { lean_object* x_1; lean_object* x_2; @@ -60120,7 +60715,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_5007____closed__36() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__36() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -60130,7 +60725,7 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__37() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__37() { _start: { lean_object* x_1; lean_object* x_2; @@ -60140,7 +60735,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_5007____closed__38() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__38() { _start: { lean_object* x_1; lean_object* x_2; @@ -60150,7 +60745,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_5007____closed__39() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__39() { _start: { lean_object* x_1; lean_object* x_2; @@ -60160,7 +60755,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_5007____closed__40() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__40() { _start: { lean_object* x_1; lean_object* x_2; @@ -60170,7 +60765,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_5007____closed__41() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__41() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -60180,7 +60775,7 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__42() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__42() { _start: { lean_object* x_1; lean_object* x_2; @@ -60190,7 +60785,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_5007____closed__43() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__43() { _start: { lean_object* x_1; lean_object* x_2; @@ -60200,7 +60795,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_5007____closed__44() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__44() { _start: { lean_object* x_1; lean_object* x_2; @@ -60210,7 +60805,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_5007____closed__45() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__45() { _start: { lean_object* x_1; lean_object* x_2; @@ -60220,14 +60815,14 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098_(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; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__1; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__1; x_3 = l_Lean_Parser_Term_letDecl___closed__2; -x_4 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__2; -x_5 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__3; +x_4 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__2; +x_5 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__3; x_6 = l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Term___hyg_195____closed__9; x_7 = l_Lean_Parser_registerAlias(x_2, x_3, x_4, x_5, x_6, x_1); if (lean_obj_tag(x_7) == 0) @@ -60237,7 +60832,7 @@ x_8 = lean_ctor_get(x_7, 1); lean_inc(x_8); lean_dec(x_7); x_9 = l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Term___hyg_195____closed__12; -x_10 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__4; +x_10 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__4; x_11 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_2, x_10, x_8); if (lean_obj_tag(x_11) == 0) { @@ -60246,7 +60841,7 @@ x_12 = lean_ctor_get(x_11, 1); lean_inc(x_12); lean_dec(x_11); x_13 = l_Lean_Parser_Tactic_initFn____x40_Lean_Parser_Term___hyg_195____closed__15; -x_14 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__5; +x_14 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__5; x_15 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_2, x_14, x_12); if (lean_obj_tag(x_15) == 0) { @@ -60254,10 +60849,10 @@ lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean 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_Term___hyg_5007____closed__6; +x_17 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__6; x_18 = l_Lean_Parser_Term_haveDecl___closed__2; -x_19 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__7; -x_20 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__8; +x_19 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__7; +x_20 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__8; x_21 = l_Lean_Parser_registerAlias(x_17, x_18, x_19, x_20, x_6, x_16); if (lean_obj_tag(x_21) == 0) { @@ -60265,7 +60860,7 @@ 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_5007____closed__9; +x_23 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__9; x_24 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_17, x_23, x_22); if (lean_obj_tag(x_24) == 0) { @@ -60273,7 +60868,7 @@ lean_object* x_25; lean_object* x_26; lean_object* x_27; 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_5007____closed__10; +x_26 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__10; x_27 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_17, x_26, x_25); if (lean_obj_tag(x_27) == 0) { @@ -60281,10 +60876,10 @@ lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean x_28 = lean_ctor_get(x_27, 1); lean_inc(x_28); lean_dec(x_27); -x_29 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__11; +x_29 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__11; x_30 = l_Lean_Parser_Term_sufficesDecl___closed__2; -x_31 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__12; -x_32 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__13; +x_31 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__12; +x_32 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__13; x_33 = l_Lean_Parser_registerAlias(x_29, x_30, x_31, x_32, x_6, x_28); if (lean_obj_tag(x_33) == 0) { @@ -60292,7 +60887,7 @@ 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_5007____closed__14; +x_35 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__14; x_36 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_29, x_35, x_34); if (lean_obj_tag(x_36) == 0) { @@ -60300,7 +60895,7 @@ 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_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__15; +x_38 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__15; x_39 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_29, x_38, x_37); if (lean_obj_tag(x_39) == 0) { @@ -60308,10 +60903,10 @@ lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean x_40 = lean_ctor_get(x_39, 1); lean_inc(x_40); lean_dec(x_39); -x_41 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__16; +x_41 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__16; x_42 = l_Lean_Parser_Term_letRecDecls___closed__2; -x_43 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__17; -x_44 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__18; +x_43 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__17; +x_44 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__18; x_45 = l_Lean_Parser_registerAlias(x_41, x_42, x_43, x_44, x_6, x_40); if (lean_obj_tag(x_45) == 0) { @@ -60319,7 +60914,7 @@ lean_object* x_46; lean_object* x_47; lean_object* x_48; x_46 = lean_ctor_get(x_45, 1); lean_inc(x_46); lean_dec(x_45); -x_47 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__19; +x_47 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__19; x_48 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_41, x_47, x_46); if (lean_obj_tag(x_48) == 0) { @@ -60327,7 +60922,7 @@ lean_object* x_49; lean_object* x_50; lean_object* x_51; x_49 = lean_ctor_get(x_48, 1); lean_inc(x_49); lean_dec(x_48); -x_50 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__20; +x_50 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__20; x_51 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_41, x_50, x_49); if (lean_obj_tag(x_51) == 0) { @@ -60335,10 +60930,10 @@ lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean 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_5007____closed__21; +x_53 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__21; x_54 = l_Lean_Parser_Term_hole___closed__2; -x_55 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__22; -x_56 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__23; +x_55 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__22; +x_56 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__23; x_57 = l_Lean_Parser_registerAlias(x_53, x_54, x_55, x_56, x_6, x_52); if (lean_obj_tag(x_57) == 0) { @@ -60346,7 +60941,7 @@ 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_Term___hyg_5007____closed__24; +x_59 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__24; x_60 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_53, x_59, x_58); if (lean_obj_tag(x_60) == 0) { @@ -60354,7 +60949,7 @@ 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_Term___hyg_5007____closed__25; +x_62 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__25; x_63 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_53, x_62, x_61); if (lean_obj_tag(x_63) == 0) { @@ -60362,10 +60957,10 @@ lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean 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_5007____closed__26; +x_65 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__26; x_66 = l_Lean_Parser_Term_syntheticHole___closed__2; -x_67 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__27; -x_68 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__28; +x_67 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__27; +x_68 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__28; x_69 = l_Lean_Parser_registerAlias(x_65, x_66, x_67, x_68, x_6, x_64); if (lean_obj_tag(x_69) == 0) { @@ -60373,7 +60968,7 @@ lean_object* x_70; lean_object* x_71; lean_object* x_72; 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_5007____closed__29; +x_71 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__29; x_72 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_65, x_71, x_70); if (lean_obj_tag(x_72) == 0) { @@ -60381,7 +60976,7 @@ lean_object* x_73; lean_object* x_74; lean_object* x_75; x_73 = lean_ctor_get(x_72, 1); lean_inc(x_73); lean_dec(x_72); -x_74 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__30; +x_74 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__30; x_75 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_65, x_74, x_73); if (lean_obj_tag(x_75) == 0) { @@ -60389,10 +60984,10 @@ lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean x_76 = lean_ctor_get(x_75, 1); lean_inc(x_76); lean_dec(x_75); -x_77 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__31; +x_77 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__31; x_78 = l_Lean_Parser_Term_matchDiscr___closed__2; -x_79 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__32; -x_80 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__33; +x_79 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__32; +x_80 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__33; x_81 = l_Lean_Parser_registerAlias(x_77, x_78, x_79, x_80, x_6, x_76); if (lean_obj_tag(x_81) == 0) { @@ -60400,7 +60995,7 @@ lean_object* x_82; lean_object* x_83; lean_object* x_84; x_82 = lean_ctor_get(x_81, 1); lean_inc(x_82); lean_dec(x_81); -x_83 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__34; +x_83 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__34; x_84 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_77, x_83, x_82); if (lean_obj_tag(x_84) == 0) { @@ -60408,7 +61003,7 @@ 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_5007____closed__35; +x_86 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__35; x_87 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_77, x_86, x_85); if (lean_obj_tag(x_87) == 0) { @@ -60416,10 +61011,10 @@ lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean 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_5007____closed__36; +x_89 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__36; x_90 = l_Lean_Parser_Term_bracketedBinder___closed__2; -x_91 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__37; -x_92 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__38; +x_91 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__37; +x_92 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__38; x_93 = l_Lean_Parser_registerAlias(x_89, x_90, x_91, x_92, x_6, x_88); if (lean_obj_tag(x_93) == 0) { @@ -60427,7 +61022,7 @@ 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_Term___hyg_5007____closed__39; +x_95 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__39; x_96 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_89, x_95, x_94); if (lean_obj_tag(x_96) == 0) { @@ -60435,7 +61030,7 @@ 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_5007____closed__40; +x_98 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__40; x_99 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_89, x_98, x_97); if (lean_obj_tag(x_99) == 0) { @@ -60443,10 +61038,10 @@ 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_Term___hyg_5007____closed__41; +x_101 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__41; x_102 = l_Lean_Parser_Term_attrKind___closed__2; -x_103 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__42; -x_104 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__43; +x_103 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__42; +x_104 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__43; x_105 = l_Lean_Parser_registerAlias(x_101, x_102, x_103, x_104, x_6, x_100); if (lean_obj_tag(x_105) == 0) { @@ -60454,7 +61049,7 @@ 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_Term___hyg_5007____closed__44; +x_107 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__44; x_108 = l_Lean_Parser_registerAliasCore___rarg(x_9, x_101, x_107, x_106); if (lean_obj_tag(x_108) == 0) { @@ -60462,7 +61057,7 @@ 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_Term___hyg_5007____closed__45; +x_110 = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__45; x_111 = l_Lean_Parser_registerAliasCore___rarg(x_13, x_101, x_110, x_109); return x_111; } @@ -62179,6 +62774,8 @@ lean_mark_persistent(l___regBuiltin_Lean_Parser_Term_syntheticHole_parenthesizer res = l___regBuiltin_Lean_Parser_Term_syntheticHole_parenthesizer(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +l_Lean_Parser_Term_binderIdent = _init_l_Lean_Parser_Term_binderIdent(); +lean_mark_persistent(l_Lean_Parser_Term_binderIdent); l_Lean_Parser_Term_sorry___closed__1 = _init_l_Lean_Parser_Term_sorry___closed__1(); lean_mark_persistent(l_Lean_Parser_Term_sorry___closed__1); l_Lean_Parser_Term_sorry___closed__2 = _init_l_Lean_Parser_Term_sorry___closed__2(); @@ -62373,6 +62970,8 @@ l_Lean_Parser_Term_typeAscription___closed__21 = _init_l_Lean_Parser_Term_typeAs lean_mark_persistent(l_Lean_Parser_Term_typeAscription___closed__21); l_Lean_Parser_Term_typeAscription___closed__22 = _init_l_Lean_Parser_Term_typeAscription___closed__22(); lean_mark_persistent(l_Lean_Parser_Term_typeAscription___closed__22); +l_Lean_Parser_Term_typeAscription___closed__23 = _init_l_Lean_Parser_Term_typeAscription___closed__23(); +lean_mark_persistent(l_Lean_Parser_Term_typeAscription___closed__23); l_Lean_Parser_Term_typeAscription = _init_l_Lean_Parser_Term_typeAscription(); lean_mark_persistent(l_Lean_Parser_Term_typeAscription); res = l___regBuiltin_Lean_Parser_Term_typeAscription(lean_io_mk_world()); @@ -62426,6 +63025,10 @@ l_Lean_Parser_Term_typeAscription_formatter___closed__12 = _init_l_Lean_Parser_T lean_mark_persistent(l_Lean_Parser_Term_typeAscription_formatter___closed__12); l_Lean_Parser_Term_typeAscription_formatter___closed__13 = _init_l_Lean_Parser_Term_typeAscription_formatter___closed__13(); lean_mark_persistent(l_Lean_Parser_Term_typeAscription_formatter___closed__13); +l_Lean_Parser_Term_typeAscription_formatter___closed__14 = _init_l_Lean_Parser_Term_typeAscription_formatter___closed__14(); +lean_mark_persistent(l_Lean_Parser_Term_typeAscription_formatter___closed__14); +l_Lean_Parser_Term_typeAscription_formatter___closed__15 = _init_l_Lean_Parser_Term_typeAscription_formatter___closed__15(); +lean_mark_persistent(l_Lean_Parser_Term_typeAscription_formatter___closed__15); l___regBuiltin_Lean_Parser_Term_typeAscription_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Term_typeAscription_formatter___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Term_typeAscription_formatter___closed__1); l___regBuiltin_Lean_Parser_Term_typeAscription_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Term_typeAscription_formatter___closed__2(); @@ -62459,6 +63062,10 @@ l_Lean_Parser_Term_typeAscription_parenthesizer___closed__12 = _init_l_Lean_Pars lean_mark_persistent(l_Lean_Parser_Term_typeAscription_parenthesizer___closed__12); l_Lean_Parser_Term_typeAscription_parenthesizer___closed__13 = _init_l_Lean_Parser_Term_typeAscription_parenthesizer___closed__13(); lean_mark_persistent(l_Lean_Parser_Term_typeAscription_parenthesizer___closed__13); +l_Lean_Parser_Term_typeAscription_parenthesizer___closed__14 = _init_l_Lean_Parser_Term_typeAscription_parenthesizer___closed__14(); +lean_mark_persistent(l_Lean_Parser_Term_typeAscription_parenthesizer___closed__14); +l_Lean_Parser_Term_typeAscription_parenthesizer___closed__15 = _init_l_Lean_Parser_Term_typeAscription_parenthesizer___closed__15(); +lean_mark_persistent(l_Lean_Parser_Term_typeAscription_parenthesizer___closed__15); l___regBuiltin_Lean_Parser_Term_typeAscription_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Term_typeAscription_parenthesizer___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Term_typeAscription_parenthesizer___closed__1); l___regBuiltin_Lean_Parser_Term_typeAscription_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Term_typeAscription_parenthesizer___closed__2(); @@ -62698,6 +63305,10 @@ l_Lean_Parser_Term_anonymousCtor___closed__12 = _init_l_Lean_Parser_Term_anonymo lean_mark_persistent(l_Lean_Parser_Term_anonymousCtor___closed__12); l_Lean_Parser_Term_anonymousCtor___closed__13 = _init_l_Lean_Parser_Term_anonymousCtor___closed__13(); lean_mark_persistent(l_Lean_Parser_Term_anonymousCtor___closed__13); +l_Lean_Parser_Term_anonymousCtor___closed__14 = _init_l_Lean_Parser_Term_anonymousCtor___closed__14(); +lean_mark_persistent(l_Lean_Parser_Term_anonymousCtor___closed__14); +l_Lean_Parser_Term_anonymousCtor___closed__15 = _init_l_Lean_Parser_Term_anonymousCtor___closed__15(); +lean_mark_persistent(l_Lean_Parser_Term_anonymousCtor___closed__15); l_Lean_Parser_Term_anonymousCtor = _init_l_Lean_Parser_Term_anonymousCtor(); lean_mark_persistent(l_Lean_Parser_Term_anonymousCtor); res = l___regBuiltin_Lean_Parser_Term_anonymousCtor(lean_io_mk_world()); @@ -62739,6 +63350,10 @@ l_Lean_Parser_Term_anonymousCtor_formatter___closed__6 = _init_l_Lean_Parser_Ter lean_mark_persistent(l_Lean_Parser_Term_anonymousCtor_formatter___closed__6); l_Lean_Parser_Term_anonymousCtor_formatter___closed__7 = _init_l_Lean_Parser_Term_anonymousCtor_formatter___closed__7(); lean_mark_persistent(l_Lean_Parser_Term_anonymousCtor_formatter___closed__7); +l_Lean_Parser_Term_anonymousCtor_formatter___closed__8 = _init_l_Lean_Parser_Term_anonymousCtor_formatter___closed__8(); +lean_mark_persistent(l_Lean_Parser_Term_anonymousCtor_formatter___closed__8); +l_Lean_Parser_Term_anonymousCtor_formatter___closed__9 = _init_l_Lean_Parser_Term_anonymousCtor_formatter___closed__9(); +lean_mark_persistent(l_Lean_Parser_Term_anonymousCtor_formatter___closed__9); l___regBuiltin_Lean_Parser_Term_anonymousCtor_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Term_anonymousCtor_formatter___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Term_anonymousCtor_formatter___closed__1); l___regBuiltin_Lean_Parser_Term_anonymousCtor_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Term_anonymousCtor_formatter___closed__2(); @@ -62760,6 +63375,10 @@ l_Lean_Parser_Term_anonymousCtor_parenthesizer___closed__6 = _init_l_Lean_Parser lean_mark_persistent(l_Lean_Parser_Term_anonymousCtor_parenthesizer___closed__6); l_Lean_Parser_Term_anonymousCtor_parenthesizer___closed__7 = _init_l_Lean_Parser_Term_anonymousCtor_parenthesizer___closed__7(); lean_mark_persistent(l_Lean_Parser_Term_anonymousCtor_parenthesizer___closed__7); +l_Lean_Parser_Term_anonymousCtor_parenthesizer___closed__8 = _init_l_Lean_Parser_Term_anonymousCtor_parenthesizer___closed__8(); +lean_mark_persistent(l_Lean_Parser_Term_anonymousCtor_parenthesizer___closed__8); +l_Lean_Parser_Term_anonymousCtor_parenthesizer___closed__9 = _init_l_Lean_Parser_Term_anonymousCtor_parenthesizer___closed__9(); +lean_mark_persistent(l_Lean_Parser_Term_anonymousCtor_parenthesizer___closed__9); l___regBuiltin_Lean_Parser_Term_anonymousCtor_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Term_anonymousCtor_parenthesizer___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Term_anonymousCtor_parenthesizer___closed__1); l___regBuiltin_Lean_Parser_Term_anonymousCtor_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Term_anonymousCtor_parenthesizer___closed__2(); @@ -62773,6 +63392,10 @@ l_Lean_Parser_Term_optIdent___closed__2 = _init_l_Lean_Parser_Term_optIdent___cl lean_mark_persistent(l_Lean_Parser_Term_optIdent___closed__2); l_Lean_Parser_Term_optIdent___closed__3 = _init_l_Lean_Parser_Term_optIdent___closed__3(); lean_mark_persistent(l_Lean_Parser_Term_optIdent___closed__3); +l_Lean_Parser_Term_optIdent___closed__4 = _init_l_Lean_Parser_Term_optIdent___closed__4(); +lean_mark_persistent(l_Lean_Parser_Term_optIdent___closed__4); +l_Lean_Parser_Term_optIdent___closed__5 = _init_l_Lean_Parser_Term_optIdent___closed__5(); +lean_mark_persistent(l_Lean_Parser_Term_optIdent___closed__5); l_Lean_Parser_Term_optIdent = _init_l_Lean_Parser_Term_optIdent(); lean_mark_persistent(l_Lean_Parser_Term_optIdent); l_Lean_Parser_Term_fromTerm___closed__1 = _init_l_Lean_Parser_Term_fromTerm___closed__1(); @@ -62817,6 +63440,18 @@ l_Lean_Parser_Term_sufficesDecl___closed__8 = _init_l_Lean_Parser_Term_sufficesD lean_mark_persistent(l_Lean_Parser_Term_sufficesDecl___closed__8); l_Lean_Parser_Term_sufficesDecl___closed__9 = _init_l_Lean_Parser_Term_sufficesDecl___closed__9(); lean_mark_persistent(l_Lean_Parser_Term_sufficesDecl___closed__9); +l_Lean_Parser_Term_sufficesDecl___closed__10 = _init_l_Lean_Parser_Term_sufficesDecl___closed__10(); +lean_mark_persistent(l_Lean_Parser_Term_sufficesDecl___closed__10); +l_Lean_Parser_Term_sufficesDecl___closed__11 = _init_l_Lean_Parser_Term_sufficesDecl___closed__11(); +lean_mark_persistent(l_Lean_Parser_Term_sufficesDecl___closed__11); +l_Lean_Parser_Term_sufficesDecl___closed__12 = _init_l_Lean_Parser_Term_sufficesDecl___closed__12(); +lean_mark_persistent(l_Lean_Parser_Term_sufficesDecl___closed__12); +l_Lean_Parser_Term_sufficesDecl___closed__13 = _init_l_Lean_Parser_Term_sufficesDecl___closed__13(); +lean_mark_persistent(l_Lean_Parser_Term_sufficesDecl___closed__13); +l_Lean_Parser_Term_sufficesDecl___closed__14 = _init_l_Lean_Parser_Term_sufficesDecl___closed__14(); +lean_mark_persistent(l_Lean_Parser_Term_sufficesDecl___closed__14); +l_Lean_Parser_Term_sufficesDecl___closed__15 = _init_l_Lean_Parser_Term_sufficesDecl___closed__15(); +lean_mark_persistent(l_Lean_Parser_Term_sufficesDecl___closed__15); l_Lean_Parser_Term_sufficesDecl = _init_l_Lean_Parser_Term_sufficesDecl(); lean_mark_persistent(l_Lean_Parser_Term_sufficesDecl); l_Lean_Parser_Term_suffices___closed__1 = _init_l_Lean_Parser_Term_suffices___closed__1(); @@ -62865,10 +63500,6 @@ lean_mark_persistent(l___regBuiltin_Lean_Parser_Term_suffices_declRange___closed res = l___regBuiltin_Lean_Parser_Term_suffices_declRange(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l_Lean_Parser_Term_optIdent_formatter___closed__1 = _init_l_Lean_Parser_Term_optIdent_formatter___closed__1(); -lean_mark_persistent(l_Lean_Parser_Term_optIdent_formatter___closed__1); -l_Lean_Parser_Term_optIdent_formatter___closed__2 = _init_l_Lean_Parser_Term_optIdent_formatter___closed__2(); -lean_mark_persistent(l_Lean_Parser_Term_optIdent_formatter___closed__2); l_Lean_Parser_Term_fromTerm_formatter___closed__1 = _init_l_Lean_Parser_Term_fromTerm_formatter___closed__1(); lean_mark_persistent(l_Lean_Parser_Term_fromTerm_formatter___closed__1); l_Lean_Parser_Term_fromTerm_formatter___closed__2 = _init_l_Lean_Parser_Term_fromTerm_formatter___closed__2(); @@ -62911,6 +63542,16 @@ l_Lean_Parser_Term_sufficesDecl_formatter___closed__7 = _init_l_Lean_Parser_Term lean_mark_persistent(l_Lean_Parser_Term_sufficesDecl_formatter___closed__7); l_Lean_Parser_Term_sufficesDecl_formatter___closed__8 = _init_l_Lean_Parser_Term_sufficesDecl_formatter___closed__8(); lean_mark_persistent(l_Lean_Parser_Term_sufficesDecl_formatter___closed__8); +l_Lean_Parser_Term_sufficesDecl_formatter___closed__9 = _init_l_Lean_Parser_Term_sufficesDecl_formatter___closed__9(); +lean_mark_persistent(l_Lean_Parser_Term_sufficesDecl_formatter___closed__9); +l_Lean_Parser_Term_sufficesDecl_formatter___closed__10 = _init_l_Lean_Parser_Term_sufficesDecl_formatter___closed__10(); +lean_mark_persistent(l_Lean_Parser_Term_sufficesDecl_formatter___closed__10); +l_Lean_Parser_Term_sufficesDecl_formatter___closed__11 = _init_l_Lean_Parser_Term_sufficesDecl_formatter___closed__11(); +lean_mark_persistent(l_Lean_Parser_Term_sufficesDecl_formatter___closed__11); +l_Lean_Parser_Term_sufficesDecl_formatter___closed__12 = _init_l_Lean_Parser_Term_sufficesDecl_formatter___closed__12(); +lean_mark_persistent(l_Lean_Parser_Term_sufficesDecl_formatter___closed__12); +l_Lean_Parser_Term_sufficesDecl_formatter___closed__13 = _init_l_Lean_Parser_Term_sufficesDecl_formatter___closed__13(); +lean_mark_persistent(l_Lean_Parser_Term_sufficesDecl_formatter___closed__13); l___regBuiltin_Lean_Parser_Term_sufficesDecl_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Term_sufficesDecl_formatter___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Term_sufficesDecl_formatter___closed__1); l___regBuiltin_Lean_Parser_Term_sufficesDecl_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Term_sufficesDecl_formatter___closed__2(); @@ -62949,10 +63590,6 @@ lean_mark_persistent(l___regBuiltin_Lean_Parser_Term_suffices_formatter___closed res = l___regBuiltin_Lean_Parser_Term_suffices_formatter(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l_Lean_Parser_Term_optIdent_parenthesizer___closed__1 = _init_l_Lean_Parser_Term_optIdent_parenthesizer___closed__1(); -lean_mark_persistent(l_Lean_Parser_Term_optIdent_parenthesizer___closed__1); -l_Lean_Parser_Term_optIdent_parenthesizer___closed__2 = _init_l_Lean_Parser_Term_optIdent_parenthesizer___closed__2(); -lean_mark_persistent(l_Lean_Parser_Term_optIdent_parenthesizer___closed__2); l_Lean_Parser_Term_fromTerm_parenthesizer___closed__1 = _init_l_Lean_Parser_Term_fromTerm_parenthesizer___closed__1(); lean_mark_persistent(l_Lean_Parser_Term_fromTerm_parenthesizer___closed__1); l_Lean_Parser_Term_fromTerm_parenthesizer___closed__2 = _init_l_Lean_Parser_Term_fromTerm_parenthesizer___closed__2(); @@ -62995,6 +63632,16 @@ l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__7 = _init_l_Lean_Parser_ lean_mark_persistent(l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__7); l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__8 = _init_l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__8(); lean_mark_persistent(l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__8); +l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__9 = _init_l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__9(); +lean_mark_persistent(l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__9); +l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__10 = _init_l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__10(); +lean_mark_persistent(l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__10); +l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__11 = _init_l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__11(); +lean_mark_persistent(l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__11); +l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__12 = _init_l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__12(); +lean_mark_persistent(l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__12); +l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__13 = _init_l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__13(); +lean_mark_persistent(l_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__13); l___regBuiltin_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__1); l___regBuiltin_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Term_sufficesDecl_parenthesizer___closed__2(); @@ -63159,10 +63806,6 @@ l_Lean_Parser_Term_structInstLVal___closed__13 = _init_l_Lean_Parser_Term_struct lean_mark_persistent(l_Lean_Parser_Term_structInstLVal___closed__13); l_Lean_Parser_Term_structInstLVal___closed__14 = _init_l_Lean_Parser_Term_structInstLVal___closed__14(); lean_mark_persistent(l_Lean_Parser_Term_structInstLVal___closed__14); -l_Lean_Parser_Term_structInstLVal___closed__15 = _init_l_Lean_Parser_Term_structInstLVal___closed__15(); -lean_mark_persistent(l_Lean_Parser_Term_structInstLVal___closed__15); -l_Lean_Parser_Term_structInstLVal___closed__16 = _init_l_Lean_Parser_Term_structInstLVal___closed__16(); -lean_mark_persistent(l_Lean_Parser_Term_structInstLVal___closed__16); l_Lean_Parser_Term_structInstLVal = _init_l_Lean_Parser_Term_structInstLVal(); lean_mark_persistent(l_Lean_Parser_Term_structInstLVal); l_Lean_Parser_Term_structInstField___closed__1 = _init_l_Lean_Parser_Term_structInstField___closed__1(); @@ -63802,8 +64445,6 @@ lean_mark_persistent(l___regBuiltin_Lean_Parser_Term_inaccessible_parenthesizer_ res = l___regBuiltin_Lean_Parser_Term_inaccessible_parenthesizer(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l_Lean_Parser_Term_binderIdent = _init_l_Lean_Parser_Term_binderIdent(); -lean_mark_persistent(l_Lean_Parser_Term_binderIdent); 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(); @@ -64035,8 +64676,6 @@ l_Lean_Parser_Term_explicitBinder_formatter___closed__3 = _init_l_Lean_Parser_Te lean_mark_persistent(l_Lean_Parser_Term_explicitBinder_formatter___closed__3); l_Lean_Parser_Term_explicitBinder_formatter___closed__4 = _init_l_Lean_Parser_Term_explicitBinder_formatter___closed__4(); lean_mark_persistent(l_Lean_Parser_Term_explicitBinder_formatter___closed__4); -l_Lean_Parser_Term_explicitBinder_formatter___closed__5 = _init_l_Lean_Parser_Term_explicitBinder_formatter___closed__5(); -lean_mark_persistent(l_Lean_Parser_Term_explicitBinder_formatter___closed__5); l_Lean_Parser_Term_strictImplicitLeftBracket_formatter___closed__1 = _init_l_Lean_Parser_Term_strictImplicitLeftBracket_formatter___closed__1(); lean_mark_persistent(l_Lean_Parser_Term_strictImplicitLeftBracket_formatter___closed__1); l_Lean_Parser_Term_strictImplicitLeftBracket_formatter___closed__2 = _init_l_Lean_Parser_Term_strictImplicitLeftBracket_formatter___closed__2(); @@ -64061,6 +64700,10 @@ l_Lean_Parser_Term_strictImplicitBinder_formatter___closed__3 = _init_l_Lean_Par lean_mark_persistent(l_Lean_Parser_Term_strictImplicitBinder_formatter___closed__3); l_Lean_Parser_Term_implicitBinder_formatter___closed__1 = _init_l_Lean_Parser_Term_implicitBinder_formatter___closed__1(); lean_mark_persistent(l_Lean_Parser_Term_implicitBinder_formatter___closed__1); +l_Lean_Parser_Term_optIdent_formatter___closed__1 = _init_l_Lean_Parser_Term_optIdent_formatter___closed__1(); +lean_mark_persistent(l_Lean_Parser_Term_optIdent_formatter___closed__1); +l_Lean_Parser_Term_optIdent_formatter___closed__2 = _init_l_Lean_Parser_Term_optIdent_formatter___closed__2(); +lean_mark_persistent(l_Lean_Parser_Term_optIdent_formatter___closed__2); l_Lean_Parser_Term_instBinder_formatter___closed__1 = _init_l_Lean_Parser_Term_instBinder_formatter___closed__1(); lean_mark_persistent(l_Lean_Parser_Term_instBinder_formatter___closed__1); l_Lean_Parser_Term_instBinder_formatter___closed__2 = _init_l_Lean_Parser_Term_instBinder_formatter___closed__2(); @@ -64077,6 +64720,8 @@ l_Lean_Parser_Term_instBinder_formatter___closed__7 = _init_l_Lean_Parser_Term_i lean_mark_persistent(l_Lean_Parser_Term_instBinder_formatter___closed__7); l_Lean_Parser_Term_instBinder_formatter___closed__8 = _init_l_Lean_Parser_Term_instBinder_formatter___closed__8(); lean_mark_persistent(l_Lean_Parser_Term_instBinder_formatter___closed__8); +l_Lean_Parser_Term_instBinder_formatter___closed__9 = _init_l_Lean_Parser_Term_instBinder_formatter___closed__9(); +lean_mark_persistent(l_Lean_Parser_Term_instBinder_formatter___closed__9); l___regBuiltin_Lean_Parser_Term_instBinder_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Term_instBinder_formatter___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Term_instBinder_formatter___closed__1); l___regBuiltin_Lean_Parser_Term_instBinder_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Term_instBinder_formatter___closed__2(); @@ -64147,8 +64792,6 @@ l_Lean_Parser_Term_explicitBinder_parenthesizer___closed__4 = _init_l_Lean_Parse lean_mark_persistent(l_Lean_Parser_Term_explicitBinder_parenthesizer___closed__4); l_Lean_Parser_Term_explicitBinder_parenthesizer___closed__5 = _init_l_Lean_Parser_Term_explicitBinder_parenthesizer___closed__5(); lean_mark_persistent(l_Lean_Parser_Term_explicitBinder_parenthesizer___closed__5); -l_Lean_Parser_Term_explicitBinder_parenthesizer___closed__6 = _init_l_Lean_Parser_Term_explicitBinder_parenthesizer___closed__6(); -lean_mark_persistent(l_Lean_Parser_Term_explicitBinder_parenthesizer___closed__6); l_Lean_Parser_Term_strictImplicitLeftBracket_parenthesizer___closed__1 = _init_l_Lean_Parser_Term_strictImplicitLeftBracket_parenthesizer___closed__1(); lean_mark_persistent(l_Lean_Parser_Term_strictImplicitLeftBracket_parenthesizer___closed__1); l_Lean_Parser_Term_strictImplicitLeftBracket_parenthesizer___closed__2 = _init_l_Lean_Parser_Term_strictImplicitLeftBracket_parenthesizer___closed__2(); @@ -64173,6 +64816,10 @@ l_Lean_Parser_Term_strictImplicitBinder_parenthesizer___closed__3 = _init_l_Lean lean_mark_persistent(l_Lean_Parser_Term_strictImplicitBinder_parenthesizer___closed__3); l_Lean_Parser_Term_implicitBinder_parenthesizer___closed__1 = _init_l_Lean_Parser_Term_implicitBinder_parenthesizer___closed__1(); lean_mark_persistent(l_Lean_Parser_Term_implicitBinder_parenthesizer___closed__1); +l_Lean_Parser_Term_optIdent_parenthesizer___closed__1 = _init_l_Lean_Parser_Term_optIdent_parenthesizer___closed__1(); +lean_mark_persistent(l_Lean_Parser_Term_optIdent_parenthesizer___closed__1); +l_Lean_Parser_Term_optIdent_parenthesizer___closed__2 = _init_l_Lean_Parser_Term_optIdent_parenthesizer___closed__2(); +lean_mark_persistent(l_Lean_Parser_Term_optIdent_parenthesizer___closed__2); l_Lean_Parser_Term_instBinder_parenthesizer___closed__1 = _init_l_Lean_Parser_Term_instBinder_parenthesizer___closed__1(); lean_mark_persistent(l_Lean_Parser_Term_instBinder_parenthesizer___closed__1); l_Lean_Parser_Term_instBinder_parenthesizer___closed__2 = _init_l_Lean_Parser_Term_instBinder_parenthesizer___closed__2(); @@ -64185,6 +64832,8 @@ l_Lean_Parser_Term_instBinder_parenthesizer___closed__5 = _init_l_Lean_Parser_Te lean_mark_persistent(l_Lean_Parser_Term_instBinder_parenthesizer___closed__5); l_Lean_Parser_Term_instBinder_parenthesizer___closed__6 = _init_l_Lean_Parser_Term_instBinder_parenthesizer___closed__6(); lean_mark_persistent(l_Lean_Parser_Term_instBinder_parenthesizer___closed__6); +l_Lean_Parser_Term_instBinder_parenthesizer___closed__7 = _init_l_Lean_Parser_Term_instBinder_parenthesizer___closed__7(); +lean_mark_persistent(l_Lean_Parser_Term_instBinder_parenthesizer___closed__7); l___regBuiltin_Lean_Parser_Term_instBinder_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Term_instBinder_parenthesizer___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Term_instBinder_parenthesizer___closed__1); l___regBuiltin_Lean_Parser_Term_instBinder_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Term_instBinder_parenthesizer___closed__2(); @@ -64525,6 +65174,10 @@ l_Lean_Parser_Term_match___closed__16 = _init_l_Lean_Parser_Term_match___closed_ lean_mark_persistent(l_Lean_Parser_Term_match___closed__16); l_Lean_Parser_Term_match___closed__17 = _init_l_Lean_Parser_Term_match___closed__17(); lean_mark_persistent(l_Lean_Parser_Term_match___closed__17); +l_Lean_Parser_Term_match___closed__18 = _init_l_Lean_Parser_Term_match___closed__18(); +lean_mark_persistent(l_Lean_Parser_Term_match___closed__18); +l_Lean_Parser_Term_match___closed__19 = _init_l_Lean_Parser_Term_match___closed__19(); +lean_mark_persistent(l_Lean_Parser_Term_match___closed__19); l_Lean_Parser_Term_match = _init_l_Lean_Parser_Term_match(); lean_mark_persistent(l_Lean_Parser_Term_match); res = l___regBuiltin_Lean_Parser_Term_match(lean_io_mk_world()); @@ -64683,6 +65336,8 @@ l_Lean_Parser_Term_match_formatter___closed__12 = _init_l_Lean_Parser_Term_match lean_mark_persistent(l_Lean_Parser_Term_match_formatter___closed__12); l_Lean_Parser_Term_match_formatter___closed__13 = _init_l_Lean_Parser_Term_match_formatter___closed__13(); lean_mark_persistent(l_Lean_Parser_Term_match_formatter___closed__13); +l_Lean_Parser_Term_match_formatter___closed__14 = _init_l_Lean_Parser_Term_match_formatter___closed__14(); +lean_mark_persistent(l_Lean_Parser_Term_match_formatter___closed__14); l___regBuiltin_Lean_Parser_Term_match_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Term_match_formatter___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Term_match_formatter___closed__1); l___regBuiltin_Lean_Parser_Term_match_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Term_match_formatter___closed__2(); @@ -64821,6 +65476,8 @@ l_Lean_Parser_Term_match_parenthesizer___closed__12 = _init_l_Lean_Parser_Term_m lean_mark_persistent(l_Lean_Parser_Term_match_parenthesizer___closed__12); l_Lean_Parser_Term_match_parenthesizer___closed__13 = _init_l_Lean_Parser_Term_match_parenthesizer___closed__13(); lean_mark_persistent(l_Lean_Parser_Term_match_parenthesizer___closed__13); +l_Lean_Parser_Term_match_parenthesizer___closed__14 = _init_l_Lean_Parser_Term_match_parenthesizer___closed__14(); +lean_mark_persistent(l_Lean_Parser_Term_match_parenthesizer___closed__14); l___regBuiltin_Lean_Parser_Term_match_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Term_match_parenthesizer___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Term_match_parenthesizer___closed__1); l___regBuiltin_Lean_Parser_Term_match_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Term_match_parenthesizer___closed__2(); @@ -64981,8 +65638,6 @@ l_Lean_Parser_Term_basicFun___closed__13 = _init_l_Lean_Parser_Term_basicFun___c lean_mark_persistent(l_Lean_Parser_Term_basicFun___closed__13); l_Lean_Parser_Term_basicFun___closed__14 = _init_l_Lean_Parser_Term_basicFun___closed__14(); lean_mark_persistent(l_Lean_Parser_Term_basicFun___closed__14); -l_Lean_Parser_Term_basicFun___closed__15 = _init_l_Lean_Parser_Term_basicFun___closed__15(); -lean_mark_persistent(l_Lean_Parser_Term_basicFun___closed__15); l_Lean_Parser_Term_basicFun = _init_l_Lean_Parser_Term_basicFun(); lean_mark_persistent(l_Lean_Parser_Term_basicFun); l_Lean_Parser_Term_fun___closed__1 = _init_l_Lean_Parser_Term_fun___closed__1(); @@ -65087,8 +65742,6 @@ l_Lean_Parser_Term_basicFun_formatter___closed__9 = _init_l_Lean_Parser_Term_bas lean_mark_persistent(l_Lean_Parser_Term_basicFun_formatter___closed__9); l_Lean_Parser_Term_basicFun_formatter___closed__10 = _init_l_Lean_Parser_Term_basicFun_formatter___closed__10(); lean_mark_persistent(l_Lean_Parser_Term_basicFun_formatter___closed__10); -l_Lean_Parser_Term_basicFun_formatter___closed__11 = _init_l_Lean_Parser_Term_basicFun_formatter___closed__11(); -lean_mark_persistent(l_Lean_Parser_Term_basicFun_formatter___closed__11); l___regBuiltin_Lean_Parser_Term_basicFun_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Term_basicFun_formatter___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Term_basicFun_formatter___closed__1); l___regBuiltin_Lean_Parser_Term_basicFun_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Term_basicFun_formatter___closed__2(); @@ -65173,8 +65826,6 @@ l_Lean_Parser_Term_basicFun_parenthesizer___closed__9 = _init_l_Lean_Parser_Term lean_mark_persistent(l_Lean_Parser_Term_basicFun_parenthesizer___closed__9); l_Lean_Parser_Term_basicFun_parenthesizer___closed__10 = _init_l_Lean_Parser_Term_basicFun_parenthesizer___closed__10(); lean_mark_persistent(l_Lean_Parser_Term_basicFun_parenthesizer___closed__10); -l_Lean_Parser_Term_basicFun_parenthesizer___closed__11 = _init_l_Lean_Parser_Term_basicFun_parenthesizer___closed__11(); -lean_mark_persistent(l_Lean_Parser_Term_basicFun_parenthesizer___closed__11); l___regBuiltin_Lean_Parser_Term_basicFun_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Term_basicFun_parenthesizer___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Term_basicFun_parenthesizer___closed__1); l___regBuiltin_Lean_Parser_Term_basicFun_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Term_basicFun_parenthesizer___closed__2(); @@ -65235,6 +65886,12 @@ l_Lean_Parser_Term_withAnonymousAntiquot___closed__10 = _init_l_Lean_Parser_Term lean_mark_persistent(l_Lean_Parser_Term_withAnonymousAntiquot___closed__10); l_Lean_Parser_Term_withAnonymousAntiquot___closed__11 = _init_l_Lean_Parser_Term_withAnonymousAntiquot___closed__11(); lean_mark_persistent(l_Lean_Parser_Term_withAnonymousAntiquot___closed__11); +l_Lean_Parser_Term_withAnonymousAntiquot___closed__12 = _init_l_Lean_Parser_Term_withAnonymousAntiquot___closed__12(); +lean_mark_persistent(l_Lean_Parser_Term_withAnonymousAntiquot___closed__12); +l_Lean_Parser_Term_withAnonymousAntiquot___closed__13 = _init_l_Lean_Parser_Term_withAnonymousAntiquot___closed__13(); +lean_mark_persistent(l_Lean_Parser_Term_withAnonymousAntiquot___closed__13); +l_Lean_Parser_Term_withAnonymousAntiquot___closed__14 = _init_l_Lean_Parser_Term_withAnonymousAntiquot___closed__14(); +lean_mark_persistent(l_Lean_Parser_Term_withAnonymousAntiquot___closed__14); l_Lean_Parser_Term_withAnonymousAntiquot = _init_l_Lean_Parser_Term_withAnonymousAntiquot(); lean_mark_persistent(l_Lean_Parser_Term_withAnonymousAntiquot); l_Lean_Parser_Term_leading__parser___closed__1 = _init_l_Lean_Parser_Term_leading__parser___closed__1(); @@ -65259,8 +65916,6 @@ l_Lean_Parser_Term_leading__parser___closed__10 = _init_l_Lean_Parser_Term_leadi lean_mark_persistent(l_Lean_Parser_Term_leading__parser___closed__10); l_Lean_Parser_Term_leading__parser___closed__11 = _init_l_Lean_Parser_Term_leading__parser___closed__11(); lean_mark_persistent(l_Lean_Parser_Term_leading__parser___closed__11); -l_Lean_Parser_Term_leading__parser___closed__12 = _init_l_Lean_Parser_Term_leading__parser___closed__12(); -lean_mark_persistent(l_Lean_Parser_Term_leading__parser___closed__12); l_Lean_Parser_Term_leading__parser = _init_l_Lean_Parser_Term_leading__parser(); lean_mark_persistent(l_Lean_Parser_Term_leading__parser); res = l___regBuiltin_Lean_Parser_Term_leading__parser(lean_io_mk_world()); @@ -65303,6 +65958,10 @@ l_Lean_Parser_Term_withAnonymousAntiquot_formatter___closed__6 = _init_l_Lean_Pa lean_mark_persistent(l_Lean_Parser_Term_withAnonymousAntiquot_formatter___closed__6); l_Lean_Parser_Term_withAnonymousAntiquot_formatter___closed__7 = _init_l_Lean_Parser_Term_withAnonymousAntiquot_formatter___closed__7(); lean_mark_persistent(l_Lean_Parser_Term_withAnonymousAntiquot_formatter___closed__7); +l_Lean_Parser_Term_withAnonymousAntiquot_formatter___closed__8 = _init_l_Lean_Parser_Term_withAnonymousAntiquot_formatter___closed__8(); +lean_mark_persistent(l_Lean_Parser_Term_withAnonymousAntiquot_formatter___closed__8); +l_Lean_Parser_Term_withAnonymousAntiquot_formatter___closed__9 = _init_l_Lean_Parser_Term_withAnonymousAntiquot_formatter___closed__9(); +lean_mark_persistent(l_Lean_Parser_Term_withAnonymousAntiquot_formatter___closed__9); l___regBuiltin_Lean_Parser_Term_withAnonymousAntiquot_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Term_withAnonymousAntiquot_formatter___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Term_withAnonymousAntiquot_formatter___closed__1); l___regBuiltin_Lean_Parser_Term_withAnonymousAntiquot_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Term_withAnonymousAntiquot_formatter___closed__2(); @@ -65353,6 +66012,10 @@ l_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___closed__6 = _init_l_Lea lean_mark_persistent(l_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___closed__6); l_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___closed__7 = _init_l_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___closed__7(); lean_mark_persistent(l_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___closed__7); +l_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___closed__8 = _init_l_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___closed__8(); +lean_mark_persistent(l_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___closed__8); +l_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___closed__9 = _init_l_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___closed__9(); +lean_mark_persistent(l_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___closed__9); l___regBuiltin_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___closed__1); l___regBuiltin_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Term_withAnonymousAntiquot_parenthesizer___closed__2(); @@ -65403,8 +66066,6 @@ l_Lean_Parser_Term_trailing__parser___closed__9 = _init_l_Lean_Parser_Term_trail lean_mark_persistent(l_Lean_Parser_Term_trailing__parser___closed__9); l_Lean_Parser_Term_trailing__parser___closed__10 = _init_l_Lean_Parser_Term_trailing__parser___closed__10(); lean_mark_persistent(l_Lean_Parser_Term_trailing__parser___closed__10); -l_Lean_Parser_Term_trailing__parser___closed__11 = _init_l_Lean_Parser_Term_trailing__parser___closed__11(); -lean_mark_persistent(l_Lean_Parser_Term_trailing__parser___closed__11); l_Lean_Parser_Term_trailing__parser = _init_l_Lean_Parser_Term_trailing__parser(); lean_mark_persistent(l_Lean_Parser_Term_trailing__parser); res = l___regBuiltin_Lean_Parser_Term_trailing__parser(lean_io_mk_world()); @@ -66433,6 +67094,8 @@ l_Lean_Parser_Term_haveDecl___closed__7 = _init_l_Lean_Parser_Term_haveDecl___cl lean_mark_persistent(l_Lean_Parser_Term_haveDecl___closed__7); l_Lean_Parser_Term_haveDecl___closed__8 = _init_l_Lean_Parser_Term_haveDecl___closed__8(); lean_mark_persistent(l_Lean_Parser_Term_haveDecl___closed__8); +l_Lean_Parser_Term_haveDecl___closed__9 = _init_l_Lean_Parser_Term_haveDecl___closed__9(); +lean_mark_persistent(l_Lean_Parser_Term_haveDecl___closed__9); l_Lean_Parser_Term_haveDecl = _init_l_Lean_Parser_Term_haveDecl(); lean_mark_persistent(l_Lean_Parser_Term_haveDecl); l_Lean_Parser_Term_have___closed__1 = _init_l_Lean_Parser_Term_have___closed__1(); @@ -66455,8 +67118,6 @@ l_Lean_Parser_Term_have___closed__9 = _init_l_Lean_Parser_Term_have___closed__9( lean_mark_persistent(l_Lean_Parser_Term_have___closed__9); l_Lean_Parser_Term_have___closed__10 = _init_l_Lean_Parser_Term_have___closed__10(); lean_mark_persistent(l_Lean_Parser_Term_have___closed__10); -l_Lean_Parser_Term_have___closed__11 = _init_l_Lean_Parser_Term_have___closed__11(); -lean_mark_persistent(l_Lean_Parser_Term_have___closed__11); l_Lean_Parser_Term_have = _init_l_Lean_Parser_Term_have(); lean_mark_persistent(l_Lean_Parser_Term_have); res = l___regBuiltin_Lean_Parser_Term_have(lean_io_mk_world()); @@ -66523,6 +67184,8 @@ l_Lean_Parser_Term_haveDecl_formatter___closed__3 = _init_l_Lean_Parser_Term_hav lean_mark_persistent(l_Lean_Parser_Term_haveDecl_formatter___closed__3); l_Lean_Parser_Term_haveDecl_formatter___closed__4 = _init_l_Lean_Parser_Term_haveDecl_formatter___closed__4(); lean_mark_persistent(l_Lean_Parser_Term_haveDecl_formatter___closed__4); +l_Lean_Parser_Term_haveDecl_formatter___closed__5 = _init_l_Lean_Parser_Term_haveDecl_formatter___closed__5(); +lean_mark_persistent(l_Lean_Parser_Term_haveDecl_formatter___closed__5); l___regBuiltin_Lean_Parser_Term_haveDecl_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Term_haveDecl_formatter___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Term_haveDecl_formatter___closed__1); l___regBuiltin_Lean_Parser_Term_haveDecl_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Term_haveDecl_formatter___closed__2(); @@ -66593,6 +67256,8 @@ l_Lean_Parser_Term_haveDecl_parenthesizer___closed__3 = _init_l_Lean_Parser_Term lean_mark_persistent(l_Lean_Parser_Term_haveDecl_parenthesizer___closed__3); l_Lean_Parser_Term_haveDecl_parenthesizer___closed__4 = _init_l_Lean_Parser_Term_haveDecl_parenthesizer___closed__4(); lean_mark_persistent(l_Lean_Parser_Term_haveDecl_parenthesizer___closed__4); +l_Lean_Parser_Term_haveDecl_parenthesizer___closed__5 = _init_l_Lean_Parser_Term_haveDecl_parenthesizer___closed__5(); +lean_mark_persistent(l_Lean_Parser_Term_haveDecl_parenthesizer___closed__5); l___regBuiltin_Lean_Parser_Term_haveDecl_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Term_haveDecl_parenthesizer___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Term_haveDecl_parenthesizer___closed__1); l___regBuiltin_Lean_Parser_Term_haveDecl_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Term_haveDecl_parenthesizer___closed__2(); @@ -66719,6 +67384,10 @@ l_Lean_Parser_Term_attributes___closed__11 = _init_l_Lean_Parser_Term_attributes lean_mark_persistent(l_Lean_Parser_Term_attributes___closed__11); l_Lean_Parser_Term_attributes___closed__12 = _init_l_Lean_Parser_Term_attributes___closed__12(); lean_mark_persistent(l_Lean_Parser_Term_attributes___closed__12); +l_Lean_Parser_Term_attributes___closed__13 = _init_l_Lean_Parser_Term_attributes___closed__13(); +lean_mark_persistent(l_Lean_Parser_Term_attributes___closed__13); +l_Lean_Parser_Term_attributes___closed__14 = _init_l_Lean_Parser_Term_attributes___closed__14(); +lean_mark_persistent(l_Lean_Parser_Term_attributes___closed__14); l_Lean_Parser_Term_attributes = _init_l_Lean_Parser_Term_attributes(); lean_mark_persistent(l_Lean_Parser_Term_attributes); l_Lean_Parser_Term_letRecDecl___closed__1 = _init_l_Lean_Parser_Term_letRecDecl___closed__1(); @@ -66904,6 +67573,8 @@ l_Lean_Parser_Term_attributes_formatter___closed__6 = _init_l_Lean_Parser_Term_a lean_mark_persistent(l_Lean_Parser_Term_attributes_formatter___closed__6); l_Lean_Parser_Term_attributes_formatter___closed__7 = _init_l_Lean_Parser_Term_attributes_formatter___closed__7(); lean_mark_persistent(l_Lean_Parser_Term_attributes_formatter___closed__7); +l_Lean_Parser_Term_attributes_formatter___closed__8 = _init_l_Lean_Parser_Term_attributes_formatter___closed__8(); +lean_mark_persistent(l_Lean_Parser_Term_attributes_formatter___closed__8); l___regBuiltin_Lean_Parser_Term_attributes_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Term_attributes_formatter___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Term_attributes_formatter___closed__1); l___regBuiltin_Lean_Parser_Term_attributes_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Term_attributes_formatter___closed__2(); @@ -67059,6 +67730,8 @@ l_Lean_Parser_Term_attributes_parenthesizer___closed__6 = _init_l_Lean_Parser_Te lean_mark_persistent(l_Lean_Parser_Term_attributes_parenthesizer___closed__6); l_Lean_Parser_Term_attributes_parenthesizer___closed__7 = _init_l_Lean_Parser_Term_attributes_parenthesizer___closed__7(); lean_mark_persistent(l_Lean_Parser_Term_attributes_parenthesizer___closed__7); +l_Lean_Parser_Term_attributes_parenthesizer___closed__8 = _init_l_Lean_Parser_Term_attributes_parenthesizer___closed__8(); +lean_mark_persistent(l_Lean_Parser_Term_attributes_parenthesizer___closed__8); l___regBuiltin_Lean_Parser_Term_attributes_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Term_attributes_parenthesizer___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Term_attributes_parenthesizer___closed__1); l___regBuiltin_Lean_Parser_Term_attributes_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Term_attributes_parenthesizer___closed__2(); @@ -67139,6 +67812,10 @@ l_Lean_Parser_Term_whereDecls_formatter___closed__8 = _init_l_Lean_Parser_Term_w lean_mark_persistent(l_Lean_Parser_Term_whereDecls_formatter___closed__8); l_Lean_Parser_Term_whereDecls_formatter___closed__9 = _init_l_Lean_Parser_Term_whereDecls_formatter___closed__9(); lean_mark_persistent(l_Lean_Parser_Term_whereDecls_formatter___closed__9); +l_Lean_Parser_Term_whereDecls_formatter___closed__10 = _init_l_Lean_Parser_Term_whereDecls_formatter___closed__10(); +lean_mark_persistent(l_Lean_Parser_Term_whereDecls_formatter___closed__10); +l_Lean_Parser_Term_whereDecls_formatter___closed__11 = _init_l_Lean_Parser_Term_whereDecls_formatter___closed__11(); +lean_mark_persistent(l_Lean_Parser_Term_whereDecls_formatter___closed__11); l___regBuiltin_Lean_Parser_Term_whereDecls_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Term_whereDecls_formatter___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Term_whereDecls_formatter___closed__1); l___regBuiltin_Lean_Parser_Term_whereDecls_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Term_whereDecls_formatter___closed__2(); @@ -67158,6 +67835,10 @@ l_Lean_Parser_Term_whereDecls_parenthesizer___closed__5 = _init_l_Lean_Parser_Te lean_mark_persistent(l_Lean_Parser_Term_whereDecls_parenthesizer___closed__5); l_Lean_Parser_Term_whereDecls_parenthesizer___closed__6 = _init_l_Lean_Parser_Term_whereDecls_parenthesizer___closed__6(); lean_mark_persistent(l_Lean_Parser_Term_whereDecls_parenthesizer___closed__6); +l_Lean_Parser_Term_whereDecls_parenthesizer___closed__7 = _init_l_Lean_Parser_Term_whereDecls_parenthesizer___closed__7(); +lean_mark_persistent(l_Lean_Parser_Term_whereDecls_parenthesizer___closed__7); +l_Lean_Parser_Term_whereDecls_parenthesizer___closed__8 = _init_l_Lean_Parser_Term_whereDecls_parenthesizer___closed__8(); +lean_mark_persistent(l_Lean_Parser_Term_whereDecls_parenthesizer___closed__8); l___regBuiltin_Lean_Parser_Term_whereDecls_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Term_whereDecls_parenthesizer___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Term_whereDecls_parenthesizer___closed__1); l___regBuiltin_Lean_Parser_Term_whereDecls_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Term_whereDecls_parenthesizer___closed__2(); @@ -67185,6 +67866,8 @@ l_Lean_Parser_Term_whereDecls___closed__9 = _init_l_Lean_Parser_Term_whereDecls_ lean_mark_persistent(l_Lean_Parser_Term_whereDecls___closed__9); l_Lean_Parser_Term_whereDecls___closed__10 = _init_l_Lean_Parser_Term_whereDecls___closed__10(); lean_mark_persistent(l_Lean_Parser_Term_whereDecls___closed__10); +l_Lean_Parser_Term_whereDecls___closed__11 = _init_l_Lean_Parser_Term_whereDecls___closed__11(); +lean_mark_persistent(l_Lean_Parser_Term_whereDecls___closed__11); l_Lean_Parser_Term_whereDecls = _init_l_Lean_Parser_Term_whereDecls(); lean_mark_persistent(l_Lean_Parser_Term_whereDecls); l_Lean_Parser_Term_matchAltsWhereDecls_formatter___closed__1 = _init_l_Lean_Parser_Term_matchAltsWhereDecls_formatter___closed__1(); @@ -67329,6 +68012,8 @@ l_Lean_Parser_Term_binrel___closed__11 = _init_l_Lean_Parser_Term_binrel___close lean_mark_persistent(l_Lean_Parser_Term_binrel___closed__11); l_Lean_Parser_Term_binrel___closed__12 = _init_l_Lean_Parser_Term_binrel___closed__12(); lean_mark_persistent(l_Lean_Parser_Term_binrel___closed__12); +l_Lean_Parser_Term_binrel___closed__13 = _init_l_Lean_Parser_Term_binrel___closed__13(); +lean_mark_persistent(l_Lean_Parser_Term_binrel___closed__13); l_Lean_Parser_Term_binrel = _init_l_Lean_Parser_Term_binrel(); lean_mark_persistent(l_Lean_Parser_Term_binrel); res = l___regBuiltin_Lean_Parser_Term_binrel(lean_io_mk_world()); @@ -67386,6 +68071,8 @@ l_Lean_Parser_Term_binrel_parenthesizer___closed__6 = _init_l_Lean_Parser_Term_b lean_mark_persistent(l_Lean_Parser_Term_binrel_parenthesizer___closed__6); l_Lean_Parser_Term_binrel_parenthesizer___closed__7 = _init_l_Lean_Parser_Term_binrel_parenthesizer___closed__7(); lean_mark_persistent(l_Lean_Parser_Term_binrel_parenthesizer___closed__7); +l_Lean_Parser_Term_binrel_parenthesizer___closed__8 = _init_l_Lean_Parser_Term_binrel_parenthesizer___closed__8(); +lean_mark_persistent(l_Lean_Parser_Term_binrel_parenthesizer___closed__8); l___regBuiltin_Lean_Parser_Term_binrel_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Term_binrel_parenthesizer___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Term_binrel_parenthesizer___closed__1); l___regBuiltin_Lean_Parser_Term_binrel_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Term_binrel_parenthesizer___closed__2(); @@ -67628,8 +68315,6 @@ l_Lean_Parser_Term_unop___closed__9 = _init_l_Lean_Parser_Term_unop___closed__9( lean_mark_persistent(l_Lean_Parser_Term_unop___closed__9); l_Lean_Parser_Term_unop___closed__10 = _init_l_Lean_Parser_Term_unop___closed__10(); lean_mark_persistent(l_Lean_Parser_Term_unop___closed__10); -l_Lean_Parser_Term_unop___closed__11 = _init_l_Lean_Parser_Term_unop___closed__11(); -lean_mark_persistent(l_Lean_Parser_Term_unop___closed__11); l_Lean_Parser_Term_unop = _init_l_Lean_Parser_Term_unop(); lean_mark_persistent(l_Lean_Parser_Term_unop); res = l___regBuiltin_Lean_Parser_Term_unop(lean_io_mk_world()); @@ -67679,8 +68364,6 @@ l_Lean_Parser_Term_unop_parenthesizer___closed__4 = _init_l_Lean_Parser_Term_uno lean_mark_persistent(l_Lean_Parser_Term_unop_parenthesizer___closed__4); l_Lean_Parser_Term_unop_parenthesizer___closed__5 = _init_l_Lean_Parser_Term_unop_parenthesizer___closed__5(); lean_mark_persistent(l_Lean_Parser_Term_unop_parenthesizer___closed__5); -l_Lean_Parser_Term_unop_parenthesizer___closed__6 = _init_l_Lean_Parser_Term_unop_parenthesizer___closed__6(); -lean_mark_persistent(l_Lean_Parser_Term_unop_parenthesizer___closed__6); l___regBuiltin_Lean_Parser_Term_unop_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Term_unop_parenthesizer___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Term_unop_parenthesizer___closed__1); l___regBuiltin_Lean_Parser_Term_unop_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Term_unop_parenthesizer___closed__2(); @@ -67966,8 +68649,6 @@ l_Lean_Parser_Term_withDeclName_formatter___closed__5 = _init_l_Lean_Parser_Term lean_mark_persistent(l_Lean_Parser_Term_withDeclName_formatter___closed__5); l_Lean_Parser_Term_withDeclName_formatter___closed__6 = _init_l_Lean_Parser_Term_withDeclName_formatter___closed__6(); lean_mark_persistent(l_Lean_Parser_Term_withDeclName_formatter___closed__6); -l_Lean_Parser_Term_withDeclName_formatter___closed__7 = _init_l_Lean_Parser_Term_withDeclName_formatter___closed__7(); -lean_mark_persistent(l_Lean_Parser_Term_withDeclName_formatter___closed__7); l___regBuiltin_Lean_Parser_Term_withDeclName_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Term_withDeclName_formatter___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Term_withDeclName_formatter___closed__1); l___regBuiltin_Lean_Parser_Term_withDeclName_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Term_withDeclName_formatter___closed__2(); @@ -68353,6 +69034,8 @@ l_Lean_Parser_Term_clear_formatter___closed__5 = _init_l_Lean_Parser_Term_clear_ lean_mark_persistent(l_Lean_Parser_Term_clear_formatter___closed__5); l_Lean_Parser_Term_clear_formatter___closed__6 = _init_l_Lean_Parser_Term_clear_formatter___closed__6(); lean_mark_persistent(l_Lean_Parser_Term_clear_formatter___closed__6); +l_Lean_Parser_Term_clear_formatter___closed__7 = _init_l_Lean_Parser_Term_clear_formatter___closed__7(); +lean_mark_persistent(l_Lean_Parser_Term_clear_formatter___closed__7); l___regBuiltin_Lean_Parser_Term_clear_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Term_clear_formatter___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Term_clear_formatter___closed__1); l___regBuiltin_Lean_Parser_Term_clear_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Term_clear_formatter___closed__2(); @@ -68372,6 +69055,8 @@ l_Lean_Parser_Term_clear_parenthesizer___closed__5 = _init_l_Lean_Parser_Term_cl lean_mark_persistent(l_Lean_Parser_Term_clear_parenthesizer___closed__5); l_Lean_Parser_Term_clear_parenthesizer___closed__6 = _init_l_Lean_Parser_Term_clear_parenthesizer___closed__6(); lean_mark_persistent(l_Lean_Parser_Term_clear_parenthesizer___closed__6); +l_Lean_Parser_Term_clear_parenthesizer___closed__7 = _init_l_Lean_Parser_Term_clear_parenthesizer___closed__7(); +lean_mark_persistent(l_Lean_Parser_Term_clear_parenthesizer___closed__7); l___regBuiltin_Lean_Parser_Term_clear_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Term_clear_parenthesizer___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Term_clear_parenthesizer___closed__1); l___regBuiltin_Lean_Parser_Term_clear_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Term_clear_parenthesizer___closed__2(); @@ -68904,6 +69589,10 @@ l_Lean_Parser_Term_ellipsis___closed__8 = _init_l_Lean_Parser_Term_ellipsis___cl lean_mark_persistent(l_Lean_Parser_Term_ellipsis___closed__8); l_Lean_Parser_Term_ellipsis___closed__9 = _init_l_Lean_Parser_Term_ellipsis___closed__9(); lean_mark_persistent(l_Lean_Parser_Term_ellipsis___closed__9); +l_Lean_Parser_Term_ellipsis___closed__10 = _init_l_Lean_Parser_Term_ellipsis___closed__10(); +lean_mark_persistent(l_Lean_Parser_Term_ellipsis___closed__10); +l_Lean_Parser_Term_ellipsis___closed__11 = _init_l_Lean_Parser_Term_ellipsis___closed__11(); +lean_mark_persistent(l_Lean_Parser_Term_ellipsis___closed__11); l_Lean_Parser_Term_ellipsis = _init_l_Lean_Parser_Term_ellipsis(); lean_mark_persistent(l_Lean_Parser_Term_ellipsis); l_Lean_Parser_Term_argument___closed__1 = _init_l_Lean_Parser_Term_argument___closed__1(); @@ -68983,6 +69672,8 @@ l_Lean_Parser_Term_ellipsis_formatter___closed__3 = _init_l_Lean_Parser_Term_ell lean_mark_persistent(l_Lean_Parser_Term_ellipsis_formatter___closed__3); l_Lean_Parser_Term_ellipsis_formatter___closed__4 = _init_l_Lean_Parser_Term_ellipsis_formatter___closed__4(); lean_mark_persistent(l_Lean_Parser_Term_ellipsis_formatter___closed__4); +l_Lean_Parser_Term_ellipsis_formatter___closed__5 = _init_l_Lean_Parser_Term_ellipsis_formatter___closed__5(); +lean_mark_persistent(l_Lean_Parser_Term_ellipsis_formatter___closed__5); l___regBuiltin_Lean_Parser_Term_ellipsis_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Term_ellipsis_formatter___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Term_ellipsis_formatter___closed__1); l___regBuiltin_Lean_Parser_Term_ellipsis_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Term_ellipsis_formatter___closed__2(); @@ -69034,6 +69725,8 @@ l_Lean_Parser_Term_ellipsis_parenthesizer___closed__3 = _init_l_Lean_Parser_Term lean_mark_persistent(l_Lean_Parser_Term_ellipsis_parenthesizer___closed__3); l_Lean_Parser_Term_ellipsis_parenthesizer___closed__4 = _init_l_Lean_Parser_Term_ellipsis_parenthesizer___closed__4(); lean_mark_persistent(l_Lean_Parser_Term_ellipsis_parenthesizer___closed__4); +l_Lean_Parser_Term_ellipsis_parenthesizer___closed__5 = _init_l_Lean_Parser_Term_ellipsis_parenthesizer___closed__5(); +lean_mark_persistent(l_Lean_Parser_Term_ellipsis_parenthesizer___closed__5); l___regBuiltin_Lean_Parser_Term_ellipsis_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Term_ellipsis_parenthesizer___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Term_ellipsis_parenthesizer___closed__1); l___regBuiltin_Lean_Parser_Term_ellipsis_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Term_ellipsis_parenthesizer___closed__2(); @@ -70011,6 +70704,8 @@ l_Lean_Parser_Term_stateRefT___closed__9 = _init_l_Lean_Parser_Term_stateRefT___ lean_mark_persistent(l_Lean_Parser_Term_stateRefT___closed__9); l_Lean_Parser_Term_stateRefT___closed__10 = _init_l_Lean_Parser_Term_stateRefT___closed__10(); lean_mark_persistent(l_Lean_Parser_Term_stateRefT___closed__10); +l_Lean_Parser_Term_stateRefT___closed__11 = _init_l_Lean_Parser_Term_stateRefT___closed__11(); +lean_mark_persistent(l_Lean_Parser_Term_stateRefT___closed__11); l_Lean_Parser_Term_stateRefT = _init_l_Lean_Parser_Term_stateRefT(); lean_mark_persistent(l_Lean_Parser_Term_stateRefT); res = l___regBuiltin_Lean_Parser_Term_stateRefT(lean_io_mk_world()); @@ -70062,6 +70757,8 @@ l_Lean_Parser_Term_stateRefT_formatter___closed__5 = _init_l_Lean_Parser_Term_st lean_mark_persistent(l_Lean_Parser_Term_stateRefT_formatter___closed__5); l_Lean_Parser_Term_stateRefT_formatter___closed__6 = _init_l_Lean_Parser_Term_stateRefT_formatter___closed__6(); lean_mark_persistent(l_Lean_Parser_Term_stateRefT_formatter___closed__6); +l_Lean_Parser_Term_stateRefT_formatter___closed__7 = _init_l_Lean_Parser_Term_stateRefT_formatter___closed__7(); +lean_mark_persistent(l_Lean_Parser_Term_stateRefT_formatter___closed__7); l___regBuiltin_Lean_Parser_Term_stateRefT_formatter___closed__1 = _init_l___regBuiltin_Lean_Parser_Term_stateRefT_formatter___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Term_stateRefT_formatter___closed__1); l___regBuiltin_Lean_Parser_Term_stateRefT_formatter___closed__2 = _init_l___regBuiltin_Lean_Parser_Term_stateRefT_formatter___closed__2(); @@ -70100,6 +70797,8 @@ l_Lean_Parser_Term_stateRefT_parenthesizer___closed__5 = _init_l_Lean_Parser_Ter lean_mark_persistent(l_Lean_Parser_Term_stateRefT_parenthesizer___closed__5); l_Lean_Parser_Term_stateRefT_parenthesizer___closed__6 = _init_l_Lean_Parser_Term_stateRefT_parenthesizer___closed__6(); lean_mark_persistent(l_Lean_Parser_Term_stateRefT_parenthesizer___closed__6); +l_Lean_Parser_Term_stateRefT_parenthesizer___closed__7 = _init_l_Lean_Parser_Term_stateRefT_parenthesizer___closed__7(); +lean_mark_persistent(l_Lean_Parser_Term_stateRefT_parenthesizer___closed__7); l___regBuiltin_Lean_Parser_Term_stateRefT_parenthesizer___closed__1 = _init_l___regBuiltin_Lean_Parser_Term_stateRefT_parenthesizer___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Parser_Term_stateRefT_parenthesizer___closed__1); l___regBuiltin_Lean_Parser_Term_stateRefT_parenthesizer___closed__2 = _init_l___regBuiltin_Lean_Parser_Term_stateRefT_parenthesizer___closed__2(); @@ -70467,97 +71166,97 @@ lean_mark_persistent(l___regBuiltin_Lean_Parser_Tactic_quotSeq_parenthesizer___c res = l___regBuiltin_Lean_Parser_Tactic_quotSeq_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_5007____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__1(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__1); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__2 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__2(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__2); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__3 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__3(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__3); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__4 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__4(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__4); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__5 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__5(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__5); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__6 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__6(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__6); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__7 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__7(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__7); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__8 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__8(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__8); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__9 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__9(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__9); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__10 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__10(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__10); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__11 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__11(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__11); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__12 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__12(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__12); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__13 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__13(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__13); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__14 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__14(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__14); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__15 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__15(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__15); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__16 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__16(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__16); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__17 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__17(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__17); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__18 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__18(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__18); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__19 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__19(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__19); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__20 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__20(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__20); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__21 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__21(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__21); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__22 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__22(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__22); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__23 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__23(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__23); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__24 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__24(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__24); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__25 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__25(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__25); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__26 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__26(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__26); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__27 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__27(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__27); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__28 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__28(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__28); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__29 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__29(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__29); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__30 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__30(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__30); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__31 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__31(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__31); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__32 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__32(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__32); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__33 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__33(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__33); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__34 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__34(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__34); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__35 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__35(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__35); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__36 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__36(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__36); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__37 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__37(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__37); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__38 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__38(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__38); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__39 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__39(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__39); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__40 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__40(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__40); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__41 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__41(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__41); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__42 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__42(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__42); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__43 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__43(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__43); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__44 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__44(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__44); -l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__45 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__45(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007____closed__45); -res = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5007_(lean_io_mk_world()); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__1(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__1); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__2 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__2(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__2); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__3 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__3(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__3); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__4 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__4(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__4); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__5 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__5(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__5); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__6 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__6(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__6); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__7 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__7(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__7); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__8 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__8(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__8); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__9 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__9(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__9); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__10 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__10(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__10); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__11 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__11(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__11); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__12 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__12(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__12); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__13 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__13(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__13); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__14 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__14(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__14); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__15 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__15(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__15); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__16 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__16(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__16); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__17 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__17(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__17); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__18 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__18(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__18); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__19 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__19(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__19); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__20 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__20(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__20); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__21 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__21(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__21); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__22 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__22(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__22); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__23 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__23(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__23); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__24 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__24(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__24); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__25 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__25(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__25); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__26 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__26(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__26); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__27 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__27(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__27); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__28 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__28(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__28); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__29 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__29(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__29); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__30 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__30(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__30); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__31 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__31(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__31); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__32 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__32(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__32); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__33 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__33(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__33); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__34 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__34(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__34); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__35 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__35(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__35); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__36 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__36(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__36); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__37 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__37(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__37); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__38 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__38(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__38); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__39 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__39(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__39); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__40 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__40(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__40); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__41 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__41(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__41); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__42 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__42(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__42); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__43 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__43(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__43); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__44 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__44(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__44); +l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__45 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__45(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098____closed__45); +res = l_Lean_Parser_initFn____x40_Lean_Parser_Term___hyg_5098_(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/Types.c b/stage0/stdlib/Lean/Parser/Types.c index 7356f8471744..8e056e252c5d 100644 --- a/stage0/stdlib/Lean/Parser/Types.c +++ b/stage0/stdlib/Lean/Parser/Types.c @@ -47,7 +47,6 @@ size_t lean_uint64_to_usize(uint64_t); LEAN_EXPORT lean_object* l_Lean_Parser_Error_toString(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_instBEqParserCacheKey; LEAN_EXPORT lean_object* l_Lean_Parser_adaptCacheableContextFn(lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l_List_beq___at___private_Init_Meta_0__Lean_Syntax_beqPreresolved____x40_Init_Meta___hyg_3131____spec__1(lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__beqOption____x40_Init_Data_Option_Basic___hyg_1116____at___private_Lean_Parser_Types_0__Lean_Parser_beqCacheableParserContext____x40_Lean_Parser_Types___hyg_235____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Parser_SyntaxNodeKindSet_insert___spec__2(lean_object*, size_t, size_t, lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); @@ -60,6 +59,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_SyntaxStack_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_mkIdent(lean_object*, lean_object*, lean_object*); lean_object* lean_mk_array(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ParserModuleContext_currNamespace___default; +uint8_t l_List_beq___at___private_Init_Meta_0__Lean_Syntax_beqPreresolved____x40_Init_Meta___hyg_3135____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Parser_instInhabitedParserInfo___lambda__2___closed__2; LEAN_EXPORT lean_object* l_Lean_AssocList_contains___at_Lean_Parser_withCacheFn___spec__2___boxed(lean_object*, lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); @@ -1344,7 +1344,7 @@ return x_8; else { uint8_t x_9; -x_9 = l_List_beq___at___private_Init_Meta_0__Lean_Syntax_beqPreresolved____x40_Init_Meta___hyg_3131____spec__1(x_4, x_6); +x_9 = l_List_beq___at___private_Init_Meta_0__Lean_Syntax_beqPreresolved____x40_Init_Meta___hyg_3135____spec__1(x_4, x_6); return x_9; } } @@ -1646,7 +1646,7 @@ lean_dec(x_1); x_4 = l_Lean_Parser_instInhabitedInputContext___closed__2; x_5 = lean_string_dec_eq(x_2, x_4); x_6 = lean_box(0); -x_7 = l_List_beq___at___private_Init_Meta_0__Lean_Syntax_beqPreresolved____x40_Init_Meta___hyg_3131____spec__1(x_3, x_6); +x_7 = l_List_beq___at___private_Init_Meta_0__Lean_Syntax_beqPreresolved____x40_Init_Meta___hyg_3135____spec__1(x_3, x_6); if (x_5 == 0) { lean_object* x_8; diff --git a/stage0/stdlib/Lean/PrettyPrinter.c b/stage0/stdlib/Lean/PrettyPrinter.c index 4fd303d7e70e..9f4bcb388a87 100644 --- a/stage0/stdlib/Lean/PrettyPrinter.c +++ b/stage0/stdlib/Lean/PrettyPrinter.c @@ -63,7 +63,6 @@ LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_0__Lean_PrettyPrinter_wi lean_object* l_List_mapTR_loop___at_Lean_mkConstWithLevelParams___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_0__Lean_PrettyPrinter_noContext(lean_object*); static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_1222____closed__4; -lean_object* l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_1364____closed__7; lean_object* l_instHashableArray___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_ppUsing(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -131,6 +130,7 @@ lean_object* l_Lean_sanitizeSyntax(lean_object*, lean_object*); extern lean_object* l_Lean_PrettyPrinter_combinatorFormatterAttribute; lean_object* l_Lean_PrettyPrinter_delabCore(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_1222____lambda__2(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PPContext_runMetaM___rarg___closed__2; extern lean_object* l_Lean_PrettyPrinter_combinatorParenthesizerAttribute; static lean_object* l_Lean_PPContext_runCoreM___rarg___closed__17; @@ -977,7 +977,7 @@ x_14 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_ppUsing___lambda__1), 7, lean_closure_set(x_14, 0, x_2); lean_closure_set(x_14, 1, x_1); x_15 = l_Lean_PPContext_runCoreM___rarg___closed__1; -x_16 = l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__3___rarg(x_13, x_15, x_14, x_3, x_4, x_5, x_6, x_7); +x_16 = l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2___rarg(x_13, x_15, x_14, x_3, x_4, x_5, x_6, x_7); return x_16; } } @@ -1131,7 +1131,7 @@ lean_closure_set(x_15, 0, x_1); lean_closure_set(x_15, 1, x_2); lean_closure_set(x_15, 2, x_3); x_16 = l_Lean_PPContext_runCoreM___rarg___closed__1; -x_17 = l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__3___rarg(x_14, x_16, x_15, x_4, x_5, x_6, x_7, x_8); +x_17 = l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2___rarg(x_14, x_16, x_15, x_4, x_5, x_6, x_7, x_8); return x_17; } } diff --git a/stage0/stdlib/Lean/PrettyPrinter/Delaborator/Basic.c b/stage0/stdlib/Lean/PrettyPrinter/Delaborator/Basic.c index c534f3dd5562..8648ac65abee 100644 --- a/stage0/stdlib/Lean/PrettyPrinter/Delaborator/Basic.c +++ b/stage0/stdlib/Lean/PrettyPrinter/Delaborator/Basic.c @@ -168,7 +168,6 @@ lean_object* l_Lean_Attribute_Builtin_getIdent(lean_object*, lean_object*, lean_ lean_object* lean_st_mk_ref(lean_object*, lean_object*); lean_object* l_Lean_getPPProofs___boxed(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_instAlternativeDelabM___closed__2; -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_instMonadStateOfHoleIteratorDelabM___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delab(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -251,6 +250,7 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_instMonadQuotationDelabM___ uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_failure___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_instMonadStateOfHoleIteratorDelabM___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_getExprKind___closed__10; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_failure___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_PrettyPrinter_Delaborator_withBindingBodyUnusedName___spec__3(lean_object*); @@ -9029,7 +9029,7 @@ return x_14; else { lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_15 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_3, x_6, x_7, x_8, x_9, x_10); +x_15 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_3, x_6, x_7, x_8, x_9, x_10); x_16 = lean_ctor_get(x_15, 0); lean_inc(x_16); x_17 = lean_ctor_get(x_15, 1); @@ -9087,7 +9087,7 @@ return x_30; else { lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_31 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_3, x_6, x_7, x_29, x_9, x_10); +x_31 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_3, x_6, x_7, x_29, x_9, x_10); x_32 = lean_ctor_get(x_31, 0); lean_inc(x_32); x_33 = lean_ctor_get(x_31, 1); diff --git a/stage0/stdlib/Lean/PrettyPrinter/Delaborator/TopDownAnalyze.c b/stage0/stdlib/Lean/PrettyPrinter/Delaborator/TopDownAnalyze.c index 14d02c249c40..606a6764a571 100644 --- a/stage0/stdlib/Lean/PrettyPrinter/Delaborator/TopDownAnalyze.c +++ b/stage0/stdlib/Lean/PrettyPrinter/Delaborator/TopDownAnalyze.c @@ -18,6 +18,7 @@ static lean_object* l_Lean_getPPAnalysisLetVarType___closed__2; static lean_object* l_Lean_getPPAnalysisHole___closed__2; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore_maybeSetExplicit___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore_hBinOpHeuristic(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_inspectOutParams___closed__1; static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_replaceLPsWithVars___lambda__1___closed__4; lean_object* l___private_Init_Util_0__outOfBounds___rarg(lean_object*); @@ -36,6 +37,7 @@ LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at_Lean_PrettyPrinter_Delaborato LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_analyzeProj___spec__7(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_TopDownAnalyze_analyze___closed__3; LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_analyzeLet___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_profiler; static lean_object* l_Lean_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownAnalyze___hyg_212____closed__4; LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownAnalyze___hyg_332_(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_isHBinOp___lambda__1___closed__19; @@ -44,7 +46,6 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_isHBinOp___lambda__1___closed__10; lean_object* lean_mk_empty_array_with_capacity(lean_object*); lean_object* l_Lean_mkAppN(lean_object*, lean_object*); -static lean_object* l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___closed__3; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_analyzeLam___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_SubExpr_withLetValue___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_analyzeLet___spec__27___closed__1; LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownAnalyze___hyg_252_(lean_object*); @@ -54,6 +55,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_getExpr___at_L LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_annotateBoolAt___lambda__1(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_mkFreshLevelMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); +static lean_object* l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__3___closed__4; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore_analyzeArg___lambda__5(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_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore_collectBottomUps___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*); static lean_object* l_Lean_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownAnalyze___hyg_332____closed__5; @@ -92,7 +94,6 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze___cl LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownAnalyze___hyg_92_(lean_object*); uint8_t l_Lean_Level_hasParam(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withNaryArg___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore_applyFunBinderHeuristic___spec__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_withOptProfile___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__3___closed__1; LEAN_EXPORT uint8_t l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_isHBinOp(lean_object*); uint8_t l_Lean_Expr_isAppOfArity(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withLetVarType___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_analyzeLet___spec__14___closed__2; @@ -104,6 +105,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_topDownAnalyze___lambd LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownAnalyze___hyg_4166_(lean_object*); uint8_t l_Lean_Expr_isApp(lean_object*); static lean_object* l_Lean_getPPAnalysisNamedArg___closed__1; +double l_Lean_trace_profiler_threshold_getSecs(lean_object*); static lean_object* l_Lean_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownAnalyze___hyg_452____closed__5; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_isSimpleHOFun(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownAnalyze___hyg_8____closed__8; @@ -136,6 +138,7 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_anal LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_analyzeMData___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_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_inspectOutParams(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_withLetBody___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_analyzeLet___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_MonadExcept_ofExcept___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__4___boxed(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___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_panic___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_analyzeLet___spec__20(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_TopDownAnalyze_analyze_analyzeLam___closed__1; @@ -170,6 +173,7 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_isHBinOp___l LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_getPos___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore_analyzeArg___spec__3___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_getPPAnalyzeTypeAscriptions___boxed(lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore___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*); +uint8_t lean_float_decLt(double, double); static lean_object* l_Lean_getPPAnalysisSkip___closed__1; LEAN_EXPORT uint8_t l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_isHBinOp___lambda__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___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*); @@ -184,7 +188,6 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze LEAN_EXPORT lean_object* l_Lean_getPPAnalysisNeedsType___boxed(lean_object*); static lean_object* l_Lean_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownAnalyze___hyg_132____closed__5; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withOptProfile___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_analyzeApp___lambda__1(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_annotateBoolAt___closed__4; LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_analyzeLet___spec__38(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -205,6 +208,7 @@ LEAN_EXPORT uint8_t l_Lean_AssocList_contains___at_Lean_PrettyPrinter_Delaborato uint8_t l___private_Lean_Util_Trace_0__Lean_checkTraceOption(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_analyzeConst___closed__7; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore_applyFunBinderHeuristic_core___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*); +static lean_object* l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__3___closed__1; static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_isHBinOp___lambda__1___closed__13; LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_analyzeLet___spec__10___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*); static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_isHigherOrder___closed__1; @@ -226,6 +230,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze static lean_object* l_Lean_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownAnalyze___hyg_252____closed__1; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore_maybeSetExplicit___spec__5___boxed(lean_object**); uint8_t l_Lean_Expr_isAtomic(lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, double, 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__9(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_replaceLPsWithVars___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownAnalyze___hyg_292_(lean_object*); @@ -258,7 +263,6 @@ size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT uint8_t l_Lean_getPPAnalyzeTypeAscriptions(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_instMonadWithReaderOfSubExprAnalyzeM(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore_collectHigherOrders(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___closed__1; LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_analyzeProj___spec__2(lean_object*, 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_lambdaTelescopeImp___rarg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_analyzeProj___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -301,12 +305,12 @@ static lean_object* l_Lean_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownA LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_analyzeSort___rarg(lean_object*); static lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_analyzeConst___spec__1___closed__2; static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_isHBinOp___lambda__1___closed__21; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__4(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownAnalyze___hyg_12870____closed__15; lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withType___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_registerInternalExceptionId(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore_analyzeArg___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_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___closed__4; static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_isHBinOp___lambda__1___closed__42; lean_object* l_Lean_Meta_SavedState_restore(lean_object*, lean_object*, lean_object*, 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*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -333,6 +337,7 @@ LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_PrettyPrinter_Delaborator_Top LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_annotateBoolAt(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownAnalyze___hyg_452____closed__2; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore_collectBottomUps(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___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*, lean_object*, lean_object*); extern uint8_t l_Lean_instInhabitedBinderInfo; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore_collectBottomUps___lambda__1___boxed(lean_object*); lean_object* l_ReaderT_instMonadLiftReaderT(lean_object*, lean_object*, lean_object*); @@ -347,7 +352,6 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_instInhabitedContext___closed__3; 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*); static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_isHBinOp___lambda__1___closed__33; -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownAnalyze___hyg_452____closed__1; static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_isSubstLike___closed__2; uint8_t l_List_isEmpty___rarg(lean_object*); @@ -359,7 +363,6 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_anal LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_getExpr___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_maybeAddBlockImplicit___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze___closed__6; LEAN_EXPORT uint8_t l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_isSubstLike(lean_object*); -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_getPPAnalyzeOmitMax(lean_object*); static lean_object* l_Lean_getPPAnalysisLetVarType___closed__1; static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_analyzeApp___closed__3; @@ -412,7 +415,6 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingDom static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_isTrivialBottomUp___closed__5; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_hasMVarAtCurrDepth(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_TopDownAnalyze_analyzeAppStagedCore_collectHigherOrders___spec__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_object*, lean_object*, lean_object*); -lean_object* l_Lean_getBoolOption___at_Lean_Meta_processPostponed___spec__4(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_analyzeLet___spec__10___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_panic___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_analyzeLet___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_PrettyPrinter_Delaborator_TopDownAnalyze_isFunLike___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -433,6 +435,7 @@ LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_TopDownAnal LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_annotateBool(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_checkpointDefEq___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_checkpointDefEq___spec__1___closed__4; static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_analyzeApp___closed__4; +uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_topDownAnalyze___lambda__3___closed__4; static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_isHBinOp___lambda__1___closed__25; static lean_object* l_Lean_PrettyPrinter_Delaborator_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownAnalyze___hyg_12870____closed__7; @@ -443,7 +446,6 @@ LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_PrettyPrinter_Delaborator_T lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownAnalyze___hyg_8____closed__1; lean_object* l_Lean_Meta_getResetPostponed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_withOptProfile___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__3___closed__2; LEAN_EXPORT uint8_t l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_hasMVarAtCurrDepth___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_annotateBoolAt___spec__5(lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownAnalyze___hyg_12870____closed__11; @@ -451,6 +453,7 @@ LEAN_EXPORT uint8_t l_Lean_getPPAnalyzeExplicitHoles(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore(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_analyzeLet___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_TopDownAnalyze_isHBinOp___lambda__1___closed__15; +static lean_object* l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__3___closed__3; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore_analyzeArg___lambda__5___boxed(lean_object**); LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_analyzeLet___spec__26(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_saveState___rarg(lean_object*, lean_object*, lean_object*, lean_object*); @@ -477,7 +480,6 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_isHBinOp___l static lean_object* l_Lean_getPPAnalysisNeedsType___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_analyzeLet___spec__10(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_analyzeApp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static double l_Lean_withSeconds___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__4___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_tryUnify___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getPPAnalyzeTrustOfScientific___boxed(lean_object*); uint8_t l_Lean_getPPInstanceTypes(lean_object*); @@ -526,6 +528,7 @@ LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_PrettyPrinter_Delaborator_Top static lean_object* l_Lean_PrettyPrinter_Delaborator_topDownAnalyze___lambda__3___closed__1; static lean_object* l_Lean_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownAnalyze___hyg_8____closed__3; lean_object* l_StateT_lift___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_processPostponed___spec__4(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_panic___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_analyzeMData___spec__11(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_getExpr___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore_hBinOpHeuristic___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownAnalyze___hyg_212____closed__1; @@ -566,6 +569,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_returnsPi___lambda__1_ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_hasLevelMVarAtCurrDepth___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_replaceLPsWithVars___spec__8(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getPPAnalyzeTrustOfNat___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_withSeconds___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); lean_object* l_Lean_MVarId_getDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_analyzeConst___spec__1___closed__3; @@ -578,6 +582,8 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_isHBinOp___l uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withNewMCtxDepth___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__1(lean_object*); lean_object* lean_float_to_string(double); +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___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*); static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_analyzeApp___closed__1; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_isHBinOp___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_analyzeMData(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -596,6 +602,7 @@ LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delabora LEAN_EXPORT uint8_t l_Lean_getPPAnalysisHole(lean_object*); lean_object* l_Array_shrink___rarg(lean_object*, lean_object*); lean_object* l_Lean_FindLevelMVar_main(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__3___closed__2; LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownAnalyze___hyg_212_(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withNaryArg___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore_analyzeArg___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_Core_checkMaxHeartbeats(lean_object*, lean_object*, lean_object*, lean_object*); @@ -692,7 +699,6 @@ size_t lean_usize_add(size_t, size_t); lean_object* l_Lean_MetavarContext_findLevelDepth_x3f(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_containsBadMax(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_analyzeConst___closed__2; -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getPPAnalyzeTrustId___boxed(lean_object*); uint8_t l_Lean_Expr_hasLevelParam(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore_analyzeArg___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*); @@ -705,6 +711,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze static lean_object* l_Lean_PrettyPrinter_Delaborator_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownAnalyze___hyg_12870____closed__3; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_topDownAnalyze___lambda__3(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_analyzeProj___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_trace_profiler; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_analyzeConst(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_checkKnowsType(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*); @@ -795,10 +802,10 @@ static lean_object* l_Lean_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownA LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_analyzePi(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_mvarId_x21(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore_analyzeArg(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_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___closed__2; lean_object* l_Lean_Meta_setPostponed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_isFOLike(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___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*); +static lean_object* l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__4___closed__1; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_returnsPi___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofName(lean_object*); @@ -817,9 +824,9 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_replace LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_analyzeLam___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_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_withKnowing___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_PrettyPrinter_Delaborator_SubExpr_getPos___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore_applyFunBinderHeuristic___spec__2___boxed(lean_object*); +static lean_object* l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__3___closed__5; static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_isHBinOp___lambda__1___closed__12; static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_isHBinOp___lambda__1___closed__3; -LEAN_EXPORT lean_object* l_Lean_withSeconds___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_isProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_annotateBoolAt___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_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_analyzeAppStaged___lambda__1___closed__2; @@ -837,7 +844,7 @@ uint8_t l_Lean_getPPProofsWithType(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_isHBinOp___lambda__1___closed__6; static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_isHBinOp___lambda__1___closed__11; static lean_object* l_Lean_PrettyPrinter_Delaborator_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownAnalyze___hyg_12870____closed__13; -lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_processPostponed___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static double l_Lean_withSeconds___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__3___closed__1; lean_object* l_Lean_Meta_processPostponed(uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Array_isEmpty___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_instMonadReaderOfSubExprAnalyzeM___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -3594,7 +3601,7 @@ LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_TopDownAnaly _start: { lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_7 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_1, x_2, x_3, x_4, x_5, x_6); +x_7 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_1, x_2, x_3, x_4, x_5, x_6); x_8 = lean_ctor_get(x_7, 0); lean_inc(x_8); x_9 = lean_ctor_get(x_7, 1); @@ -30414,7 +30421,7 @@ x_2 = lean_alloc_closure((void*)(l_Lean_Meta_withNewMCtxDepth___at_Lean_PrettyPr return x_2; } } -static double _init_l_Lean_withSeconds___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__4___closed__1() { +static double _init_l_Lean_withSeconds___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__3___closed__1() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; double x_4; @@ -30425,7 +30432,7 @@ x_4 = l_Float_ofScientific(x_1, x_2, x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_withSeconds___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___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_EXPORT lean_object* l_Lean_withSeconds___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___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; lean_object* x_8; lean_object* x_9; lean_object* x_10; @@ -30457,7 +30464,7 @@ x_17 = 0; x_18 = lean_unsigned_to_nat(0u); x_19 = l_Float_ofScientific(x_16, x_17, x_18); lean_dec(x_16); -x_20 = l_Lean_withSeconds___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__4___closed__1; +x_20 = l_Lean_withSeconds___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__3___closed__1; x_21 = lean_float_div(x_19, x_20); x_22 = lean_box_float(x_21); x_23 = lean_alloc_ctor(0, 2, 0); @@ -30481,7 +30488,7 @@ x_27 = 0; x_28 = lean_unsigned_to_nat(0u); x_29 = l_Float_ofScientific(x_26, x_27, x_28); lean_dec(x_26); -x_30 = l_Lean_withSeconds___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__4___closed__1; +x_30 = l_Lean_withSeconds___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__3___closed__1; x_31 = lean_float_div(x_29, x_30); x_32 = lean_box_float(x_31); x_33 = lean_alloc_ctor(0, 2, 0); @@ -30518,201 +30525,7 @@ return x_38; } } } -static lean_object* _init_l_Lean_withOptProfile___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__3___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("profiler", 8); -return x_1; -} -} -static lean_object* _init_l_Lean_withOptProfile___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___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_Lean_withOptProfile___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__3___closed__1; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Lean_withOptProfile___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___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; uint8_t x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_7 = l_Lean_withOptProfile___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__3___closed__2; -x_8 = 0; -x_9 = l_Lean_getBoolOption___at_Lean_Meta_processPostponed___spec__4(x_7, x_8, x_2, x_3, x_4, x_5, x_6); -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -x_11 = lean_unbox(x_10); -lean_dec(x_10); -if (x_11 == 0) -{ -lean_object* x_12; lean_object* x_13; -x_12 = lean_ctor_get(x_9, 1); -lean_inc(x_12); -lean_dec(x_9); -x_13 = lean_apply_5(x_1, x_2, x_3, x_4, x_5, x_12); -if (lean_obj_tag(x_13) == 0) -{ -uint8_t x_14; -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_13, 0); -x_16 = lean_box(0); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_15); -lean_ctor_set(x_17, 1, x_16); -lean_ctor_set(x_13, 0, x_17); -return x_13; -} -else -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_18 = lean_ctor_get(x_13, 0); -x_19 = lean_ctor_get(x_13, 1); -lean_inc(x_19); -lean_inc(x_18); -lean_dec(x_13); -x_20 = lean_box(0); -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(0, 2, 0); -lean_ctor_set(x_22, 0, x_21); -lean_ctor_set(x_22, 1, x_19); -return x_22; -} -} -else -{ -uint8_t x_23; -x_23 = !lean_is_exclusive(x_13); -if (x_23 == 0) -{ -return x_13; -} -else -{ -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_13, 0); -x_25 = lean_ctor_get(x_13, 1); -lean_inc(x_25); -lean_inc(x_24); -lean_dec(x_13); -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_27; lean_object* x_28; -x_27 = lean_ctor_get(x_9, 1); -lean_inc(x_27); -lean_dec(x_9); -x_28 = l_Lean_withSeconds___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__4(x_1, x_2, x_3, x_4, x_5, x_27); -if (lean_obj_tag(x_28) == 0) -{ -uint8_t x_29; -x_29 = !lean_is_exclusive(x_28); -if (x_29 == 0) -{ -lean_object* x_30; uint8_t x_31; -x_30 = lean_ctor_get(x_28, 0); -x_31 = !lean_is_exclusive(x_30); -if (x_31 == 0) -{ -lean_object* x_32; lean_object* x_33; -x_32 = lean_ctor_get(x_30, 1); -x_33 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_33, 0, x_32); -lean_ctor_set(x_30, 1, x_33); -return x_28; -} -else -{ -lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_34 = lean_ctor_get(x_30, 0); -x_35 = lean_ctor_get(x_30, 1); -lean_inc(x_35); -lean_inc(x_34); -lean_dec(x_30); -x_36 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_36, 0, x_35); -x_37 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_37, 0, x_34); -lean_ctor_set(x_37, 1, x_36); -lean_ctor_set(x_28, 0, x_37); -return x_28; -} -} -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; -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_ctor_get(x_38, 0); -lean_inc(x_40); -x_41 = lean_ctor_get(x_38, 1); -lean_inc(x_41); -if (lean_is_exclusive(x_38)) { - lean_ctor_release(x_38, 0); - lean_ctor_release(x_38, 1); - x_42 = x_38; -} else { - lean_dec_ref(x_38); - x_42 = lean_box(0); -} -x_43 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_43, 0, x_41); -if (lean_is_scalar(x_42)) { - x_44 = lean_alloc_ctor(0, 2, 0); -} else { - x_44 = x_42; -} -lean_ctor_set(x_44, 0, x_40); -lean_ctor_set(x_44, 1, 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_39); -return x_45; -} -} -else -{ -uint8_t x_46; -x_46 = !lean_is_exclusive(x_28); -if (x_46 == 0) -{ -return x_28; -} -else -{ -lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_47 = lean_ctor_get(x_28, 0); -x_48 = lean_ctor_get(x_28, 1); -lean_inc(x_48); -lean_inc(x_47); -lean_dec(x_28); -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_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___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_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___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) { _start: { if (lean_obj_tag(x_1) == 0) @@ -30803,237 +30616,565 @@ return x_21; } } } -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___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) { -_start: -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; -lean_inc(x_10); -x_13 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_processPostponed___spec__6(x_1, x_2, x_3, x_6, x_4, x_8, x_9, x_10, x_11, x_12); -x_14 = lean_ctor_get(x_13, 1); -lean_inc(x_14); -lean_dec(x_13); -x_15 = l_MonadExcept_ofExcept___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__5(x_5, x_8, x_9, x_10, x_11, x_14); -lean_dec(x_10); -return x_15; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___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) { +_start: +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; +lean_inc(x_10); +x_13 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_processPostponed___spec__4(x_1, x_2, x_3, x_6, x_4, x_8, x_9, x_10, x_11, x_12); +x_14 = lean_ctor_get(x_13, 1); +lean_inc(x_14); +lean_dec(x_13); +x_15 = l_MonadExcept_ofExcept___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__4(x_5, x_8, x_9, x_10, x_11, x_14); +lean_dec(x_10); +return x_15; +} +} +static lean_object* _init_l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_profiler; +return x_1; +} +} +static lean_object* _init_l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__3___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("[", 1); +return x_1; +} +} +static lean_object* _init_l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__3___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__3___closed__2; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__3___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("s] ", 3); +return x_1; +} +} +static lean_object* _init_l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__3___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__3___closed__4; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, lean_object* x_6, uint8_t x_7, double 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; lean_object* x_16; +lean_dec(x_9); +x_15 = lean_ctor_get(x_12, 5); +lean_inc(x_15); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_2); +x_16 = lean_apply_6(x_1, x_2, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_16) == 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); +x_19 = l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__3___closed__1; +x_20 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_6, x_19); +lean_dec(x_6); +if (x_20 == 0) +{ +if (x_7 == 0) +{ +lean_object* x_21; lean_object* x_22; +x_21 = lean_box(0); +x_22 = l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__2(x_3, x_4, x_15, x_5, x_2, x_17, x_21, x_10, x_11, x_12, x_13, x_18); +lean_dec(x_13); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_2); +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; +x_23 = lean_float_to_string(x_8); +x_24 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_24, 0, x_23); +x_25 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_25, 0, x_24); +x_26 = l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__3___closed__3; +x_27 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_25); +x_28 = l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__3___closed__5; +x_29 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +x_30 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_17); +x_31 = l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_annotateBoolAt___closed__4; +x_32 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +x_33 = lean_box(0); +x_34 = l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__2(x_3, x_4, x_15, x_5, x_2, x_32, x_33, x_10, x_11, x_12, x_13, x_18); +lean_dec(x_13); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_2); +return x_34; +} +} +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; +x_35 = lean_float_to_string(x_8); +x_36 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_36, 0, x_35); +x_37 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_37, 0, x_36); +x_38 = l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__3___closed__3; +x_39 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_37); +x_40 = l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__3___closed__5; +x_41 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_40); +x_42 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_17); +x_43 = l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_annotateBoolAt___closed__4; +x_44 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_44, 0, x_42); +lean_ctor_set(x_44, 1, x_43); +x_45 = lean_box(0); +x_46 = l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__2(x_3, x_4, x_15, x_5, x_2, x_44, x_45, x_10, x_11, x_12, x_13, x_18); +lean_dec(x_13); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_2); +return x_46; +} +} +else +{ +uint8_t x_47; +lean_dec(x_15); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_47 = !lean_is_exclusive(x_16); +if (x_47 == 0) +{ +return x_16; +} +else +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_16, 0); +x_49 = lean_ctor_get(x_16, 1); +lean_inc(x_49); +lean_inc(x_48); +lean_dec(x_16); +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; +} +} +} +} +static lean_object* _init_l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__4___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_trace_profiler; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t 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, 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_dec(x_7); +x_13 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_processPostponed___spec__2___rarg(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_alloc_closure((void*)(l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__1), 6, 1); +lean_closure_set(x_16, 0, x_1); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +x_17 = l_Lean_withSeconds___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__3(x_16, x_8, x_9, x_10, x_11, x_15); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; lean_object* x_67; uint8_t x_68; +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_ctor_get(x_18, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_18, 1); +lean_inc(x_21); +lean_dec(x_18); +x_67 = l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__4___closed__1; +x_68 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_5, x_67); +if (x_68 == 0) +{ +uint8_t x_69; +x_69 = 0; +x_22 = x_69; +goto block_66; +} +else +{ +double x_70; double x_71; uint8_t x_72; +x_70 = l_Lean_trace_profiler_threshold_getSecs(x_5); +x_71 = lean_unbox_float(x_21); +x_72 = lean_float_decLt(x_70, x_71); +x_22 = x_72; +goto block_66; +} +block_66: +{ +if (x_6 == 0) +{ +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; +lean_dec(x_21); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +x_23 = lean_st_ref_take(x_11, 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_is_exclusive(x_24); +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_24, 3); +x_28 = l_Lean_PersistentArray_append___rarg(x_14, x_27); +lean_ctor_set(x_24, 3, x_28); +x_29 = lean_st_ref_set(x_11, x_24, x_25); +x_30 = lean_ctor_get(x_29, 1); +lean_inc(x_30); +lean_dec(x_29); +x_31 = l_MonadExcept_ofExcept___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__4(x_20, x_8, x_9, x_10, x_11, x_30); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_20); +if (lean_obj_tag(x_31) == 0) +{ +uint8_t x_32; +x_32 = !lean_is_exclusive(x_31); +if (x_32 == 0) +{ +return x_31; +} +else +{ +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(0, 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_36; +x_36 = !lean_is_exclusive(x_31); +if (x_36 == 0) +{ +return x_31; +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_31, 0); +x_38 = lean_ctor_get(x_31, 1); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_31); +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_40; lean_object* x_41; 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_40 = lean_ctor_get(x_24, 0); +x_41 = lean_ctor_get(x_24, 1); +x_42 = lean_ctor_get(x_24, 2); +x_43 = lean_ctor_get(x_24, 3); +x_44 = lean_ctor_get(x_24, 4); +x_45 = lean_ctor_get(x_24, 5); +x_46 = lean_ctor_get(x_24, 6); +lean_inc(x_46); +lean_inc(x_45); +lean_inc(x_44); +lean_inc(x_43); +lean_inc(x_42); +lean_inc(x_41); +lean_inc(x_40); +lean_dec(x_24); +x_47 = l_Lean_PersistentArray_append___rarg(x_14, x_43); +x_48 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_48, 0, x_40); +lean_ctor_set(x_48, 1, x_41); +lean_ctor_set(x_48, 2, x_42); +lean_ctor_set(x_48, 3, x_47); +lean_ctor_set(x_48, 4, x_44); +lean_ctor_set(x_48, 5, x_45); +lean_ctor_set(x_48, 6, x_46); +x_49 = lean_st_ref_set(x_11, x_48, x_25); +x_50 = lean_ctor_get(x_49, 1); +lean_inc(x_50); +lean_dec(x_49); +x_51 = l_MonadExcept_ofExcept___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__4(x_20, x_8, x_9, x_10, x_11, x_50); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_20); +if (lean_obj_tag(x_51) == 0) +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_52 = lean_ctor_get(x_51, 0); +lean_inc(x_52); +x_53 = lean_ctor_get(x_51, 1); +lean_inc(x_53); +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_51); + x_54 = lean_box(0); +} +if (lean_is_scalar(x_54)) { + x_55 = lean_alloc_ctor(0, 2, 0); +} else { + x_55 = x_54; +} +lean_ctor_set(x_55, 0, x_52); +lean_ctor_set(x_55, 1, x_53); +return x_55; +} +else +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_56 = lean_ctor_get(x_51, 0); +lean_inc(x_56); +x_57 = lean_ctor_get(x_51, 1); +lean_inc(x_57); +if (lean_is_exclusive(x_51)) { + lean_ctor_release(x_51, 0); + lean_ctor_release(x_51, 1); + x_58 = x_51; +} else { + lean_dec_ref(x_51); + x_58 = lean_box(0); +} +if (lean_is_scalar(x_58)) { + x_59 = lean_alloc_ctor(1, 2, 0); +} else { + x_59 = x_58; +} +lean_ctor_set(x_59, 0, x_56); +lean_ctor_set(x_59, 1, x_57); +return x_59; +} } } -static lean_object* _init_l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("[", 1); -return x_1; +lean_object* x_60; double x_61; lean_object* x_62; +x_60 = lean_box(0); +x_61 = lean_unbox_float(x_21); +lean_dec(x_21); +x_62 = l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__3(x_2, x_20, x_14, x_3, x_4, x_5, x_22, x_61, x_60, x_8, x_9, x_10, x_11, x_19); +return x_62; } } -static lean_object* _init_l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___closed__2() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___closed__1; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +lean_object* x_63; double x_64; lean_object* x_65; +x_63 = lean_box(0); +x_64 = lean_unbox_float(x_21); +lean_dec(x_21); +x_65 = l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__3(x_2, x_20, x_14, x_3, x_4, x_5, x_22, x_64, x_63, x_8, x_9, x_10, x_11, x_19); +return x_65; } } -static lean_object* _init_l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("s] ", 3); -return x_1; } +else +{ +uint8_t x_73; +lean_dec(x_14); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +x_73 = !lean_is_exclusive(x_17); +if (x_73 == 0) +{ +return x_17; } -static lean_object* _init_l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___closed__4() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___closed__3; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_74 = lean_ctor_get(x_17, 0); +x_75 = lean_ctor_get(x_17, 1); +lean_inc(x_75); +lean_inc(x_74); +lean_dec(x_17); +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_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__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) { _start: { -lean_object* x_10; lean_object* x_11; uint8_t x_12; +lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_10 = lean_ctor_get(x_7, 2); +lean_inc(x_10); lean_inc(x_1); -x_10 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_1, 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_unbox(x_11); +x_11 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___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_unbox(x_12); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); lean_dec(x_11); -if (x_12 == 0) +x_15 = l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__4___closed__1; +x_16 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_10, x_15); +if (x_16 == 0) { -lean_object* x_13; lean_object* x_14; +lean_object* x_17; +lean_dec(x_12); +lean_dec(x_10); lean_dec(x_2); lean_dec(x_1); -x_13 = lean_ctor_get(x_10, 1); -lean_inc(x_13); -lean_dec(x_10); -x_14 = lean_apply_5(x_3, x_5, x_6, x_7, x_8, x_13); -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; -x_15 = lean_ctor_get(x_10, 1); -lean_inc(x_15); -lean_dec(x_10); -x_16 = lean_ctor_get(x_7, 5); -lean_inc(x_16); -x_17 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_processPostponed___spec__2___rarg(x_8, 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_alloc_closure((void*)(l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__1), 6, 1); -lean_closure_set(x_20, 0, x_3); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_21 = l_Lean_withOptProfile___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__3(x_20, x_5, x_6, x_7, x_8, 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; -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_22, 0); -lean_inc(x_24); -x_25 = lean_ctor_get(x_22, 1); -lean_inc(x_25); -lean_dec(x_22); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_24); -x_26 = lean_apply_6(x_2, x_24, x_5, x_6, x_7, x_8, x_23); -if (lean_obj_tag(x_26) == 0) +x_17 = lean_apply_5(x_3, x_5, x_6, x_7, x_8, x_14); +if (lean_obj_tag(x_17) == 0) { -if (lean_obj_tag(x_25) == 0) +uint8_t x_18; +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; -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_box(0); -x_30 = l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__2(x_18, x_1, x_16, x_4, x_24, x_27, x_29, x_5, x_6, x_7, x_8, x_28); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_24); -return x_30; +return x_17; } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; double x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; 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_31 = lean_ctor_get(x_26, 0); -lean_inc(x_31); -x_32 = lean_ctor_get(x_26, 1); -lean_inc(x_32); -lean_dec(x_26); -x_33 = lean_ctor_get(x_25, 0); -lean_inc(x_33); -lean_dec(x_25); -x_34 = lean_unbox_float(x_33); -lean_dec(x_33); -x_35 = lean_float_to_string(x_34); -x_36 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_36, 0, x_35); -x_37 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_37, 0, x_36); -x_38 = l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___closed__2; -x_39 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_39, 1, x_37); -x_40 = l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___closed__4; -x_41 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_41, 0, x_39); -lean_ctor_set(x_41, 1, x_40); -x_42 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_31); -x_43 = l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_annotateBoolAt___closed__4; -x_44 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_44, 0, x_42); -lean_ctor_set(x_44, 1, x_43); -x_45 = lean_box(0); -x_46 = l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__2(x_18, x_1, x_16, x_4, x_24, x_44, x_45, x_5, x_6, x_7, x_8, x_32); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_24); -return x_46; +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_17, 0); +x_20 = lean_ctor_get(x_17, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_17); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); +return x_21; } } else { -uint8_t x_47; -lean_dec(x_25); -lean_dec(x_24); -lean_dec(x_18); -lean_dec(x_16); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_1); -x_47 = !lean_is_exclusive(x_26); -if (x_47 == 0) +uint8_t x_22; +x_22 = !lean_is_exclusive(x_17); +if (x_22 == 0) { -return x_26; +return x_17; } else { -lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_48 = lean_ctor_get(x_26, 0); -x_49 = lean_ctor_get(x_26, 1); -lean_inc(x_49); -lean_inc(x_48); -lean_dec(x_26); -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; +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_17, 0); +x_24 = lean_ctor_get(x_17, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_17); +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 { -uint8_t x_51; -lean_dec(x_18); -lean_dec(x_16); -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_51 = !lean_is_exclusive(x_21); -if (x_51 == 0) -{ -return x_21; +lean_object* x_26; uint8_t x_27; lean_object* x_28; +x_26 = lean_box(0); +x_27 = lean_unbox(x_12); +lean_dec(x_12); +x_28 = l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__4(x_3, x_2, x_1, x_4, x_10, x_27, x_26, x_5, x_6, x_7, x_8, x_14); +return x_28; +} } else { -lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_52 = lean_ctor_get(x_21, 0); -x_53 = lean_ctor_get(x_21, 1); -lean_inc(x_53); -lean_inc(x_52); -lean_dec(x_21); -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_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; +x_29 = lean_ctor_get(x_11, 1); +lean_inc(x_29); +lean_dec(x_11); +x_30 = lean_box(0); +x_31 = lean_unbox(x_12); +lean_dec(x_12); +x_32 = l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__4(x_3, x_2, x_1, x_4, x_10, x_31, x_30, x_5, x_6, x_7, x_8, x_29); +return x_32; } } } @@ -31464,11 +31605,11 @@ x_11 = l_Lean_Meta_withNewMCtxDepth___at_Lean_PrettyPrinter_Delaborator_topDownA return x_11; } } -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___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_MonadExcept_ofExcept___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___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: { lean_object* x_7; -x_7 = l_MonadExcept_ofExcept___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__5(x_1, x_2, x_3, x_4, x_5, x_6); +x_7 = l_MonadExcept_ofExcept___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__4(x_1, x_2, x_3, x_4, x_5, x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); @@ -31492,6 +31633,32 @@ lean_dec(x_5); return x_14; } } +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___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_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +uint8_t x_15; uint8_t x_16; double x_17; lean_object* x_18; +x_15 = lean_unbox(x_5); +lean_dec(x_5); +x_16 = lean_unbox(x_7); +lean_dec(x_7); +x_17 = lean_unbox_float(x_8); +lean_dec(x_8); +x_18 = l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__3(x_1, x_2, x_3, x_4, x_15, x_6, x_16, x_17, x_9, x_10, x_11, x_12, x_13, x_14); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___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, lean_object* x_11, lean_object* x_12) { +_start: +{ +uint8_t x_13; uint8_t x_14; lean_object* x_15; +x_13 = lean_unbox(x_4); +lean_dec(x_4); +x_14 = lean_unbox(x_6); +lean_dec(x_6); +x_15 = l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__4(x_1, x_2, x_3, x_13, x_5, x_14, x_7, x_8, x_9, x_10, x_11, x_12); +return x_15; +} +} LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___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: { @@ -32383,19 +32550,19 @@ l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore_maybeSetExp lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore_maybeSetExplicit___closed__5); l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore_maybeSetExplicit___closed__6 = _init_l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore_maybeSetExplicit___closed__6(); lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore_maybeSetExplicit___closed__6); -l_Lean_withSeconds___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__4___closed__1 = _init_l_Lean_withSeconds___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__4___closed__1(); -l_Lean_withOptProfile___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__3___closed__1 = _init_l_Lean_withOptProfile___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__3___closed__1(); -lean_mark_persistent(l_Lean_withOptProfile___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__3___closed__1); -l_Lean_withOptProfile___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__3___closed__2 = _init_l_Lean_withOptProfile___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__3___closed__2(); -lean_mark_persistent(l_Lean_withOptProfile___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__3___closed__2); -l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___closed__1 = _init_l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___closed__1(); -lean_mark_persistent(l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___closed__1); -l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___closed__2 = _init_l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___closed__2(); -lean_mark_persistent(l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___closed__2); -l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___closed__3 = _init_l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___closed__3(); -lean_mark_persistent(l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___closed__3); -l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___closed__4 = _init_l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___closed__4(); -lean_mark_persistent(l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___closed__4); +l_Lean_withSeconds___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__3___closed__1 = _init_l_Lean_withSeconds___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__3___closed__1(); +l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__3___closed__1 = _init_l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__3___closed__1); +l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__3___closed__2 = _init_l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__3___closed__2); +l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__3___closed__3 = _init_l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__3___closed__3(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__3___closed__3); +l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__3___closed__4 = _init_l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__3___closed__4(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__3___closed__4); +l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__3___closed__5 = _init_l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__3___closed__5(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__3___closed__5); +l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__4___closed__1 = _init_l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__4___closed__1); l_Lean_PrettyPrinter_Delaborator_topDownAnalyze___lambda__3___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_topDownAnalyze___lambda__3___closed__1(); lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_topDownAnalyze___lambda__3___closed__1); l_Lean_PrettyPrinter_Delaborator_topDownAnalyze___lambda__3___closed__2 = _init_l_Lean_PrettyPrinter_Delaborator_topDownAnalyze___lambda__3___closed__2(); diff --git a/stage0/stdlib/Lean/PrettyPrinter/Formatter.c b/stage0/stdlib/Lean/PrettyPrinter/Formatter.c index 28be967e8f12..a16bcef03f88 100644 --- a/stage0/stdlib/Lean/PrettyPrinter/Formatter.c +++ b/stage0/stdlib/Lean/PrettyPrinter/Formatter.c @@ -14,6 +14,7 @@ extern "C" { #endif static lean_object* l_Lean_PrettyPrinter_Formatter_categoryFormatterCore___lambda__4___closed__1; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_format___lambda__4(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*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_instMonadTraverserFormatterM___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_pushToken(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -21,8 +22,9 @@ static lean_object* l_Lean_PrettyPrinter_Formatter_State_stack___default___close LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_setLhsPrec_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_PrettyPrinter_Formatter_unicodeSymbolNoAntiquot_formatter___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_string_utf8_extract(lean_object*, lean_object*, lean_object*); +lean_object* l___private_Init_Util_0__outOfBounds___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_withAntiquotSuffixSplice_formatter___boxed(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_format___lambda__1(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*); uint8_t l_String_endsWith(lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Formatter_checkKind___closed__7; @@ -41,7 +43,6 @@ LEAN_EXPORT lean_object* l_Nat_forM_loop___at_Lean_PrettyPrinter_Formatter_manyN lean_object* lean_mk_empty_array_with_capacity(lean_object*); static lean_object* l_Lean_PrettyPrinter_mkCombinatorFormatterAttribute___closed__3; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_setLhsPrec_formatter___rarg(lean_object*); -static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__5; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_State_leadWord___default; static lean_object* l_Lean_PrettyPrinter_formatCommand___closed__2; LEAN_EXPORT lean_object* l_List_forM___at_Lean_PrettyPrinter_Formatter_sepByNoAntiquot_formatter___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -58,6 +59,7 @@ static lean_object* l_panic___at_Lean_PrettyPrinter_Formatter_symbolNoAntiquot_f LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Formatter_interpolatedStr_formatter___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_Lean_PrettyPrinter_Formatter_fill(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_6401_(lean_object*); static lean_object* l_Lean_PrettyPrinter_mkFormatterAttribute___closed__10; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_getStackSize___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Formatter_checkKind___closed__3; @@ -67,9 +69,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_mkFormatterAttribute___lambda__1__ lean_object* l_Lean_Syntax_formatStxAux(lean_object*, uint8_t, lean_object*, lean_object*); lean_object* l_Lean_Syntax_Traverser_fromSyntax(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_mkFormatterAttribute___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_PrettyPrinter_format___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Syntax_MonadTraverser_goUp___at_Lean_PrettyPrinter_Formatter_visitArgs___spec__4___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__9; static lean_object* l_Lean_PrettyPrinter_Formatter_instMonadTraverserFormatterM___closed__1; static lean_object* l_Lean_PrettyPrinter_Formatter_many1Unbox_formatter___closed__2; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_checkKind(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -80,6 +80,7 @@ uint8_t l_Lean_Syntax_isAntiquotSuffixSplice(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_instOrElseFormatterM(lean_object*); static lean_object* l_Lean_PrettyPrinter_Formatter_fieldIdx_formatter___closed__2; lean_object* l_Lean_KeyedDeclsAttribute_init___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_PrettyPrinter_format___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_categoryFormatterCore___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_toString(lean_object*, uint8_t); extern lean_object* l_Lean_Parser_builtinTokenTable; @@ -94,6 +95,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_sepByNoAntiquot_formatte LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_PrettyPrinter_Formatter_interpolatedStr_formatter___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_visitArgs___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_checkColGe_formatter___rarg(lean_object*); +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_instMonadTraverserFormatterM___lambda__2(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_PrettyPrinter_Formatter_checkLineEq_formatter(lean_object*, lean_object*, lean_object*, lean_object*); @@ -101,7 +103,6 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_skip_formatter(lean_obje lean_object* l_Lean_Syntax_getArgs(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Formatter_0__Lean_PrettyPrinter_Formatter_getExprPos_x3f(lean_object*); LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_PrettyPrinter_Formatter_categoryFormatterCore___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__10; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_checkLineEq_formatter___rarg(lean_object*); lean_object* l_String_trim(lean_object*); lean_object* l_instInhabitedReaderT___rarg___boxed(lean_object*, lean_object*); @@ -119,7 +120,6 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_pushToken___lambda__5(le lean_object* lean_environment_find(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_formatterForKindUnsafe(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_throwBacktrack(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_PrettyPrinter_format___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_forM_loop___at_Lean_PrettyPrinter_Formatter_manyNoAntiquot_formatter___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_Formatter_getStackSize___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_checkNoWsBefore_formatter___boxed(lean_object*); @@ -127,12 +127,14 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_format(lean_object*, lean_object*, static lean_object* l_Lean_PrettyPrinter_Formatter_symbolNoAntiquot_formatter___closed__4; lean_object* l_Lean_Elab_addConstInfo___at_Lean_registerInitAttrUnsafe___spec__13(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_error_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_6740____closed__3; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_FormatterM_orElse(lean_object*); static lean_object* l_Lean_PrettyPrinter_Formatter_parseToken___closed__2; 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_Lean_PrettyPrinter_Formatter_checkStackTop_formatter___rarg(lean_object*); LEAN_EXPORT uint8_t l_Lean_PrettyPrinter_Formatter_State_isUngrouped___default; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_State_stack___default; +static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__5; LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_PrettyPrinter_Formatter_categoryFormatterCore___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_checkLhsPrec_formatter(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); @@ -152,6 +154,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_pushToken___lambda__2(le LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_getStack(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_rawCh_formatter(uint32_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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_format___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_checkNoWsBefore_formatter___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_charLitNoAntiquot_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_string_utf8_byte_size(lean_object*); @@ -161,21 +164,24 @@ lean_object* lean_string_push(lean_object*, uint32_t); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_checkTailWs_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Std_Format_isNil(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_withMaybeTag(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_PrettyPrinter_format___lambda__4___closed__4; extern lean_object* l_instInhabitedNat; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_skip_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_instInhabitedPUnit; +LEAN_EXPORT uint8_t l_Lean_PrettyPrinter_format___lambda__2(uint32_t); static lean_object* l_Lean_PrettyPrinter_Formatter_parseToken___closed__3; static lean_object* l_Lean_PrettyPrinter_Formatter_formatterForKindUnsafe___closed__2; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_getStack___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_trailingNode_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_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__6; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_pushToken___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_PrettyPrinter_format___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_trailingNode_formatter(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_Formatter_checkLinebreakBefore_formatter___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_withoutInfo_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_FormatterM_orElse___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_withFn_formatter(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_format___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_checkTailWs_formatter(lean_object*, lean_object*, lean_object*, lean_object*); 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*); @@ -183,12 +189,12 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_categoryFormatterCore(le lean_object* lean_pretty_printer_formatter_interpret_parser_descr(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Formatter_0__Lean_PrettyPrinter_Formatter_getExprPos_x3f___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_formatTerm(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_6740_(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_errorAtSavedPos_formatter___rarg(lean_object*); static lean_object* l_Lean_PrettyPrinter_Formatter_categoryFormatterCore___lambda__4___closed__4; static lean_object* l_Lean_PrettyPrinter_Formatter_trailingNode_formatter___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_pushNone_formatter___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Formatter_categoryFormatterCore___closed__4; -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_PrettyPrinter_format___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_notFollowedBy_formatter___rarg(lean_object*); static lean_object* l_Lean_PrettyPrinter_instOrElseFormatterM___closed__1; LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_PrettyPrinter_Formatter_categoryFormatterCore___spec__2(lean_object*, lean_object*); @@ -197,6 +203,7 @@ LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Formatter_0__Lean_Pretty lean_object* l_Array_reverse___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_checkLhsPrec_formatter___rarg(lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_hygieneInfoNoAntiquot_formatter___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Formatter_strLitNoAntiquot_formatter___closed__1; LEAN_EXPORT lean_object* l_Lean_Syntax_MonadTraverser_getCur___at_Lean_PrettyPrinter_Formatter_visitArgs___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*); @@ -210,9 +217,11 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_getStack___boxed(lean_ob LEAN_EXPORT uint8_t l_Lean_PrettyPrinter_Formatter_State_mustBeGrouped___default; static lean_object* l_Lean_PrettyPrinter_mkFormatterAttribute___closed__2; lean_object* lean_nat_to_int(lean_object*); +static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_6740____closed__1; lean_object* l_Lean_MessageData_ofSyntax(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_pushWhitespace___boxed(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___boxed(lean_object*); +static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__7; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_eoi_formatter___rarg(lean_object*); 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*); @@ -223,6 +232,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_concat___lambda__1(lean_ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_checkPrec_formatter(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_withAntiquotSuffixSplice_formatter(lean_object*); lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean_object*); +static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__4; lean_object* l_List_range(lean_object*); static lean_object* l_Lean_PrettyPrinter_Formatter_numLitNoAntiquot_formatter___closed__2; static lean_object* l_Lean_PrettyPrinter_Formatter_pushToken___lambda__4___closed__1; @@ -239,12 +249,15 @@ static lean_object* l_Lean_PrettyPrinter_mkFormatterAttribute___closed__5; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_visitAtom(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_formatTactic___closed__2; 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*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_hygieneInfoNoAntiquot_formatter(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_mkFormatterAttribute___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_many1Unbox_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_PrettyPrinter_format___lambda__4___closed__1; lean_object* l_Lean_Parser_tokenFn(lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); lean_object* l_Lean_throwError___at_Lean_registerTagAttribute___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_ite___rarg(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__11; static lean_object* l_Lean_PrettyPrinter_Formatter_nameLitNoAntiquot_formatter___closed__2; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_concat___lambda__1___boxed(lean_object*); lean_object* l_Lean_Attribute_Builtin_getIdent(lean_object*, lean_object*, lean_object*, lean_object*); @@ -253,6 +266,7 @@ static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Formatter_ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_instCoeFormatterFormatterAliasValue(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_registerAlias(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_mkFormatterAttribute___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_hygieneInfoNoAntiquot_formatter___boxed(lean_object*); lean_object* l_Lean_FileMap_ofString(lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_mkFormatterAttribute___lambda__3___closed__4; @@ -265,21 +279,20 @@ lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); lean_object* l_instInhabited___rarg(lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Formatter_checkKind___closed__1; static lean_object* l_Lean_PrettyPrinter_mkFormatterAttribute___closed__3; -static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__2; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_throwBacktrack___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_checkPrec_formatter___rarg(lean_object*); static lean_object* l_Lean_PrettyPrinter_Formatter_orelse_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*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_pushLine___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_mkFormatterAttribute___closed__8; -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_PrettyPrinter_format___spec__2___boxed(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_Formatter_unicodeSymbolNoAntiquot_formatter___closed__3; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_instMonadTraverserFormatterM___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__12; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_visitArgs___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_lookahead_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__2; static lean_object* l_Lean_PrettyPrinter_formatTactic___closed__1; +lean_object* l_Lean_isTracingEnabledFor___at_Lean_addDecl___spec__8(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_nonReservedSymbolNoAntiquot_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_checkColGe_formatter(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); @@ -291,10 +304,8 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_checkColGt_formatter___b lean_object* lean_mk_antiquot_formatter(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_formatTerm___closed__1; static lean_object* l_Lean_PrettyPrinter_mkFormatterAttribute___closed__9; -static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__7; -static lean_object* l_Lean_PrettyPrinter_format___lambda__1___closed__3; +uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, 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_EXPORT lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025_(lean_object*); lean_object* l_Lean_Syntax_getArg(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_PrettyPrinter_Formatter_symbolNoAntiquot_formatter___closed__7; @@ -307,20 +318,20 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Forma static lean_object* l_Lean_PrettyPrinter_Formatter_concat___closed__1; 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*); static lean_object* l_panic___at_Lean_PrettyPrinter_Formatter_symbolNoAntiquot_formatter___spec__1___closed__3; +static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_6740____closed__6; static lean_object* l_Lean_PrettyPrinter_Formatter_symbolNoAntiquot_formatter___closed__2; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_identEq_formatter___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_MonadTraverser_getCur___at_Lean_PrettyPrinter_Formatter_visitArgs___spec__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_pushToken___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_pushNone_formatter___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Formatter_pushToken___lambda__2___closed__3; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_pp_oneline; static lean_object* l_Lean_PrettyPrinter_Formatter_trailingNode_formatter___lambda__1___closed__1; static lean_object* l_Lean_PrettyPrinter_mkCombinatorFormatterAttribute___closed__5; uint8_t l_Lean_Environment_contains(lean_object*, lean_object*); lean_object* lean_int_sub(lean_object*, lean_object*); extern lean_object* l_Lean_inheritedTraceOptions; -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_6389_(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_pushToken___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__8; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_notFollowedBy_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_mkCombinatorFormatterAttribute___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_formatterAliasesRef; @@ -328,11 +339,12 @@ LEAN_EXPORT lean_object* l_Lean_Syntax_MonadTraverser_goUp___at_Lean_PrettyPrint 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_setStack___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__4; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_checkColEq_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272_(lean_object*); static lean_object* l_Lean_PrettyPrinter_Formatter_symbolNoAntiquot_formatter___closed__6; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_instMonadTraverserFormatterM___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_extract___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__12; 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_ite(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_pushNone_formatter___boxed(lean_object*); @@ -343,6 +355,7 @@ uint8_t l_Lean_Syntax_isToken(lean_object*, lean_object*); lean_object* l_Lean_ParserCompiler_registerCombinatorAttribute(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Formatter_identNoAntiquot_formatter___closed__4; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_checkTailWs_formatter___rarg(lean_object*); +lean_object* l_String_findAux(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Formatter_scientificLitNoAntiquot_formatter___closed__2; lean_object* l_Lean_Parser_registerAliasCore___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_categoryFormatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -352,19 +365,20 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_lookahead_formatter(lean 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_scientificLitNoAntiquot_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_group(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Std_Format_pretty_x27(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_categoryFormatterCore___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_checkLineEq_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__11; +static lean_object* l_Lean_PrettyPrinter_format___lambda__4___closed__6; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_checkStackTop_formatter(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_categoryFormatterCore___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Formatter_instMonadTraverserFormatterM___closed__3; static lean_object* l_Lean_PrettyPrinter_Formatter_orelse_formatter___closed__1; +static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__8; LEAN_EXPORT lean_object* l_Lean_Syntax_MonadTraverser_getCur___at_Lean_PrettyPrinter_Formatter_visitArgs___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_visitArgs(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_fold(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_errorAtSavedPos_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_PrettyPrinter_format___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_setStack(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Formatter_symbolNoAntiquot_formatter___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); @@ -378,45 +392,51 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_checkNoWsBefore_formatte LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_checkNoImmediateColon_formatter(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_sepBy1NoAntiquot_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_checkNoImmediateColon_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t lean_uint32_dec_eq(uint32_t, uint32_t); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_categoryFormatterCore___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__9; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_eoi_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Formatter_pushToken___lambda__2___closed__1; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Formatter_charLitNoAntiquot_formatter___closed__2; lean_object* l_Array_shrink___rarg(lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Formatter_categoryFormatterCore___closed__2; +static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_6740____closed__4; LEAN_EXPORT lean_object* l_Lean_Syntax_MonadTraverser_goDown___at_Lean_PrettyPrinter_Formatter_visitArgs___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_6740____closed__5; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_categoryFormatterCore___lambda__4(lean_object*, lean_object*, 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_Lean_throwError___at_Lean_PrettyPrinter_format___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_withAntiquotSuffixSplice_formatter___rarg(lean_object*, lean_object*, lean_object*, 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_categoryFormatterCore___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_PrettyPrinter_format___lambda__4___closed__2; static lean_object* l_Lean_PrettyPrinter_Formatter_identNoAntiquot_formatter___closed__2; lean_object* lean_nat_sub(lean_object*, lean_object*); +lean_object* l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_interpretParserDescr_x27___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Std_Format_getIndent(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_MonadTraverser_goLeft___at_Lean_PrettyPrinter_Formatter_visitArgs___spec__2(lean_object*); static lean_object* l_Lean_PrettyPrinter_Formatter_checkKind___closed__4; extern lean_object* l_Lean_PrettyPrinter_backtrackExceptionId; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_skip_formatter___rarg(lean_object*); -static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_andthen_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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_PrettyPrinter_Formatter_pushToken___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_checkLhsPrec_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_PrettyPrinter_format___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_instMonadTraverserFormatterM___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_indent(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_mkFormatterAttribute___closed__12; +LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_PrettyPrinter_format___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_checkColGe_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_6740____closed__2; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_mkAntiquot_formatter_x27___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_Formatter_pushNone_formatter(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_mkFormatterAttribute___lambda__3(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_mkFormatterAttribute___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_strLitNoAntiquot_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Core_instMonadCoreM; -LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_PrettyPrinter_format___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_mkFormatterAttribute___lambda__4(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_visitAtom___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_getStackSize___boxed(lean_object*); @@ -430,6 +450,8 @@ static lean_object* l_Lean_PrettyPrinter_mkFormatterAttribute___lambda__3___clos LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_error_formatter___rarg(lean_object*); 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(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_PrettyPrinter_format___lambda__4___closed__5; +static lean_object* l_Lean_PrettyPrinter_format___lambda__4___closed__7; static lean_object* l_Lean_PrettyPrinter_Formatter_visitArgs___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_getStackSize(lean_object*); size_t lean_usize_add(size_t, size_t); @@ -443,13 +465,17 @@ static lean_object* l_Lean_PrettyPrinter_Formatter_identNoAntiquot_formatter___c LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_checkKind___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_instInhabitedSyntax; lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__3; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_formatterAttribute; +lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__6; lean_object* l_String_trimRight(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_checkNoImmediateColon_formatter___rarg(lean_object*); +static lean_object* l_Lean_PrettyPrinter_format___lambda__4___closed__3; +static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__10; static lean_object* l_panic___at_Lean_PrettyPrinter_Formatter_symbolNoAntiquot_formatter___spec__1___closed__1; static lean_object* l_Lean_PrettyPrinter_Formatter_many1Unbox_formatter___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_pushAlign___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_hygieneInfoNoAntiquot_formatter___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Formatter_registerAlias___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_instMonadTraverserFormatterM___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_instCoeForAllFormatterFormatterAliasValue(lean_object*); @@ -470,17 +496,19 @@ static lean_object* l_Lean_PrettyPrinter_Formatter_unicodeSymbolNoAntiquot_forma LEAN_EXPORT lean_object* l_Lean_Syntax_MonadTraverser_getCur___at_Lean_PrettyPrinter_Formatter_visitArgs___spec__1(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_pushToken___lambda__2___closed__2; +static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_withFn_formatter___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_unicodeSymbolNoAntiquot_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_error_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_PrettyPrinter_format___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Formatter_formatterForKindUnsafe___closed__5; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_identEq_formatter(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_withFn_formatter___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_PrettyPrinter_Formatter_categoryFormatterCore___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_format___lambda__2___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_ite___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_formatTactic(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_MonadTraverser_goDown___at_Lean_PrettyPrinter_Formatter_visitArgs___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__3; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_setLhsPrec_formatter(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofName(lean_object*); static lean_object* l_Lean_PrettyPrinter_Formatter_strLitNoAntiquot_formatter___closed__2; @@ -9627,6 +9655,42 @@ lean_dec(x_1); return x_2; } } +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_hygieneInfoNoAntiquot_formatter___rarg(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_Syntax_MonadTraverser_goLeft___at_Lean_PrettyPrinter_Formatter_visitArgs___spec__2___rarg(x_1, x_2, x_3, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_hygieneInfoNoAntiquot_formatter(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Formatter_hygieneInfoNoAntiquot_formatter___rarg___boxed), 4, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_hygieneInfoNoAntiquot_formatter___rarg___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_PrettyPrinter_Formatter_hygieneInfoNoAntiquot_formatter___rarg(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_PrettyPrinter_Formatter_hygieneInfoNoAntiquot_formatter___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_PrettyPrinter_Formatter_hygieneInfoNoAntiquot_formatter(x_1); +lean_dec(x_1); +return x_2; +} +} LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_PrettyPrinter_Formatter_interpolatedStr_formatter___spec__1___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: { @@ -9904,7 +9968,7 @@ x_10 = l_Lean_PrettyPrinter_Formatter_ite___rarg(x_9, x_2, x_3, x_4, x_5, x_6, x return x_10; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_6389_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_6401_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; uint8_t x_4; @@ -9974,44 +10038,79 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_PrettyPrinter_format___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_6740____closed__1() { _start: { -lean_object* x_5; lean_object* x_6; uint8_t x_7; -x_5 = l_Lean_isTracingEnabledFor___at_Lean_PrettyPrinter_Formatter_categoryFormatterCore___spec__1___closed__1; -x_6 = lean_st_ref_get(x_5, x_4); -x_7 = !lean_is_exclusive(x_6); -if (x_7 == 0) +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("pp", 2); +return x_1; +} +} +static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_6740____closed__2() { +_start: { -lean_object* x_8; lean_object* x_9; uint8_t x_10; lean_object* x_11; -x_8 = lean_ctor_get(x_6, 0); -x_9 = lean_ctor_get(x_2, 2); -x_10 = l___private_Lean_Util_Trace_0__Lean_checkTraceOption(x_8, x_9, x_1); -lean_dec(x_8); -x_11 = lean_box(x_10); -lean_ctor_set(x_6, 0, x_11); -return x_6; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("oneline", 7); +return x_1; } -else +} +static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_6740____closed__3() { +_start: { -lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; -x_12 = lean_ctor_get(x_6, 0); -x_13 = lean_ctor_get(x_6, 1); -lean_inc(x_13); -lean_inc(x_12); -lean_dec(x_6); -x_14 = lean_ctor_get(x_2, 2); -x_15 = l___private_Lean_Util_Trace_0__Lean_checkTraceOption(x_12, x_14, x_1); -lean_dec(x_12); -x_16 = lean_box(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_13); -return x_17; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_6740____closed__1; +x_2 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_6740____closed__2; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_6740____closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("(pretty printer) elide all but first line of pretty printer output", 66); +return x_1; +} +} +static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_6740____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_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_6740____closed__1; +x_3 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_6740____closed__4; +x_4 = lean_box(x_1); +x_5 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_5, 0, x_4); +lean_ctor_set(x_5, 1, x_2); +lean_ctor_set(x_5, 2, x_3); +return x_5; +} +} +static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_6740____closed__6() { +_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_PrettyPrinter_mkFormatterAttribute___closed__5; +x_2 = l_Lean_PrettyPrinter_mkFormatterAttribute___closed__6; +x_3 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_6740____closed__1; +x_4 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_6740____closed__2; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} } +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_6740_(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_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_6740____closed__3; +x_3 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_6740____closed__5; +x_4 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_6740____closed__6; +x_5 = l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__1(x_2, x_3, x_4, x_1); +return x_5; } } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_PrettyPrinter_format___spec__2(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_PrettyPrinter_format___spec__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; uint8_t x_7; @@ -10049,7 +10148,7 @@ return x_13; } } } -LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_PrettyPrinter_format___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_addTrace___at_Lean_PrettyPrinter_format___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; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; @@ -10169,7 +10268,41 @@ return x_44; } } } -static lean_object* _init_l_Lean_PrettyPrinter_format___lambda__1___closed__1() { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_format___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_6; lean_object* x_7; lean_object* x_8; +x_6 = 1; +x_7 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_7, 0, x_1); +lean_ctor_set_uint8(x_7, sizeof(void*)*1, 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_5); +return x_8; +} +} +LEAN_EXPORT uint8_t l_Lean_PrettyPrinter_format___lambda__2(uint32_t x_1) { +_start: +{ +uint32_t x_2; uint8_t x_3; +x_2 = 10; +x_3 = lean_uint32_dec_eq(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_format___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) { +_start: +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_8 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_8, 0, x_3); +x_9 = lean_box(0); +x_10 = lean_apply_5(x_1, x_8, x_9, x_5, x_6, x_7); +return x_10; +} +} +static lean_object* _init_l_Lean_PrettyPrinter_format___lambda__4___closed__1() { _start: { lean_object* x_1; @@ -10177,7 +10310,23 @@ x_1 = l_Lean_Parser_builtinTokenTable; return x_1; } } -static lean_object* _init_l_Lean_PrettyPrinter_format___lambda__1___closed__2() { +static lean_object* _init_l_Lean_PrettyPrinter_format___lambda__4___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_format___lambda__1___boxed), 5, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_PrettyPrinter_format___lambda__4___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_PrettyPrinter_pp_oneline; +return x_1; +} +} +static lean_object* _init_l_Lean_PrettyPrinter_format___lambda__4___closed__4() { _start: { lean_object* x_1; @@ -10185,29 +10334,46 @@ x_1 = lean_mk_string_from_bytes("format: uncaught backtrack exception", 36); return x_1; } } -static lean_object* _init_l_Lean_PrettyPrinter_format___lambda__1___closed__3() { +static lean_object* _init_l_Lean_PrettyPrinter_format___lambda__4___closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_PrettyPrinter_format___lambda__1___closed__2; +x_1 = l_Lean_PrettyPrinter_format___lambda__4___closed__4; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_format___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) { +static lean_object* _init_l_Lean_PrettyPrinter_format___lambda__4___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_format___lambda__2___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_PrettyPrinter_format___lambda__4___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(" [...]", 6); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_format___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) { _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; uint8_t x_15; uint8_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_dec(x_3); x_7 = lean_ctor_get(x_4, 2); lean_inc(x_7); -x_8 = l_Lean_PrettyPrinter_format___lambda__1___closed__1; +x_8 = l_Lean_PrettyPrinter_format___lambda__4___closed__1; x_9 = lean_st_ref_get(x_8, 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); +lean_inc(x_7); x_12 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_12, 0, x_7); lean_ctor_set(x_12, 1, x_10); @@ -10235,105 +10401,452 @@ lean_inc(x_20); x_23 = l_Lean_PrettyPrinter_Formatter_fold(x_22, x_2, x_12, x_20, x_4, x_5, x_21); if (lean_obj_tag(x_23) == 0) { -lean_object* x_24; lean_object* x_25; uint8_t x_26; -lean_dec(x_5); -lean_dec(x_4); +lean_object* x_24; 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; uint8_t x_34; lean_object* x_35; x_24 = lean_ctor_get(x_23, 1); lean_inc(x_24); lean_dec(x_23); x_25 = lean_st_ref_get(x_20, x_24); lean_dec(x_20); -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; uint8_t x_32; lean_object* x_33; -x_27 = lean_ctor_get(x_25, 0); -x_28 = lean_ctor_get(x_27, 2); +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, 2); lean_inc(x_28); -lean_dec(x_27); -x_29 = l_Std_instInhabitedFormat; +lean_dec(x_26); +x_29 = lean_array_get_size(x_28); x_30 = lean_unsigned_to_nat(0u); -x_31 = lean_array_get(x_29, x_28, x_30); +x_31 = lean_nat_dec_lt(x_30, x_29); +lean_dec(x_29); +x_32 = l_Lean_PrettyPrinter_format___lambda__4___closed__2; +x_33 = l_Lean_PrettyPrinter_format___lambda__4___closed__3; +x_34 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_7, x_33); +if (x_31 == 0) +{ +lean_object* x_112; lean_object* x_113; lean_dec(x_28); -x_32 = 1; -x_33 = lean_alloc_ctor(6, 1, 1); -lean_ctor_set(x_33, 0, x_31); -lean_ctor_set_uint8(x_33, sizeof(void*)*1, x_32); -lean_ctor_set(x_25, 0, x_33); -return x_25; +x_112 = l_Std_instInhabitedFormat; +x_113 = l___private_Init_Util_0__outOfBounds___rarg(x_112); +x_35 = x_113; +goto block_111; } 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; uint8_t 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 = lean_ctor_get(x_34, 2); -lean_inc(x_36); -lean_dec(x_34); -x_37 = l_Std_instInhabitedFormat; -x_38 = lean_unsigned_to_nat(0u); -x_39 = lean_array_get(x_37, x_36, x_38); -lean_dec(x_36); -x_40 = 1; -x_41 = lean_alloc_ctor(6, 1, 1); -lean_ctor_set(x_41, 0, x_39); -lean_ctor_set_uint8(x_41, sizeof(void*)*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_35); +lean_object* x_114; +x_114 = lean_array_fget(x_28, x_30); +lean_dec(x_28); +x_35 = x_114; +goto block_111; +} +block_111: +{ +if (x_34 == 0) +{ +lean_object* x_36; lean_object* x_37; +lean_dec(x_7); +x_36 = lean_box(0); +lean_inc(x_5); +lean_inc(x_4); +x_37 = lean_apply_5(x_32, x_35, x_36, x_4, x_5, x_27); +if (lean_obj_tag(x_37) == 0) +{ +lean_dec(x_5); +lean_dec(x_4); +return x_37; +} +else +{ +lean_object* x_38; +x_38 = lean_ctor_get(x_37, 0); +lean_inc(x_38); +if (lean_obj_tag(x_38) == 0) +{ +uint8_t x_39; +lean_dec(x_5); +lean_dec(x_4); +x_39 = !lean_is_exclusive(x_37); +if (x_39 == 0) +{ +lean_object* x_40; +x_40 = lean_ctor_get(x_37, 0); +lean_dec(x_40); +return x_37; +} +else +{ +lean_object* x_41; lean_object* x_42; +x_41 = lean_ctor_get(x_37, 1); +lean_inc(x_41); +lean_dec(x_37); +x_42 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_42, 0, x_38); +lean_ctor_set(x_42, 1, x_41); return x_42; } } else { -lean_object* x_43; +uint8_t x_43; +x_43 = !lean_is_exclusive(x_37); +if (x_43 == 0) +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; +x_44 = lean_ctor_get(x_37, 1); +x_45 = lean_ctor_get(x_37, 0); +lean_dec(x_45); +x_46 = lean_ctor_get(x_38, 0); +lean_inc(x_46); +x_47 = l_Lean_PrettyPrinter_FormatterM_orElse___rarg___closed__1; +x_48 = lean_nat_dec_eq(x_47, x_46); +lean_dec(x_46); +if (x_48 == 0) +{ +lean_dec(x_5); +lean_dec(x_4); +return x_37; +} +else +{ +lean_object* x_49; lean_object* x_50; +lean_free_object(x_37); +lean_dec(x_38); +x_49 = l_Lean_PrettyPrinter_format___lambda__4___closed__5; +x_50 = l_Lean_throwError___at_Lean_PrettyPrinter_format___spec__1(x_49, x_4, x_5, x_44); +lean_dec(x_5); +lean_dec(x_4); +return x_50; +} +} +else +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; uint8_t x_54; +x_51 = lean_ctor_get(x_37, 1); +lean_inc(x_51); +lean_dec(x_37); +x_52 = lean_ctor_get(x_38, 0); +lean_inc(x_52); +x_53 = l_Lean_PrettyPrinter_FormatterM_orElse___rarg___closed__1; +x_54 = lean_nat_dec_eq(x_53, x_52); +lean_dec(x_52); +if (x_54 == 0) +{ +lean_object* x_55; +lean_dec(x_5); +lean_dec(x_4); +x_55 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_55, 0, x_38); +lean_ctor_set(x_55, 1, x_51); +return x_55; +} +else +{ +lean_object* x_56; lean_object* x_57; +lean_dec(x_38); +x_56 = l_Lean_PrettyPrinter_format___lambda__4___closed__5; +x_57 = l_Lean_throwError___at_Lean_PrettyPrinter_format___spec__1(x_56, x_4, x_5, x_51); +lean_dec(x_5); +lean_dec(x_4); +return x_57; +} +} +} +} +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; +lean_inc(x_35); +x_58 = l_Std_Format_pretty_x27(x_35, x_7); +lean_dec(x_7); +x_59 = l_String_trim(x_58); +lean_dec(x_58); +x_60 = lean_string_utf8_byte_size(x_59); +x_61 = l_Lean_PrettyPrinter_format___lambda__4___closed__6; +x_62 = l_String_findAux(x_59, x_61, x_60, x_30); +x_63 = lean_nat_dec_lt(x_62, x_60); +lean_dec(x_60); +if (x_63 == 0) +{ +lean_object* x_64; lean_object* x_65; +lean_dec(x_62); +x_64 = lean_box(0); +lean_inc(x_5); +lean_inc(x_4); +x_65 = l_Lean_PrettyPrinter_format___lambda__3(x_32, x_35, x_59, x_64, x_4, x_5, x_27); +lean_dec(x_35); +if (lean_obj_tag(x_65) == 0) +{ +lean_dec(x_5); +lean_dec(x_4); +return x_65; +} +else +{ +lean_object* x_66; +x_66 = lean_ctor_get(x_65, 0); +lean_inc(x_66); +if (lean_obj_tag(x_66) == 0) +{ +uint8_t x_67; +lean_dec(x_5); +lean_dec(x_4); +x_67 = !lean_is_exclusive(x_65); +if (x_67 == 0) +{ +lean_object* x_68; +x_68 = lean_ctor_get(x_65, 0); +lean_dec(x_68); +return x_65; +} +else +{ +lean_object* x_69; lean_object* x_70; +x_69 = lean_ctor_get(x_65, 1); +lean_inc(x_69); +lean_dec(x_65); +x_70 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_70, 0, x_66); +lean_ctor_set(x_70, 1, x_69); +return x_70; +} +} +else +{ +uint8_t x_71; +x_71 = !lean_is_exclusive(x_65); +if (x_71 == 0) +{ +lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; uint8_t x_76; +x_72 = lean_ctor_get(x_65, 1); +x_73 = lean_ctor_get(x_65, 0); +lean_dec(x_73); +x_74 = lean_ctor_get(x_66, 0); +lean_inc(x_74); +x_75 = l_Lean_PrettyPrinter_FormatterM_orElse___rarg___closed__1; +x_76 = lean_nat_dec_eq(x_75, x_74); +lean_dec(x_74); +if (x_76 == 0) +{ +lean_dec(x_5); +lean_dec(x_4); +return x_65; +} +else +{ +lean_object* x_77; lean_object* x_78; +lean_free_object(x_65); +lean_dec(x_66); +x_77 = l_Lean_PrettyPrinter_format___lambda__4___closed__5; +x_78 = l_Lean_throwError___at_Lean_PrettyPrinter_format___spec__1(x_77, x_4, x_5, x_72); +lean_dec(x_5); +lean_dec(x_4); +return x_78; +} +} +else +{ +lean_object* x_79; lean_object* x_80; lean_object* x_81; uint8_t x_82; +x_79 = lean_ctor_get(x_65, 1); +lean_inc(x_79); +lean_dec(x_65); +x_80 = lean_ctor_get(x_66, 0); +lean_inc(x_80); +x_81 = l_Lean_PrettyPrinter_FormatterM_orElse___rarg___closed__1; +x_82 = lean_nat_dec_eq(x_81, x_80); +lean_dec(x_80); +if (x_82 == 0) +{ +lean_object* x_83; +lean_dec(x_5); +lean_dec(x_4); +x_83 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_83, 0, x_66); +lean_ctor_set(x_83, 1, x_79); +return x_83; +} +else +{ +lean_object* x_84; lean_object* x_85; +lean_dec(x_66); +x_84 = l_Lean_PrettyPrinter_format___lambda__4___closed__5; +x_85 = l_Lean_throwError___at_Lean_PrettyPrinter_format___spec__1(x_84, x_4, x_5, x_79); +lean_dec(x_5); +lean_dec(x_4); +return x_85; +} +} +} +} +} +else +{ +lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; +x_86 = lean_string_utf8_extract(x_59, x_30, x_62); +lean_dec(x_62); +lean_dec(x_59); +x_87 = l_Lean_PrettyPrinter_format___lambda__4___closed__7; +x_88 = lean_string_append(x_86, x_87); +x_89 = lean_box(0); +lean_inc(x_5); +lean_inc(x_4); +x_90 = l_Lean_PrettyPrinter_format___lambda__3(x_32, x_35, x_88, x_89, x_4, x_5, x_27); +lean_dec(x_35); +if (lean_obj_tag(x_90) == 0) +{ +lean_dec(x_5); +lean_dec(x_4); +return x_90; +} +else +{ +lean_object* x_91; +x_91 = lean_ctor_get(x_90, 0); +lean_inc(x_91); +if (lean_obj_tag(x_91) == 0) +{ +uint8_t x_92; +lean_dec(x_5); +lean_dec(x_4); +x_92 = !lean_is_exclusive(x_90); +if (x_92 == 0) +{ +lean_object* x_93; +x_93 = lean_ctor_get(x_90, 0); +lean_dec(x_93); +return x_90; +} +else +{ +lean_object* x_94; lean_object* x_95; +x_94 = lean_ctor_get(x_90, 1); +lean_inc(x_94); +lean_dec(x_90); +x_95 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_95, 0, x_91); +lean_ctor_set(x_95, 1, x_94); +return x_95; +} +} +else +{ +uint8_t x_96; +x_96 = !lean_is_exclusive(x_90); +if (x_96 == 0) +{ +lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; uint8_t x_101; +x_97 = lean_ctor_get(x_90, 1); +x_98 = lean_ctor_get(x_90, 0); +lean_dec(x_98); +x_99 = lean_ctor_get(x_91, 0); +lean_inc(x_99); +x_100 = l_Lean_PrettyPrinter_FormatterM_orElse___rarg___closed__1; +x_101 = lean_nat_dec_eq(x_100, x_99); +lean_dec(x_99); +if (x_101 == 0) +{ +lean_dec(x_5); +lean_dec(x_4); +return x_90; +} +else +{ +lean_object* x_102; lean_object* x_103; +lean_free_object(x_90); +lean_dec(x_91); +x_102 = l_Lean_PrettyPrinter_format___lambda__4___closed__5; +x_103 = l_Lean_throwError___at_Lean_PrettyPrinter_format___spec__1(x_102, x_4, x_5, x_97); +lean_dec(x_5); +lean_dec(x_4); +return x_103; +} +} +else +{ +lean_object* x_104; lean_object* x_105; lean_object* x_106; uint8_t x_107; +x_104 = lean_ctor_get(x_90, 1); +lean_inc(x_104); +lean_dec(x_90); +x_105 = lean_ctor_get(x_91, 0); +lean_inc(x_105); +x_106 = l_Lean_PrettyPrinter_FormatterM_orElse___rarg___closed__1; +x_107 = lean_nat_dec_eq(x_106, x_105); +lean_dec(x_105); +if (x_107 == 0) +{ +lean_object* x_108; +lean_dec(x_5); +lean_dec(x_4); +x_108 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_108, 0, x_91); +lean_ctor_set(x_108, 1, x_104); +return x_108; +} +else +{ +lean_object* x_109; lean_object* x_110; +lean_dec(x_91); +x_109 = l_Lean_PrettyPrinter_format___lambda__4___closed__5; +x_110 = l_Lean_throwError___at_Lean_PrettyPrinter_format___spec__1(x_109, x_4, x_5, x_104); +lean_dec(x_5); +lean_dec(x_4); +return x_110; +} +} +} +} +} +} +} +} +else +{ +lean_object* x_115; lean_dec(x_20); -x_43 = lean_ctor_get(x_23, 0); -lean_inc(x_43); -if (lean_obj_tag(x_43) == 0) +lean_dec(x_7); +x_115 = lean_ctor_get(x_23, 0); +lean_inc(x_115); +if (lean_obj_tag(x_115) == 0) { -uint8_t x_44; +uint8_t x_116; lean_dec(x_5); lean_dec(x_4); -x_44 = !lean_is_exclusive(x_23); -if (x_44 == 0) +x_116 = !lean_is_exclusive(x_23); +if (x_116 == 0) { -lean_object* x_45; -x_45 = lean_ctor_get(x_23, 0); -lean_dec(x_45); +lean_object* x_117; +x_117 = lean_ctor_get(x_23, 0); +lean_dec(x_117); return x_23; } else { -lean_object* x_46; lean_object* x_47; -x_46 = lean_ctor_get(x_23, 1); -lean_inc(x_46); +lean_object* x_118; lean_object* x_119; +x_118 = lean_ctor_get(x_23, 1); +lean_inc(x_118); lean_dec(x_23); -x_47 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_47, 0, x_43); -lean_ctor_set(x_47, 1, x_46); -return x_47; +x_119 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_119, 0, x_115); +lean_ctor_set(x_119, 1, x_118); +return x_119; } } else { -uint8_t x_48; -x_48 = !lean_is_exclusive(x_23); -if (x_48 == 0) +uint8_t x_120; +x_120 = !lean_is_exclusive(x_23); +if (x_120 == 0) { -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_23, 1); -x_50 = lean_ctor_get(x_23, 0); -lean_dec(x_50); -x_51 = lean_ctor_get(x_43, 0); -lean_inc(x_51); -x_52 = l_Lean_PrettyPrinter_FormatterM_orElse___rarg___closed__1; -x_53 = lean_nat_dec_eq(x_52, x_51); -lean_dec(x_51); -if (x_53 == 0) +lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; uint8_t x_125; +x_121 = lean_ctor_get(x_23, 1); +x_122 = lean_ctor_get(x_23, 0); +lean_dec(x_122); +x_123 = lean_ctor_get(x_115, 0); +lean_inc(x_123); +x_124 = l_Lean_PrettyPrinter_FormatterM_orElse___rarg___closed__1; +x_125 = lean_nat_dec_eq(x_124, x_123); +lean_dec(x_123); +if (x_125 == 0) { lean_dec(x_5); lean_dec(x_4); @@ -10341,46 +10854,46 @@ return x_23; } else { -lean_object* x_54; lean_object* x_55; +lean_object* x_126; lean_object* x_127; lean_free_object(x_23); -lean_dec(x_43); -x_54 = l_Lean_PrettyPrinter_format___lambda__1___closed__3; -x_55 = l_Lean_throwError___at_Lean_PrettyPrinter_format___spec__2(x_54, x_4, x_5, x_49); +lean_dec(x_115); +x_126 = l_Lean_PrettyPrinter_format___lambda__4___closed__5; +x_127 = l_Lean_throwError___at_Lean_PrettyPrinter_format___spec__1(x_126, x_4, x_5, x_121); lean_dec(x_5); lean_dec(x_4); -return x_55; +return x_127; } } else { -lean_object* x_56; lean_object* x_57; lean_object* x_58; uint8_t x_59; -x_56 = lean_ctor_get(x_23, 1); -lean_inc(x_56); +lean_object* x_128; lean_object* x_129; lean_object* x_130; uint8_t x_131; +x_128 = lean_ctor_get(x_23, 1); +lean_inc(x_128); lean_dec(x_23); -x_57 = lean_ctor_get(x_43, 0); -lean_inc(x_57); -x_58 = l_Lean_PrettyPrinter_FormatterM_orElse___rarg___closed__1; -x_59 = lean_nat_dec_eq(x_58, x_57); -lean_dec(x_57); -if (x_59 == 0) -{ -lean_object* x_60; +x_129 = lean_ctor_get(x_115, 0); +lean_inc(x_129); +x_130 = l_Lean_PrettyPrinter_FormatterM_orElse___rarg___closed__1; +x_131 = lean_nat_dec_eq(x_130, x_129); +lean_dec(x_129); +if (x_131 == 0) +{ +lean_object* x_132; lean_dec(x_5); lean_dec(x_4); -x_60 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_60, 0, x_43); -lean_ctor_set(x_60, 1, x_56); -return x_60; +x_132 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_132, 0, x_115); +lean_ctor_set(x_132, 1, x_128); +return x_132; } else { -lean_object* x_61; lean_object* x_62; -lean_dec(x_43); -x_61 = l_Lean_PrettyPrinter_format___lambda__1___closed__3; -x_62 = l_Lean_throwError___at_Lean_PrettyPrinter_format___spec__2(x_61, x_4, x_5, x_56); +lean_object* x_133; lean_object* x_134; +lean_dec(x_115); +x_133 = l_Lean_PrettyPrinter_format___lambda__4___closed__5; +x_134 = l_Lean_throwError___at_Lean_PrettyPrinter_format___spec__1(x_133, x_4, x_5, x_128); lean_dec(x_5); lean_dec(x_4); -return x_62; +return x_134; } } } @@ -10411,7 +10924,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_format(lean_object* x_1, lean_obje { lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; x_6 = l_Lean_PrettyPrinter_format___closed__2; -x_7 = l_Lean_isTracingEnabledFor___at_Lean_PrettyPrinter_format___spec__1(x_6, x_3, x_4, x_5); +x_7 = l_Lean_isTracingEnabledFor___at_Lean_addDecl___spec__8(x_6, x_3, x_4, x_5); x_8 = lean_ctor_get(x_7, 0); lean_inc(x_8); x_9 = lean_unbox(x_8); @@ -10423,7 +10936,7 @@ x_10 = lean_ctor_get(x_7, 1); lean_inc(x_10); lean_dec(x_7); x_11 = lean_box(0); -x_12 = l_Lean_PrettyPrinter_format___lambda__1(x_2, x_1, x_11, x_3, x_4, x_10); +x_12 = l_Lean_PrettyPrinter_format___lambda__4(x_2, x_1, x_11, x_3, x_4, x_10); return x_12; } else @@ -10446,47 +10959,69 @@ lean_ctor_set(x_20, 1, x_18); x_21 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_21, 0, x_20); lean_ctor_set(x_21, 1, x_19); -x_22 = l_Lean_addTrace___at_Lean_PrettyPrinter_format___spec__3(x_6, x_21, x_3, x_4, x_13); +x_22 = l_Lean_addTrace___at_Lean_PrettyPrinter_format___spec__2(x_6, x_21, x_3, x_4, 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_Lean_PrettyPrinter_format___lambda__1(x_2, x_1, x_23, x_3, x_4, x_24); +x_25 = l_Lean_PrettyPrinter_format___lambda__4(x_2, x_1, x_23, x_3, x_4, x_24); return x_25; } } } -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_PrettyPrinter_format___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_throwError___at_Lean_PrettyPrinter_format___spec__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_isTracingEnabledFor___at_Lean_PrettyPrinter_format___spec__1(x_1, x_2, x_3, x_4); +x_5 = l_Lean_throwError___at_Lean_PrettyPrinter_format___spec__1(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_throwError___at_Lean_PrettyPrinter_format___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_addTrace___at_Lean_PrettyPrinter_format___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_5; -x_5 = l_Lean_throwError___at_Lean_PrettyPrinter_format___spec__2(x_1, x_2, x_3, x_4); +lean_object* x_6; +x_6 = l_Lean_addTrace___at_Lean_PrettyPrinter_format___spec__2(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -return x_5; +return x_6; } } -LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_PrettyPrinter_format___spec__3___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_PrettyPrinter_format___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_addTrace___at_Lean_PrettyPrinter_format___spec__3(x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean_PrettyPrinter_format___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_Lean_PrettyPrinter_format___lambda__2___boxed(lean_object* x_1) { +_start: +{ +uint32_t x_2; uint8_t x_3; lean_object* x_4; +x_2 = lean_unbox_uint32(x_1); +lean_dec(x_1); +x_3 = l_Lean_PrettyPrinter_format___lambda__2(x_2); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_format___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) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_PrettyPrinter_format___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_4); +lean_dec(x_2); +return x_8; +} +} LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_formatCategory(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { @@ -10578,7 +11113,7 @@ x_6 = l_Lean_PrettyPrinter_formatCategory(x_5, x_1, x_2, x_3, x_4); return x_6; } } -static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__1() { +static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -10588,17 +11123,17 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__2() { +static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__1; +x_1 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__1; x_2 = l_Lean_PrettyPrinter_mkFormatterAttribute___closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__3() { +static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__3() { _start: { lean_object* x_1; @@ -10606,17 +11141,17 @@ x_1 = lean_mk_string_from_bytes("initFn", 6); return x_1; } } -static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__4() { +static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__2; -x_2 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__3; +x_1 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__2; +x_2 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__3; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__5() { +static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__5() { _start: { lean_object* x_1; @@ -10624,47 +11159,47 @@ x_1 = lean_mk_string_from_bytes("_@", 2); return x_1; } } -static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__6() { +static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__4; -x_2 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__5; +x_1 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__4; +x_2 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__7() { +static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__6; +x_1 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__6; x_2 = l_Lean_PrettyPrinter_mkFormatterAttribute___closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__8() { +static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__7; +x_1 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__7; x_2 = l_Lean_PrettyPrinter_mkFormatterAttribute___closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__9() { +static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__8; +x_1 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__8; x_2 = l_Lean_PrettyPrinter_mkFormatterAttribute___closed__7; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__10() { +static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__10() { _start: { lean_object* x_1; @@ -10672,33 +11207,33 @@ x_1 = lean_mk_string_from_bytes("_hyg", 4); return x_1; } } -static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__11() { +static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__9; -x_2 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__10; +x_1 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__9; +x_2 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__10; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__12() { +static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__11; -x_2 = lean_unsigned_to_nat(7025u); +x_1 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__11; +x_2 = lean_unsigned_to_nat(7272u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; x_2 = l_Lean_PrettyPrinter_Formatter_categoryFormatterCore___closed__2; x_3 = 0; -x_4 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__12; +x_4 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__12; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); if (lean_obj_tag(x_5) == 0) { @@ -11033,19 +11568,44 @@ l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Formatter_interpolatedStr_form 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_6389_(lean_io_mk_world()); +if (builtin) {res = l_Lean_PrettyPrinter_Formatter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_6401_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_PrettyPrinter_Formatter_formatterAliasesRef = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_PrettyPrinter_Formatter_formatterAliasesRef); lean_dec_ref(res); }l_Lean_PrettyPrinter_Formatter_registerAlias___closed__1 = _init_l_Lean_PrettyPrinter_Formatter_registerAlias___closed__1(); lean_mark_persistent(l_Lean_PrettyPrinter_Formatter_registerAlias___closed__1); -l_Lean_PrettyPrinter_format___lambda__1___closed__1 = _init_l_Lean_PrettyPrinter_format___lambda__1___closed__1(); -lean_mark_persistent(l_Lean_PrettyPrinter_format___lambda__1___closed__1); -l_Lean_PrettyPrinter_format___lambda__1___closed__2 = _init_l_Lean_PrettyPrinter_format___lambda__1___closed__2(); -lean_mark_persistent(l_Lean_PrettyPrinter_format___lambda__1___closed__2); -l_Lean_PrettyPrinter_format___lambda__1___closed__3 = _init_l_Lean_PrettyPrinter_format___lambda__1___closed__3(); -lean_mark_persistent(l_Lean_PrettyPrinter_format___lambda__1___closed__3); +l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_6740____closed__1 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_6740____closed__1(); +lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_6740____closed__1); +l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_6740____closed__2 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_6740____closed__2(); +lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_6740____closed__2); +l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_6740____closed__3 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_6740____closed__3(); +lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_6740____closed__3); +l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_6740____closed__4 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_6740____closed__4(); +lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_6740____closed__4); +l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_6740____closed__5 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_6740____closed__5(); +lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_6740____closed__5); +l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_6740____closed__6 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_6740____closed__6(); +lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_6740____closed__6); +if (builtin) {res = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_6740_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +l_Lean_PrettyPrinter_pp_oneline = lean_io_result_get_value(res); +lean_mark_persistent(l_Lean_PrettyPrinter_pp_oneline); +lean_dec_ref(res); +}l_Lean_PrettyPrinter_format___lambda__4___closed__1 = _init_l_Lean_PrettyPrinter_format___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_PrettyPrinter_format___lambda__4___closed__1); +l_Lean_PrettyPrinter_format___lambda__4___closed__2 = _init_l_Lean_PrettyPrinter_format___lambda__4___closed__2(); +lean_mark_persistent(l_Lean_PrettyPrinter_format___lambda__4___closed__2); +l_Lean_PrettyPrinter_format___lambda__4___closed__3 = _init_l_Lean_PrettyPrinter_format___lambda__4___closed__3(); +lean_mark_persistent(l_Lean_PrettyPrinter_format___lambda__4___closed__3); +l_Lean_PrettyPrinter_format___lambda__4___closed__4 = _init_l_Lean_PrettyPrinter_format___lambda__4___closed__4(); +lean_mark_persistent(l_Lean_PrettyPrinter_format___lambda__4___closed__4); +l_Lean_PrettyPrinter_format___lambda__4___closed__5 = _init_l_Lean_PrettyPrinter_format___lambda__4___closed__5(); +lean_mark_persistent(l_Lean_PrettyPrinter_format___lambda__4___closed__5); +l_Lean_PrettyPrinter_format___lambda__4___closed__6 = _init_l_Lean_PrettyPrinter_format___lambda__4___closed__6(); +lean_mark_persistent(l_Lean_PrettyPrinter_format___lambda__4___closed__6); +l_Lean_PrettyPrinter_format___lambda__4___closed__7 = _init_l_Lean_PrettyPrinter_format___lambda__4___closed__7(); +lean_mark_persistent(l_Lean_PrettyPrinter_format___lambda__4___closed__7); l_Lean_PrettyPrinter_format___closed__1 = _init_l_Lean_PrettyPrinter_format___closed__1(); lean_mark_persistent(l_Lean_PrettyPrinter_format___closed__1); l_Lean_PrettyPrinter_format___closed__2 = _init_l_Lean_PrettyPrinter_format___closed__2(); @@ -11062,31 +11622,31 @@ l_Lean_PrettyPrinter_formatCommand___closed__1 = _init_l_Lean_PrettyPrinter_form lean_mark_persistent(l_Lean_PrettyPrinter_formatCommand___closed__1); l_Lean_PrettyPrinter_formatCommand___closed__2 = _init_l_Lean_PrettyPrinter_formatCommand___closed__2(); lean_mark_persistent(l_Lean_PrettyPrinter_formatCommand___closed__2); -l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__1 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__1(); -lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__1); -l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__2 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__2(); -lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__2); -l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__3 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__3(); -lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__3); -l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__4 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__4(); -lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__4); -l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__5 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__5(); -lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__5); -l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__6 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__6(); -lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__6); -l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__7 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__7(); -lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__7); -l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__8 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__8(); -lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__8); -l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__9 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__9(); -lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__9); -l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__10 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__10(); -lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__10); -l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__11 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__11(); -lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__11); -l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__12 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__12(); -lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025____closed__12); -res = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7025_(lean_io_mk_world()); +l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__1 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__1(); +lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__1); +l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__2 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__2(); +lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__2); +l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__3 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__3(); +lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__3); +l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__4 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__4(); +lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__4); +l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__5 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__5(); +lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__5); +l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__6 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__6(); +lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__6); +l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__7 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__7(); +lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__7); +l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__8 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__8(); +lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__8); +l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__9 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__9(); +lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__9); +l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__10 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__10(); +lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__10); +l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__11 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__11(); +lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__11); +l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__12 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__12(); +lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272____closed__12); +res = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Formatter___hyg_7272_(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/Parenthesizer.c b/stage0/stdlib/Lean/PrettyPrinter/Parenthesizer.c index ee56d81c88a9..e2abe8302e52 100644 --- a/stage0/stdlib/Lean/PrettyPrinter/Parenthesizer.c +++ b/stage0/stdlib/Lean/PrettyPrinter/Parenthesizer.c @@ -45,12 +45,14 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_sepBy1NoAntiquot_par static lean_object* l_Lean_PrettyPrinter_mkCategoryParenthesizerAttribute___closed__9; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_mkAntiquot_parenthesizer_x27___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_mkParenthesizerAttribute___lambda__3___closed__1; +static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__3; LEAN_EXPORT lean_object* l_Lean_Syntax_MonadTraverser_goLeft___at_Lean_PrettyPrinter_Parenthesizer_visitArgs___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_checkStackTop_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_mkParenthesizerAttribute___closed__12; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_pushNone_parenthesizer(lean_object*); static lean_object* l_Lean_PrettyPrinter_Parenthesizer_term_parenthesizer___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_parenthesizeTerm(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__6; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_strLitNoAntiquot_parenthesizer___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_skip_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getHeadInfo(lean_object*); @@ -84,9 +86,9 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_rawCh_parenthesizer_ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_symbolNoAntiquot_parenthesizer___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forM___at_Lean_PrettyPrinter_Parenthesizer_sepByNoAntiquot_parenthesizer___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_Syntax_MonadTraverser_getCur___at_Lean_PrettyPrinter_Parenthesizer_visitArgs___spec__1(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566_(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_many1Unbox_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_forM_loop___at_Lean_PrettyPrinter_Parenthesizer_manyNoAntiquot_parenthesizer___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__8; LEAN_EXPORT uint8_t l_Lean_PrettyPrinter_Parenthesizer_State_visitedToken___default; uint8_t lean_usize_dec_eq(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Syntax_MonadTraverser_goRight___at_Lean_PrettyPrinter_Parenthesizer_maybeParenthesize___spec__5___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -101,11 +103,15 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_trailingNode_parenth lean_object* lean_pretty_printer_parenthesizer_interpret_parser_descr(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_PrettyPrinter_Parenthesizer_instMonadQuotationParenthesizerM___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_fieldIdx_parenthesizer___boxed(lean_object*); +static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__8; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_level_parenthesizer___lambda__1(lean_object*); static lean_object* l_Lean_PrettyPrinter_Parenthesizer_checkKind___closed__4; LEAN_EXPORT lean_object* l_Lean_Syntax_MonadTraverser_goUp___at_Lean_PrettyPrinter_Parenthesizer_visitArgs___spec__4___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_hygieneInfoNoAntiquot_parenthesizer___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_5968_(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_maybeParenthesize___lambda__9___boxed(lean_object**); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_instMonadQuotationParenthesizerM___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__9; static lean_object* l_Lean_PrettyPrinter_Parenthesizer_maybeParenthesize___lambda__10___closed__4; static lean_object* l_Lean_PrettyPrinter_mkParenthesizerAttribute___lambda__3___closed__2; static lean_object* l_Lean_PrettyPrinter_parenthesizeCommand___closed__1; @@ -131,9 +137,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_rawCh_parenthesizer( static lean_object* l_Lean_PrettyPrinter_Parenthesizer_maybeParenthesize___closed__5; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_errorAtSavedPos_parenthesizer___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_mkParenthesizerAttribute___lambda__4(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__1; LEAN_EXPORT lean_object* l_Lean_Syntax_MonadTraverser_getCur___at_Lean_PrettyPrinter_Parenthesizer_visitArgs___spec__1___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_PrettyPrinter_parenthesize___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Parenthesizer_parenthesizerForKindUnsafe___closed__3; @@ -152,7 +156,6 @@ LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_PrettyPrinter_Parenthesizer_in LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_trailingNode_parenthesizer___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_Parenthesizer_interpolatedStr_parenthesizer___closed__1; lean_object* l_Lean_Syntax_Traverser_left(lean_object*); -static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__6; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_rawIdentNoAntiquot_parenthesizer___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_withFn_parenthesizer(lean_object*); static lean_object* l_Lean_PrettyPrinter_mkCombinatorParenthesizerAttribute___closed__4; @@ -168,6 +171,7 @@ lean_object* lean_string_utf8_byte_size(lean_object*); static lean_object* l_Lean_PrettyPrinter_Parenthesizer_term_parenthesizer___lambda__1___closed__4; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_level_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_liftCoreM___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_PrettyPrinter_parenthesize___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_checkColGe_parenthesizer___rarg(lean_object*); static lean_object* l_Lean_PrettyPrinter_Parenthesizer_tactic_parenthesizer___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_categoryParser_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -175,11 +179,9 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_checkLinebreakBefore static lean_object* l_Lean_PrettyPrinter_Parenthesizer_orelse_parenthesizer___closed__1; extern lean_object* l_instInhabitedNat; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_mkCombinatorParenthesizerAttribute(lean_object*); -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558_(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_fieldIdx_parenthesizer___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_mkCategoryParenthesizerAttribute___closed__6; static lean_object* l_Lean_PrettyPrinter_mkCombinatorParenthesizerAttribute___closed__5; -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_PrettyPrinter_parenthesize___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_withAntiquotSuffixSplice_parenthesizer___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_maybeParenthesize___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_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_checkColGt_parenthesizer___rarg(lean_object*); @@ -190,6 +192,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_withPositionAfterLin LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_tactic_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_isAntiquot(lean_object*); LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_PrettyPrinter_Parenthesizer_maybeParenthesize___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__5; size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_checkNoWsBefore_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_instMonadTraverserParenthesizerM___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -235,7 +238,6 @@ static lean_object* l_Lean_PrettyPrinter_Parenthesizer_maybeParenthesize___lambd LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_parenthesizerAttribute; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_instMonadQuotationParenthesizerM; lean_object* l_Lean_Syntax_getKind(lean_object*); -LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_PrettyPrinter_parenthesize___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_optionalNoAntiquot_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Parenthesizer_interpolatedStr_parenthesizer___spec__1___closed__2; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_error_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -256,6 +258,8 @@ static lean_object* l_Lean_PrettyPrinter_Parenthesizer_maybeParenthesize___lambd static lean_object* l_Lean_PrettyPrinter_Parenthesizer_term_parenthesizer___lambda__1___closed__6; lean_object* lean_st_ref_get(lean_object*, lean_object*); lean_object* l_Lean_throwError___at_Lean_registerTagAttribute___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_PrettyPrinter_parenthesize___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__12; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_lookahead_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Attribute_Builtin_getIdent(lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Parser_parserExtension; @@ -269,7 +273,6 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_mkCategoryParenthesizerAttribute(l LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_mkParenthesizerAttribute___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_parenthesize___closed__1; lean_object* l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__10; static lean_object* l_Lean_PrettyPrinter_Parenthesizer_maybeParenthesize___lambda__10___closed__20; lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); lean_object* l_Lean_Syntax_node3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -286,6 +289,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_maybeParenthesize___ LEAN_EXPORT lean_object* l_Lean_Syntax_MonadTraverser_goUp___at_Lean_PrettyPrinter_Parenthesizer_visitArgs___spec__4___boxed(lean_object*); static lean_object* l_Option_format___at_Lean_PrettyPrinter_Parenthesizer_maybeParenthesize___spec__6___closed__2; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_visitArgs___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__10; static lean_object* l_Lean_PrettyPrinter_Parenthesizer_checkKind___closed__1; LEAN_EXPORT lean_object* l_Lean_Syntax_MonadTraverser_getIdx___at_Lean_PrettyPrinter_Parenthesizer_maybeParenthesize___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_checkColEq_parenthesizer___rarg(lean_object*); @@ -296,9 +300,10 @@ static lean_object* l_Lean_PrettyPrinter_Parenthesizer_maybeParenthesize___lambd static lean_object* l_Lean_PrettyPrinter_mkCategoryParenthesizerAttribute___closed__1; static lean_object* l_Lean_PrettyPrinter_Parenthesizer_level_parenthesizer___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_withPosition_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_isTracingEnabledFor___at_Lean_addDecl___spec__8(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_rawIdentNoAntiquot_parenthesizer___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_checkColGt_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__9; +static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__11; static lean_object* l_Lean_PrettyPrinter_mkCategoryParenthesizerAttribute___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_errorAtSavedPos_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Parenthesizer_rawStx_parenthesizer___closed__1; @@ -354,7 +359,6 @@ extern lean_object* l_Lean_inheritedTraceOptions; static lean_object* l_Lean_PrettyPrinter_mkParenthesizerAttribute___closed__13; static lean_object* l_Lean_PrettyPrinter_Parenthesizer_checkKind___closed__3; static lean_object* l_Lean_PrettyPrinter_mkCategoryParenthesizerAttribute___closed__2; -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_5956_(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Parenthesizer_interpolatedStr_parenthesizer___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_PrettyPrinter_Parenthesizer_level_parenthesizer___closed__3; extern lean_object* l_Lean_Core_instMonadQuotationCoreM; @@ -368,7 +372,6 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_instMonadTraverserPa static lean_object* l_Option_format___at_Lean_PrettyPrinter_Parenthesizer_maybeParenthesize___spec__6___closed__4; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_interpolatedStr_parenthesizer___boxed__const__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_ite(lean_object*, lean_object*); -static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__7; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_State_minPrec___default; static lean_object* l_Lean_PrettyPrinter_Parenthesizer_tactic_parenthesizer___closed__2; static lean_object* l_Lean_PrettyPrinter_Parenthesizer_term_parenthesizer___closed__2; @@ -377,6 +380,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_unicodeSymbolNoAntiq static lean_object* l_Lean_PrettyPrinter_Parenthesizer_maybeParenthesize___lambda__10___closed__12; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_instMonadTraverserParenthesizerM___lambda__2___boxed(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_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_hygieneInfoNoAntiquot_parenthesizer___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ParserCompiler_registerCombinatorAttribute(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_instCoeForAllParenthesizerParenthesizerAliasValue__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_withoutInfo_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -386,14 +390,12 @@ lean_object* l_Lean_Parser_registerAliasCore___rarg(lean_object*, lean_object*, LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_throwBacktrack___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_rawStx_parenthesizer___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Parenthesizer_level_parenthesizer(lean_object*); -static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__12; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_combinatorParenthesizerAttribute; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_charLitNoAntiquot_parenthesizer___rarg(lean_object*, lean_object*, lean_object*, lean_object*); 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_mkParenthesizerAttribute___closed__5; lean_object* lean_string_length(lean_object*); static lean_object* l_Lean_PrettyPrinter_Parenthesizer_parenthesizerForKindUnsafe___closed__2; -LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_PrettyPrinter_parenthesize___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_scientificLitNoAntiquot_parenthesizer___boxed(lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Parenthesizer_rawStx_parenthesizer(lean_object*); @@ -406,11 +408,13 @@ static lean_object* l_Lean_PrettyPrinter_Parenthesizer_instMonadQuotationParenth 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_EXPORT lean_object* l_Lean_addTrace___at_Lean_PrettyPrinter_Parenthesizer_maybeParenthesize___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); +static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__7; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_parenthesizerAliasesRef; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Parenthesizer_level_parenthesizer___closed__2; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_maybeParenthesize(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_PrettyPrinter_Parenthesizer_rawIdentNoAntiquot_parenthesizer(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_checkWsBefore_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__4; static lean_object* l_Lean_PrettyPrinter_mkParenthesizerAttribute___closed__6; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_scientificLitNoAntiquot_parenthesizer___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Parenthesizer_maybeParenthesize___lambda__10___closed__17; @@ -430,8 +434,8 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_charLitNoAntiquot_pa LEAN_EXPORT lean_object* l_Lean_Syntax_MonadTraverser_getIdx___at_Lean_PrettyPrinter_Parenthesizer_maybeParenthesize___spec__1(lean_object*); static lean_object* l_Lean_addTrace___at_Lean_PrettyPrinter_Parenthesizer_maybeParenthesize___spec__3___closed__1; static lean_object* l_Lean_PrettyPrinter_Parenthesizer_visitArgs___closed__1; -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_PrettyPrinter_parenthesize___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); +static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__2; lean_object* l_Lean_ScopedEnvExtension_getState___rarg(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_Traverser_setCur(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_checkTailWs_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*); @@ -457,10 +461,10 @@ static lean_object* l_Lean_PrettyPrinter_Parenthesizer_throwBacktrack___rarg___c static lean_object* l_Lean_PrettyPrinter_Parenthesizer_instMonadQuotationParenthesizerM___closed__6; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_parenthesizeCategoryCore___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_instCoeForAllParenthesizerParenthesizerAliasValue(lean_object*); +static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__1; static lean_object* l_Lean_PrettyPrinter_Parenthesizer_maybeParenthesize___lambda__10___closed__15; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_parenthesize(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_mkParenthesizerAttribute___lambda__3___closed__3; -static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__3; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_strLitNoAntiquot_parenthesizer___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_tokenWithAntiquot_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_maybeParenthesize___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -471,7 +475,6 @@ size_t lean_usize_add(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_notFollowedBy_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Parenthesizer_checkKind___closed__2; lean_object* l_Lean_Syntax_Traverser_up(lean_object*); -static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__2; 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_node_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_lookahead_parenthesizer___rarg(lean_object*); @@ -490,7 +493,7 @@ static lean_object* l_Lean_PrettyPrinter_Parenthesizer_instMonadQuotationParenth static lean_object* l_Lean_PrettyPrinter_Parenthesizer_instMonadQuotationParenthesizerM___closed__2; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_rawIdentNoAntiquot_parenthesizer___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_maybeParenthesize___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*); -static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__5; +LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_PrettyPrinter_parenthesize___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_parenthesizeTactic(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_withFn_parenthesizer___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Parenthesizer_instMonadTraverserParenthesizerM___closed__3; @@ -499,7 +502,6 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_mkParenthesizerAttribute___lambda_ lean_object* l_Lean_KeyedDeclsAttribute_getValues___rarg(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_Syntax_MonadTraverser_goUp___at_Lean_PrettyPrinter_Parenthesizer_visitArgs___spec__4(lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_PrettyPrinter_parenthesize___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Parenthesizer_maybeParenthesize___closed__3; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Parenthesizer_rawStx_parenthesizer___closed__3; uint8_t l_Lean_Parser_isValidSyntaxNodeKind(lean_object*, lean_object*); @@ -509,7 +511,6 @@ lean_object* lean_array_get_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_MonadTraverser_goRight___at_Lean_PrettyPrinter_Parenthesizer_maybeParenthesize___spec__5(lean_object*); static lean_object* l_Lean_PrettyPrinter_ParenthesizerM_orElse___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_nonReservedSymbolNoAntiquot_parenthesizer___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__11; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_maybeParenthesize___lambda__9(lean_object*, lean_object*, uint8_t, 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_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_charLitNoAntiquot_parenthesizer(lean_object*); static lean_object* l_Lean_PrettyPrinter_Parenthesizer_maybeParenthesize___lambda__4___closed__3; @@ -529,7 +530,9 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_numLitNoAntiquot_par LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_checkLineEq_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_checkWsBefore_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Parenthesizer_maybeParenthesize___lambda__5___closed__2; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_hygieneInfoNoAntiquot_parenthesizer(lean_object*); lean_object* l_Lean_MessageData_ofName(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_hygieneInfoNoAntiquot_parenthesizer___boxed(lean_object*); static lean_object* l_Lean_PrettyPrinter_Parenthesizer_categoryParser_parenthesizer___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_many1NoAntiquot_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_mkParenthesizerAttribute___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -541,9 +544,9 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_throwBacktrack(lean_ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_checkTailWs_parenthesizer___rarg(lean_object*); static lean_object* l_Lean_PrettyPrinter_mkParenthesizerAttribute___closed__11; lean_object* l_Nat_repr(lean_object*); -static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__4; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_parenthesizerForKindUnsafe(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Parenthesizer_identNoAntiquot_parenthesizer___closed__1; +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_PrettyPrinter_parenthesize___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_nameLitNoAntiquot_parenthesizer___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_isAntiquotSplice(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_checkStackTop_parenthesizer___rarg(lean_object*); @@ -9427,6 +9430,42 @@ lean_dec(x_1); return x_2; } } +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_hygieneInfoNoAntiquot_parenthesizer___rarg(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_Syntax_MonadTraverser_goLeft___at_Lean_PrettyPrinter_Parenthesizer_visitArgs___spec__2___rarg(x_1, x_2, x_3, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_hygieneInfoNoAntiquot_parenthesizer(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Parenthesizer_hygieneInfoNoAntiquot_parenthesizer___rarg___boxed), 4, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_hygieneInfoNoAntiquot_parenthesizer___rarg___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_PrettyPrinter_Parenthesizer_hygieneInfoNoAntiquot_parenthesizer___rarg(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_PrettyPrinter_Parenthesizer_hygieneInfoNoAntiquot_parenthesizer___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_PrettyPrinter_Parenthesizer_hygieneInfoNoAntiquot_parenthesizer(x_1); +lean_dec(x_1); +return x_2; +} +} static lean_object* _init_l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Parenthesizer_interpolatedStr_parenthesizer___spec__1___closed__1() { _start: { @@ -9670,7 +9709,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_5956_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_5968_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; uint8_t x_4; @@ -9740,44 +9779,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_PrettyPrinter_parenthesize___spec__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; uint8_t x_7; -x_5 = l_Lean_isTracingEnabledFor___at_Lean_PrettyPrinter_Parenthesizer_maybeParenthesize___spec__2___closed__1; -x_6 = lean_st_ref_get(x_5, x_4); -x_7 = !lean_is_exclusive(x_6); -if (x_7 == 0) -{ -lean_object* x_8; lean_object* x_9; uint8_t x_10; lean_object* x_11; -x_8 = lean_ctor_get(x_6, 0); -x_9 = lean_ctor_get(x_2, 2); -x_10 = l___private_Lean_Util_Trace_0__Lean_checkTraceOption(x_8, x_9, x_1); -lean_dec(x_8); -x_11 = lean_box(x_10); -lean_ctor_set(x_6, 0, x_11); -return x_6; -} -else -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; -x_12 = lean_ctor_get(x_6, 0); -x_13 = lean_ctor_get(x_6, 1); -lean_inc(x_13); -lean_inc(x_12); -lean_dec(x_6); -x_14 = lean_ctor_get(x_2, 2); -x_15 = l___private_Lean_Util_Trace_0__Lean_checkTraceOption(x_12, x_14, x_1); -lean_dec(x_12); -x_16 = lean_box(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_13); -return x_17; -} -} -} -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_PrettyPrinter_parenthesize___spec__2(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_PrettyPrinter_parenthesize___spec__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; uint8_t x_7; @@ -9815,7 +9817,7 @@ return x_13; } } } -LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_PrettyPrinter_parenthesize___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_addTrace___at_Lean_PrettyPrinter_parenthesize___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; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; @@ -10079,7 +10081,7 @@ lean_object* x_37; lean_object* x_38; lean_free_object(x_15); lean_dec(x_26); x_37 = l_Lean_PrettyPrinter_parenthesize___lambda__1___closed__2; -x_38 = l_Lean_throwError___at_Lean_PrettyPrinter_parenthesize___spec__2(x_37, x_4, x_5, x_32); +x_38 = l_Lean_throwError___at_Lean_PrettyPrinter_parenthesize___spec__1(x_37, x_4, x_5, x_32); lean_dec(x_5); lean_dec(x_4); return x_38; @@ -10111,7 +10113,7 @@ else lean_object* x_44; lean_object* x_45; lean_dec(x_26); x_44 = l_Lean_PrettyPrinter_parenthesize___lambda__1___closed__2; -x_45 = l_Lean_throwError___at_Lean_PrettyPrinter_parenthesize___spec__2(x_44, x_4, x_5, x_39); +x_45 = l_Lean_throwError___at_Lean_PrettyPrinter_parenthesize___spec__1(x_44, x_4, x_5, x_39); lean_dec(x_5); lean_dec(x_4); return x_45; @@ -10145,7 +10147,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_parenthesize(lean_object* x_1, lea { lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; x_6 = l_Lean_PrettyPrinter_parenthesize___closed__2; -x_7 = l_Lean_isTracingEnabledFor___at_Lean_PrettyPrinter_parenthesize___spec__1(x_6, x_3, x_4, x_5); +x_7 = l_Lean_isTracingEnabledFor___at_Lean_addDecl___spec__8(x_6, x_3, x_4, x_5); x_8 = lean_ctor_get(x_7, 0); lean_inc(x_8); x_9 = lean_unbox(x_8); @@ -10180,7 +10182,7 @@ lean_ctor_set(x_20, 1, x_18); x_21 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_21, 0, x_20); lean_ctor_set(x_21, 1, x_19); -x_22 = l_Lean_addTrace___at_Lean_PrettyPrinter_parenthesize___spec__3(x_6, x_21, x_3, x_4, x_13); +x_22 = l_Lean_addTrace___at_Lean_PrettyPrinter_parenthesize___spec__2(x_6, x_21, x_3, x_4, x_13); x_23 = lean_ctor_get(x_22, 0); lean_inc(x_23); x_24 = lean_ctor_get(x_22, 1); @@ -10191,31 +10193,21 @@ return x_25; } } } -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_PrettyPrinter_parenthesize___spec__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_isTracingEnabledFor___at_Lean_PrettyPrinter_parenthesize___spec__1(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_throwError___at_Lean_PrettyPrinter_parenthesize___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_throwError___at_Lean_PrettyPrinter_parenthesize___spec__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_throwError___at_Lean_PrettyPrinter_parenthesize___spec__2(x_1, x_2, x_3, x_4); +x_5 = l_Lean_throwError___at_Lean_PrettyPrinter_parenthesize___spec__1(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_addTrace___at_Lean_PrettyPrinter_parenthesize___spec__3___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_addTrace___at_Lean_PrettyPrinter_parenthesize___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_Lean_addTrace___at_Lean_PrettyPrinter_parenthesize___spec__3(x_1, x_2, x_3, x_4, x_5); +x_6 = l_Lean_addTrace___at_Lean_PrettyPrinter_parenthesize___spec__2(x_1, x_2, x_3, x_4, x_5); lean_dec(x_4); lean_dec(x_3); return x_6; @@ -10278,7 +10270,7 @@ x_6 = l_Lean_PrettyPrinter_parenthesizeCategory(x_5, x_1, x_2, x_3, x_4); return x_6; } } -static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__1() { +static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -10288,17 +10280,17 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__2() { +static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__1; +x_1 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__1; x_2 = l_Lean_PrettyPrinter_mkParenthesizerAttribute___closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__3() { +static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__3() { _start: { lean_object* x_1; @@ -10306,17 +10298,17 @@ x_1 = lean_mk_string_from_bytes("initFn", 6); return x_1; } } -static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__4() { +static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__2; -x_2 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__3; +x_1 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__2; +x_2 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__3; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__5() { +static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__5() { _start: { lean_object* x_1; @@ -10324,47 +10316,47 @@ x_1 = lean_mk_string_from_bytes("_@", 2); return x_1; } } -static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__6() { +static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__4; -x_2 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__5; +x_1 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__4; +x_2 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__7() { +static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__6; +x_1 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__6; x_2 = l_Lean_PrettyPrinter_mkParenthesizerAttribute___closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__8() { +static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__7; +x_1 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__7; x_2 = l_Lean_PrettyPrinter_mkParenthesizerAttribute___closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__9() { +static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__8; +x_1 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__8; x_2 = l_Lean_PrettyPrinter_mkParenthesizerAttribute___closed__7; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__10() { +static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__10() { _start: { lean_object* x_1; @@ -10372,33 +10364,33 @@ x_1 = lean_mk_string_from_bytes("_hyg", 4); return x_1; } } -static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__11() { +static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__9; -x_2 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__10; +x_1 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__9; +x_2 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__10; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__12() { +static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__11; -x_2 = lean_unsigned_to_nat(6558u); +x_1 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__11; +x_2 = lean_unsigned_to_nat(6566u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; x_2 = l_Lean_PrettyPrinter_Parenthesizer_maybeParenthesize___closed__2; x_3 = 0; -x_4 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__12; +x_4 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__12; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); if (lean_obj_tag(x_5) == 0) { @@ -10821,7 +10813,7 @@ 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_5956_(lean_io_mk_world()); +if (builtin) {res = l_Lean_PrettyPrinter_Parenthesizer_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_5968_(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); @@ -10840,31 +10832,31 @@ l_Lean_PrettyPrinter_parenthesizeCommand___closed__1 = _init_l_Lean_PrettyPrinte lean_mark_persistent(l_Lean_PrettyPrinter_parenthesizeCommand___closed__1); l_Lean_PrettyPrinter_parenthesizeCommand___closed__2 = _init_l_Lean_PrettyPrinter_parenthesizeCommand___closed__2(); lean_mark_persistent(l_Lean_PrettyPrinter_parenthesizeCommand___closed__2); -l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__1 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__1(); -lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__1); -l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__2 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__2(); -lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__2); -l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__3 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__3(); -lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__3); -l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__4 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__4(); -lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__4); -l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__5 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__5(); -lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__5); -l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__6 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__6(); -lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__6); -l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__7 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__7(); -lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__7); -l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__8 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__8(); -lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__8); -l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__9 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__9(); -lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__9); -l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__10 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__10(); -lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__10); -l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__11 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__11(); -lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__11); -l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__12 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__12(); -lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558____closed__12); -res = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6558_(lean_io_mk_world()); +l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__1 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__1(); +lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__1); +l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__2 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__2(); +lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__2); +l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__3 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__3(); +lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__3); +l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__4 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__4(); +lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__4); +l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__5 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__5(); +lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__5); +l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__6 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__6(); +lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__6); +l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__7 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__7(); +lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__7); +l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__8 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__8(); +lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__8); +l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__9 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__9(); +lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__9); +l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__10 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__10(); +lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__10); +l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__11 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__11(); +lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__11); +l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__12 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__12(); +lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566____closed__12); +res = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Parenthesizer___hyg_6566_(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/CodeActions.c b/stage0/stdlib/Lean/Server/CodeActions.c index 6d209803a6a7..f385814d6da2 100644 --- a/stage0/stdlib/Lean/Server/CodeActions.c +++ b/stage0/stdlib/Lean/Server/CodeActions.c @@ -1490,7 +1490,7 @@ static lean_object* _init_l_Lean_Server_initFn____x40_Lean_Server_CodeActions___ _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("codeActionProvider", 18); +x_1 = lean_mk_string_from_bytes("code_action_provider", 20); return x_1; } } diff --git a/stage0/stdlib/Lean/Server/FileWorker.c b/stage0/stdlib/Lean/Server/FileWorker.c index 25834607bb6f..98f4143e3617 100644 --- a/stage0/stdlib/Lean/Server/FileWorker.c +++ b/stage0/stdlib/Lean/Server/FileWorker.c @@ -101,7 +101,6 @@ static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_ini LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleNotification(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static uint8_t l_Lean_Server_FileWorker_handleDidChange___closed__2; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleRpcRelease___boxed(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*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_lakeSetupSearchPath___spec__1___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_String_trim(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_updatePendingRequests___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -225,6 +224,7 @@ LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_queueRequest___boxed(lean_obje LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_compileHeader___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_List_toPArray_x27___rarg(lean_object*); lean_object* l_IO_FS_Stream_readLspMessage(lean_object*, lean_object*); +lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(lean_object*, lean_object*); lean_object* l_Lean_Server_Snapshots_Snapshot_diagnostics(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_updatePendingRequests(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_publishIleanInfoFinal(lean_object*, lean_object*, lean_object*, lean_object*); @@ -406,6 +406,7 @@ static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_ini uint8_t l_Lean_RBNode_isBlack___rarg(lean_object*); lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcConnectParams____x40_Lean_Data_Lsp_Extra___hyg_1488_(lean_object*); lean_object* lean_io_bind_task(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_toJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1395_(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_PersistentArray_mapMAux___at_Lean_Server_FileWorker_compileHeader___spec__8(uint8_t, lean_object*, 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*); @@ -496,7 +497,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_0__Lean_Server_FileW lean_object* l___private_Lean_Data_Lsp_Communication_0__IO_FS_Stream_readLspHeader(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at_Lean_Server_FileWorker_handleRpcConnect___spec__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_compileHeader___lambda__5___closed__4; -lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_toJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1412_(lean_object*); static lean_object* l_Lean_Server_FileWorker_initAndRunWorker___closed__6; static lean_object* l_Lean_Json_toStructured_x3f___at___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_publishIleanInfo___spec__3___closed__1; static lean_object* l_Lean_RBNode_foldM___at_Lean_Server_FileWorker_mainLoop___spec__1___closed__1; @@ -807,7 +807,7 @@ LEAN_EXPORT lean_object* l_Lean_Json_toStructured_x3f___at___private_Lean_Server _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_2 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_toJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1412_(x_1); +x_2 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_toJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1395_(x_1); x_3 = lean_ctor_get(x_2, 0); lean_inc(x_3); lean_dec(x_2); @@ -4588,7 +4588,7 @@ lean_dec(x_130); x_131 = lean_ctor_get(x_124, 0); lean_dec(x_131); x_132 = l_Lean_Server_FileWorker_compileHeader___lambda__1___closed__2; -x_133 = l_Lean_Option_get___at_Std_Format_pretty_x27___spec__1(x_1, x_132); +x_133 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_132); lean_dec(x_1); x_134 = lean_ctor_get(x_2, 2); lean_inc(x_134); @@ -4735,7 +4735,7 @@ lean_inc(x_180); lean_inc(x_179); lean_dec(x_124); x_182 = l_Lean_Server_FileWorker_compileHeader___lambda__1___closed__2; -x_183 = l_Lean_Option_get___at_Std_Format_pretty_x27___spec__1(x_1, x_182); +x_183 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_182); lean_dec(x_1); x_184 = lean_ctor_get(x_2, 2); lean_inc(x_184); @@ -4940,7 +4940,7 @@ lean_dec(x_23); x_24 = lean_ctor_get(x_17, 0); lean_dec(x_24); x_25 = l_Lean_Server_FileWorker_compileHeader___lambda__1___closed__2; -x_26 = l_Lean_Option_get___at_Std_Format_pretty_x27___spec__1(x_1, x_25); +x_26 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_25); lean_dec(x_1); x_27 = lean_ctor_get(x_2, 2); lean_inc(x_27); @@ -5096,7 +5096,7 @@ lean_inc(x_73); lean_inc(x_72); lean_dec(x_17); x_75 = l_Lean_Server_FileWorker_compileHeader___lambda__1___closed__2; -x_76 = l_Lean_Option_get___at_Std_Format_pretty_x27___spec__1(x_1, x_75); +x_76 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_75); lean_dec(x_1); x_77 = lean_ctor_get(x_2, 2); lean_inc(x_77); diff --git a/stage0/stdlib/Lean/Server/FileWorker/RequestHandling.c b/stage0/stdlib/Lean/Server/FileWorker/RequestHandling.c index 08cff59e69d8..36ee9657965c 100644 --- a/stage0/stdlib/Lean/Server/FileWorker/RequestHandling.c +++ b/stage0/stdlib/Lean/Server/FileWorker/RequestHandling.c @@ -89,7 +89,6 @@ lean_object* l_panic___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__2(lean_ lean_object* l_List_getLastD___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__3___closed__6; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handlePlainGoal___lambda__1(lean_object*); -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___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_Server_FileWorker_handlePlainGoal___spec__1(size_t, size_t, lean_object*); lean_object* l_Lean_findDocString_x3f(lean_object*, lean_object*, uint8_t, lean_object*); lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2401_(lean_object*); @@ -324,7 +323,6 @@ lean_object* lean_array_to_list(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleWaitForDiagnostics___lambda__2___boxed(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_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentHighlight_highlightReturn_x3f___lambda__1___boxed(lean_object*); -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleSemanticTokens___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_getInteractiveGoals___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_keywordSemanticTokenMap; @@ -437,6 +435,7 @@ LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at_Lean_Server_FileWorker_l LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleDocumentSymbol_toDocumentSymbols___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12936____closed__3; lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3960_(lean_object*); +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_isInstance(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_noHighlightKinds___closed__11; LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12936____spec__26___boxed(lean_object*, lean_object*, lean_object*); @@ -499,6 +498,7 @@ uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12936____spec__36___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleWaitForDiagnostics___closed__1; static lean_object* l_Lean_Server_FileWorker_locationLinksOfInfo___lambda__8___closed__5; +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(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_12936____spec__26(lean_object*, lean_object*, lean_object*); uint8_t lean_uint32_dec_eq(uint32_t, uint32_t); static lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker_RequestHandling___hyg_12936____spec__2___closed__2; @@ -4194,7 +4194,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = l_Lean_Elab_Info_lctx(x_3); -x_17 = lean_alloc_closure((void*)(l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2___boxed), 6, 1); +x_17 = lean_alloc_closure((void*)(l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1___boxed), 6, 1); lean_closure_set(x_17, 0, x_9); lean_inc(x_16); x_18 = l_Lean_Elab_ContextInfo_runMetaM___rarg(x_4, x_16, x_17, x_12); @@ -6095,7 +6095,7 @@ lean_inc(x_8); x_9 = lean_ctor_get(x_7, 1); lean_inc(x_9); lean_dec(x_7); -x_10 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_8, x_2, x_3, x_4, x_5, x_9); +x_10 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_8, x_2, x_3, x_4, x_5, x_9); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); @@ -6171,7 +6171,7 @@ lean_inc(x_8); x_9 = lean_ctor_get(x_7, 1); lean_inc(x_9); lean_dec(x_7); -x_10 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_8, x_2, x_3, x_4, x_5, x_9); +x_10 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_8, x_2, x_3, x_4, x_5, x_9); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); @@ -9165,7 +9165,7 @@ lean_inc(x_10); x_11 = lean_ctor_get(x_8, 1); lean_inc(x_11); lean_dec(x_8); -x_12 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_10, x_3, x_4, x_5, x_6, x_11); +x_12 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_10, x_3, x_4, x_5, x_6, x_11); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); @@ -9181,7 +9181,7 @@ lean_dec(x_8); x_14 = lean_ctor_get(x_9, 0); lean_inc(x_14); lean_dec(x_9); -x_15 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_14, x_3, x_4, x_5, x_6, x_13); +x_15 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_14, x_3, x_4, x_5, x_6, x_13); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); @@ -20052,7 +20052,7 @@ static lean_object* _init_l_Lean_Server_FileWorker_handleSemanticTokensFull___ra { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_unsigned_to_nat(1u); -x_2 = lean_unsigned_to_nat(16u); +x_2 = lean_unsigned_to_nat(31u); x_3 = lean_nat_shiftl(x_1, x_2); return x_3; } diff --git a/stage0/stdlib/Lean/Server/InfoUtils.c b/stage0/stdlib/Lean/Server/InfoUtils.c index b63c4ef8744c..a1e703863d39 100644 --- a/stage0/stdlib/Lean/Server/InfoUtils.c +++ b/stage0/stdlib/Lean/Server/InfoUtils.c @@ -69,7 +69,6 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_InfoTree_foldI LEAN_EXPORT lean_object* l_panic___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__2(lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__20___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_filterTR_loop___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__3(lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_goalsAt_x3f_hasNestedTactic___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_findDocString_x3f(lean_object*, lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -231,7 +230,6 @@ lean_object* l_Lean_getOptionDecls(lean_object*); lean_object* lean_array_to_list(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_lexOrd___at_Lean_Elab_InfoTree_hoverableInfoAt_x3f___spec__10; LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_Elab_InfoTree_goalsAt_x3f___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_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 lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_InfoTree_foldInfo_go___spec__25___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); @@ -306,6 +304,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_collectNodesBottomUp(lean_object*) 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_List_filterTR_loop___at_Lean_Elab_InfoTree_goalsAt_x3f___spec__10(lean_object*, lean_object*, lean_object*); uint8_t l_instDecidableRelLeLeOfOrd___rarg(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1___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__43___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_LocalDecl_userName(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlMAux___at_Lean_Elab_InfoTree_foldInfo_go___spec__23(lean_object*); @@ -352,6 +351,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Info_fmtHover_x3f___lambda__2(lean_object*, uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_InfoTree_hasSorry_go___spec__25(lean_object*, lean_object*, size_t, size_t, lean_object*); 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* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(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__38(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__34___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_goalsAt_x3f_isEmptyBy___boxed(lean_object*); @@ -7521,7 +7521,7 @@ lean_inc(x_9); x_10 = lean_ctor_get(x_8, 1); lean_inc(x_10); lean_dec(x_8); -x_11 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_9, x_3, x_4, x_5, x_6, x_10); +x_11 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_9, x_3, x_4, x_5, x_6, x_10); x_12 = lean_ctor_get(x_11, 0); lean_inc(x_12); x_13 = lean_ctor_get(x_11, 1); @@ -7774,7 +7774,7 @@ lean_dec(x_1); x_8 = lean_ctor_get(x_7, 3); lean_inc(x_8); lean_dec(x_7); -x_9 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_8, x_2, x_3, x_4, x_5, x_6); +x_9 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_8, x_2, x_3, x_4, x_5, x_6); x_10 = !lean_is_exclusive(x_9); if (x_10 == 0) { @@ -16699,7 +16699,7 @@ lean_inc(x_14); lean_dec(x_10); x_15 = lean_ctor_get(x_14, 3); lean_inc(x_15); -x_16 = lean_alloc_closure((void*)(l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2___boxed), 6, 1); +x_16 = lean_alloc_closure((void*)(l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1___boxed), 6, 1); lean_closure_set(x_16, 0, x_15); x_17 = l_Lean_Elab_TermInfo_runMetaM___rarg(x_14, x_13, x_16, x_3); lean_dec(x_13); diff --git a/stage0/stdlib/Lean/Server/References.c b/stage0/stdlib/Lean/Server/References.c index 34b1e7ca2cec..71836b0e5702 100644 --- a/stage0/stdlib/Lean/Server/References.c +++ b/stage0/stdlib/Lean/Server/References.c @@ -142,6 +142,7 @@ LEAN_EXPORT lean_object* l_Lean_HashMapImp_moveEntries___at_Lean_Server_Referenc LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Server_References_removeIlean___spec__5(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_Server_dedupReferences___spec__8(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_InfoTree_visitM_go___at_Lean_Server_findReferences___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_findModuleRefs___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_replace___at_Lean_Server_dedupReferences___spec__10(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_References_findAt___lambda__1(lean_object*); @@ -295,7 +296,6 @@ LEAN_EXPORT lean_object* l_Lean_Server_ModuleRefs_addRef(lean_object*, lean_obje LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_RefInfo_instCoeRefInfoRefInfo___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_RefInfo_instCoeRefInfoRefInfo(lean_object*); static lean_object* l___private_Lean_Server_References_0__Lean_Server_fromJsonIlean____x40_Lean_Server_References___hyg_1025____closed__18; -lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Server_References_referringTo___spec__5(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMapImp_moveEntries___at_Lean_Server_ModuleRefs_addRef___spec__6(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_References_definitionOf_x3f___lambda__1(lean_object*, lean_object*, lean_object*); @@ -2231,7 +2231,7 @@ x_23 = lean_ctor_get(x_14, 0); lean_inc(x_23); lean_dec(x_14); x_24 = l___private_Lean_Server_References_0__Lean_Server_fromJsonIlean____x40_Lean_Server_References___hyg_1025____closed__19; -x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227____spec__1(x_1, x_24); +x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210____spec__1(x_1, x_24); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; diff --git a/stage0/stdlib/Lean/Server/Rpc/Deriving.c b/stage0/stdlib/Lean/Server/Rpc/Deriving.c index 81c58659ec59..6d6fd2bd1111 100644 --- a/stage0/stdlib/Lean/Server/Rpc/Deriving.c +++ b/stage0/stdlib/Lean/Server/Rpc/Deriving.c @@ -23,6 +23,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_RpcEncodable_deriveInductiveInstance___lambda__1___closed__9; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__93; uint8_t l_String_endsWith(lean_object*, lean_object*); +static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__1; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__2; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__118; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__28; @@ -37,7 +38,6 @@ static lean_object* l_Lean_Server_RpcEncodable_isOptField___closed__1; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__59; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__34; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__21; -static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__13; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__33; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___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___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___closed__2; @@ -51,6 +51,7 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_D static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__14; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___closed__1; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__68; +static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__6; static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__13___lambda__1___closed__3; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___closed__8; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__17; @@ -69,6 +70,7 @@ lean_object* l_Lean_MessageData_ofList(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInstance___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__26; +LEAN_EXPORT lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228_(lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__38; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___closed__5; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__112; @@ -77,11 +79,10 @@ static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deri LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInstance___spec__3(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_RpcEncodable_deriveStructureInstance___spec__1___closed__51; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__19; -static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__6; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__91; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__17; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__53; -static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__11; +static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__9; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInstance___spec__4___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInstance___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*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__30; @@ -94,11 +95,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_RpcEncodable_deriveStructureInstance___lambda__1___closed__6; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__58; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__7(lean_object*, size_t, size_t, lean_object*); -static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__10; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__78; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__27; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__5; -static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__16; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__61; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__32; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__56; @@ -117,11 +116,9 @@ LEAN_EXPORT lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_Rpc static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__42; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__34; static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__6; -static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__7; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__3; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__120; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__14; -static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__1; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__100; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__11; static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__4___closed__1; @@ -132,12 +129,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_RpcEncodable_deriveInductiveInstance___lambda__1___closed__21; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__26; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__27; -static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__3; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__50; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__7; static lean_object* l_Lean_getConstInfoCtor___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__1___closed__3; +static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__7; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__40; -static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__4; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__52; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__111; lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); @@ -160,13 +156,16 @@ static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncod static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__6; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__5(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__15; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInstance___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___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_Array_zip___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__13___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_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__36; lean_object* l_Array_unzip___rarg(lean_object*); +static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__5; lean_object* l_Lean_SourceInfo_fromRef(lean_object*, uint8_t); +static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__4; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__1; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__31; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__25; @@ -191,8 +190,8 @@ static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deri static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__33; LEAN_EXPORT lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_matchAltTerm; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__58; -static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__15; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__12(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_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__14; lean_object* l_Lean_mkCIdentFrom(lean_object*, lean_object*, uint8_t); lean_object* l_Lean_Elab_getBetterRef(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -200,6 +199,7 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_D static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__9; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__2(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_RpcEncodable_deriveStructureInstance___spec__1___closed__44; +static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__8; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__29; lean_object* l_Lean_Elab_Command_liftTermElabM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__8; @@ -223,7 +223,6 @@ static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncod LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__13___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_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__71; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__101; -LEAN_EXPORT lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234_(lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__75; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__2; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__31; @@ -284,6 +283,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_RpcEncodable_deriveInductiveInstance___lambda__1___closed__35; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__5; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__16; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_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_RpcEncodable_deriveStructureInstance___spec__3(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_RpcEncodable_deriveStructureInstance___lambda__1___closed__81; @@ -298,6 +298,7 @@ lean_object* l_Array_append___rarg(lean_object*, lean_object*); lean_object* l_Lean_throwError___at_Lean_Elab_Command_expandDeclId___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__7; lean_object* l_Lean_MessageData_ofExpr(lean_object*); +static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__3; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__113; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInstance___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__19; @@ -311,7 +312,6 @@ static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deri LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__14(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___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_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__1; -static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__9; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__20; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__99; LEAN_EXPORT lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInstance___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -320,12 +320,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_RpcEncodable_deriveStructureInstance___lambda__1___closed__115; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__36; static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInstance___spec__3___closed__3; +static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__2; lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Term_traceAtCmdPos___spec__1(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_object* l_Lean_Meta_getFVarLocalDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__88; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__24; -static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__5; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); lean_object* lean_environment_main_module(lean_object*); LEAN_EXPORT uint8_t l_Lean_Server_RpcEncodable_isOptField(lean_object*); @@ -338,7 +338,6 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_D lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__105; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__95; -static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__8; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___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* l_Lean_Syntax_node7(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_RpcEncodable_deriveStructureInstance___spec__1___closed__57; @@ -353,11 +352,13 @@ static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deri 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___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__63; static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__2; +static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__13; static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__1; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__29; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__4; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__42; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__34; +static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__11; 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*); uint8_t l_Lean_isStructure(lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__41; @@ -372,7 +373,6 @@ LEAN_EXPORT lean_object* l_Lean_getConstInfoInduct___at___private_Lean_Server_Rp static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__13___lambda__1___closed__11; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__106; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__25; -static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__2; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__57; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__46; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__7; @@ -407,13 +407,13 @@ 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_RpcEncodable_deriveStructureInstance___lambda__1___closed__77; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__28; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__114; +static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__10; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__18; uint8_t lean_nat_dec_le(lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__13___lambda__1___closed__2; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__36; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__24; uint8_t lean_usize_dec_lt(size_t, size_t); -static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__14; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__98; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__18; static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInstance___spec__3___closed__6; @@ -424,7 +424,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_RpcEncodable_deriveInductiveInstance___lambda__1___closed__10; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_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_RpcEncodable_deriveInductiveInstance___lambda__1___closed__15; -static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__12; static lean_object* l_Lean_getConstInfoCtor___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__1___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__4(lean_object*, size_t, size_t, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__1; @@ -453,6 +452,7 @@ lean_object* l_List_toArrayAux___rarg(lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__61; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInstance___lambda__3___closed__2; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__72; +static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__12; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__87; lean_object* l_Lean_Parser_termParser(lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__2; @@ -6496,7 +6496,7 @@ lean_dec(x_10); x_58 = !lean_is_exclusive(x_57); if (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; 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; 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; size_t x_199; lean_object* x_200; lean_object* x_201; size_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; 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; size_t x_262; lean_object* x_263; lean_object* x_264; size_t 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_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; 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; size_t x_198; lean_object* x_199; lean_object* x_200; size_t 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; size_t x_261; lean_object* x_262; lean_object* x_263; size_t 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; x_59 = lean_ctor_get(x_57, 0); x_60 = lean_ctor_get(x_59, 0); lean_inc(x_60); @@ -6798,794 +6798,790 @@ lean_ctor_set(x_184, 0, x_35); lean_ctor_set(x_184, 1, x_183); lean_ctor_set(x_184, 2, x_182); lean_ctor_set(x_184, 3, x_38); -lean_inc(x_62); -lean_inc(x_35); -x_185 = l_Lean_Syntax_node2(x_35, x_52, x_184, x_62); lean_inc(x_35); -x_186 = l_Lean_Syntax_node1(x_35, x_52, x_136); -x_187 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__14; +x_185 = l_Lean_Syntax_node1(x_35, x_52, x_136); +x_186 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__14; lean_inc(x_173); lean_inc(x_140); +lean_inc(x_62); lean_inc(x_35); -x_188 = l_Lean_Syntax_node4(x_35, x_187, x_185, x_186, x_140, x_173); -x_189 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__12; +x_187 = l_Lean_Syntax_node5(x_35, x_186, x_184, x_62, x_185, x_140, x_173); +x_188 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__12; lean_inc(x_35); -x_190 = l_Lean_Syntax_node1(x_35, x_189, x_188); -x_191 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__18; +x_189 = l_Lean_Syntax_node1(x_35, x_188, x_187); +x_190 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__18; lean_inc(x_35); -x_192 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_192, 0, x_35); -lean_ctor_set(x_192, 1, x_191); -x_193 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__21; +x_191 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_191, 0, x_35); +lean_ctor_set(x_191, 1, x_190); +x_192 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__21; lean_inc(x_62); lean_inc(x_35); -x_194 = l_Lean_Syntax_node2(x_35, x_193, x_62, x_177); +x_193 = l_Lean_Syntax_node2(x_35, x_192, x_62, x_177); lean_inc(x_35); -x_195 = l_Lean_Syntax_node1(x_35, x_52, x_194); -x_196 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__22; +x_194 = l_Lean_Syntax_node1(x_35, x_52, x_193); +x_195 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__22; lean_inc(x_35); -x_197 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_197, 0, x_35); -lean_ctor_set(x_197, 1, x_196); -x_198 = lean_array_get_size(x_26); -x_199 = lean_usize_of_nat(x_198); -lean_dec(x_198); -x_200 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__16(x_38, x_199, x_18, x_26); -x_201 = lean_array_get_size(x_200); -x_202 = lean_usize_of_nat(x_201); -lean_dec(x_201); -x_203 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__17(x_202, x_18, x_200); -x_204 = l_Array_append___rarg(x_50, x_203); -lean_inc(x_35); -x_205 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_205, 0, x_35); -lean_ctor_set(x_205, 1, x_52); -lean_ctor_set(x_205, 2, x_204); -x_206 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__24; -lean_inc(x_35); -x_207 = l_Lean_Syntax_node1(x_35, x_206, x_205); -x_208 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__19; -lean_inc(x_197); +x_196 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_196, 0, x_35); +lean_ctor_set(x_196, 1, x_195); +x_197 = lean_array_get_size(x_26); +x_198 = lean_usize_of_nat(x_197); +lean_dec(x_197); +x_199 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__16(x_38, x_198, x_18, x_26); +x_200 = lean_array_get_size(x_199); +x_201 = lean_usize_of_nat(x_200); +lean_dec(x_200); +x_202 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__17(x_201, x_18, x_199); +x_203 = l_Array_append___rarg(x_50, x_202); +lean_inc(x_35); +x_204 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_204, 0, x_35); +lean_ctor_set(x_204, 1, x_52); +lean_ctor_set(x_204, 2, x_203); +x_205 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__24; +lean_inc(x_35); +x_206 = l_Lean_Syntax_node1(x_35, x_205, x_204); +x_207 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__19; +lean_inc(x_196); lean_inc_n(x_62, 2); -lean_inc(x_192); +lean_inc(x_191); lean_inc(x_35); -x_209 = l_Lean_Syntax_node6(x_35, x_208, x_192, x_62, x_62, x_195, x_197, x_207); -x_210 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__10; +x_208 = l_Lean_Syntax_node6(x_35, x_207, x_191, x_62, x_62, x_194, x_196, x_206); +x_209 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__10; lean_inc(x_62); -lean_inc(x_190); +lean_inc(x_189); lean_inc(x_180); lean_inc(x_35); -x_211 = l_Lean_Syntax_node4(x_35, x_210, x_180, x_190, x_62, x_209); -x_212 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__86; +x_210 = l_Lean_Syntax_node4(x_35, x_209, x_180, x_189, x_62, x_208); +x_211 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__86; lean_inc(x_140); lean_inc(x_62); lean_inc(x_35); -x_213 = l_Lean_Syntax_node5(x_35, x_212, x_153, x_178, x_62, x_140, x_211); -x_214 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__84; +x_212 = l_Lean_Syntax_node5(x_35, x_211, x_153, x_178, x_62, x_140, x_210); +x_213 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__84; lean_inc(x_35); -x_215 = l_Lean_Syntax_node1(x_35, x_214, x_213); -x_216 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__82; +x_214 = l_Lean_Syntax_node1(x_35, x_213, x_212); +x_215 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__82; lean_inc_n(x_62, 2); lean_inc(x_35); -x_217 = l_Lean_Syntax_node3(x_35, x_216, x_62, x_62, x_215); -x_218 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__99; +x_216 = l_Lean_Syntax_node3(x_35, x_215, x_62, x_62, x_214); +x_217 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__99; lean_inc(x_56); lean_inc(x_61); -x_219 = l_Lean_addMacroScope(x_61, x_218, x_56); -x_220 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__98; +x_218 = l_Lean_addMacroScope(x_61, x_217, x_56); +x_219 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__98; lean_inc(x_35); -x_221 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_221, 0, x_35); -lean_ctor_set(x_221, 1, x_220); -lean_ctor_set(x_221, 2, x_219); -lean_ctor_set(x_221, 3, x_38); +x_220 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_220, 0, x_35); +lean_ctor_set(x_220, 1, x_219); +lean_ctor_set(x_220, 2, x_218); +lean_ctor_set(x_220, 3, x_38); lean_inc(x_35); -x_222 = l_Lean_Syntax_node1(x_35, x_52, x_221); -x_223 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__100; +x_221 = l_Lean_Syntax_node1(x_35, x_52, x_220); +x_222 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__100; lean_inc(x_35); -x_224 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_224, 0, x_35); -lean_ctor_set(x_224, 1, x_223); -x_225 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__26; +x_223 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_223, 0, x_35); +lean_ctor_set(x_223, 1, x_222); +x_224 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__26; lean_inc(x_35); -x_226 = l_Lean_Syntax_node2(x_35, x_225, x_180, x_190); -x_227 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__105; +x_225 = l_Lean_Syntax_node2(x_35, x_224, x_180, x_189); +x_226 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__105; lean_inc(x_62); lean_inc(x_35); -x_228 = l_Lean_Syntax_node2(x_35, x_227, x_226, x_62); -x_229 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__108; +x_227 = l_Lean_Syntax_node2(x_35, x_226, x_225, x_62); +x_228 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__108; lean_inc(x_35); -x_230 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_230, 0, x_35); -lean_ctor_set(x_230, 1, x_229); -x_231 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__29; +x_229 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_229, 0, x_35); +lean_ctor_set(x_229, 1, x_228); +x_230 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__29; lean_inc(x_56); lean_inc(x_61); -x_232 = l_Lean_addMacroScope(x_61, x_231, x_56); -x_233 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__28; +x_231 = l_Lean_addMacroScope(x_61, x_230, x_56); +x_232 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__28; lean_inc(x_35); -x_234 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_234, 0, x_35); -lean_ctor_set(x_234, 1, x_233); -lean_ctor_set(x_234, 2, x_232); -lean_ctor_set(x_234, 3, x_38); +x_233 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_233, 0, x_35); +lean_ctor_set(x_233, 1, x_232); +lean_ctor_set(x_233, 2, x_231); +lean_ctor_set(x_233, 3, x_38); lean_inc(x_35); -x_235 = l_Lean_Syntax_node2(x_35, x_135, x_127, x_70); +x_234 = l_Lean_Syntax_node2(x_35, x_135, x_127, x_70); lean_inc(x_35); -x_236 = l_Lean_Syntax_node1(x_35, x_52, x_235); -x_237 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__22; +x_235 = l_Lean_Syntax_node1(x_35, x_52, x_234); +x_236 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__22; lean_inc(x_35); -x_238 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_238, 0, x_35); -lean_ctor_set(x_238, 1, x_237); -x_239 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__115; +x_237 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_237, 0, x_35); +lean_ctor_set(x_237, 1, x_236); +x_238 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__115; lean_inc(x_56); lean_inc(x_61); -x_240 = l_Lean_addMacroScope(x_61, x_239, x_56); -x_241 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__114; -x_242 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__118; -lean_inc(x_35); -x_243 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_243, 0, x_35); -lean_ctor_set(x_243, 1, x_241); -lean_ctor_set(x_243, 2, x_240); -lean_ctor_set(x_243, 3, x_242); -lean_inc(x_222); -lean_inc(x_35); -x_244 = l_Lean_Syntax_node2(x_35, x_54, x_243, x_222); -x_245 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__112; -lean_inc(x_35); -x_246 = l_Lean_Syntax_node1(x_35, x_245, x_244); -x_247 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__110; -lean_inc(x_234); -lean_inc(x_35); -x_248 = l_Lean_Syntax_node4(x_35, x_247, x_234, x_236, x_238, x_246); -x_249 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__107; +x_239 = l_Lean_addMacroScope(x_61, x_238, x_56); +x_240 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__114; +x_241 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__118; +lean_inc(x_35); +x_242 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_242, 0, x_35); +lean_ctor_set(x_242, 1, x_240); +lean_ctor_set(x_242, 2, x_239); +lean_ctor_set(x_242, 3, x_241); +lean_inc(x_221); +lean_inc(x_35); +x_243 = l_Lean_Syntax_node2(x_35, x_54, x_242, x_221); +x_244 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__112; +lean_inc(x_35); +x_245 = l_Lean_Syntax_node1(x_35, x_244, x_243); +x_246 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__110; +lean_inc(x_233); +lean_inc(x_35); +x_247 = l_Lean_Syntax_node4(x_35, x_246, x_233, x_235, x_237, x_245); +x_248 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__107; lean_inc(x_62); lean_inc(x_35); -x_250 = l_Lean_Syntax_node3(x_35, x_249, x_230, x_62, x_248); +x_249 = l_Lean_Syntax_node3(x_35, x_248, x_229, x_62, x_247); lean_inc(x_62); lean_inc(x_35); -x_251 = l_Lean_Syntax_node2(x_35, x_227, x_250, x_62); -x_252 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__34; -x_253 = l_Lean_addMacroScope(x_61, x_252, x_56); -x_254 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__33; -x_255 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__38; -lean_inc(x_35); -x_256 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_256, 0, x_35); -lean_ctor_set(x_256, 1, x_254); -lean_ctor_set(x_256, 2, x_253); -lean_ctor_set(x_256, 3, x_255); -x_257 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__39; -lean_inc(x_35); -x_258 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_258, 0, x_35); -lean_ctor_set(x_258, 1, x_257); +x_250 = l_Lean_Syntax_node2(x_35, x_226, x_249, x_62); +x_251 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__34; +x_252 = l_Lean_addMacroScope(x_61, x_251, x_56); +x_253 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__33; +x_254 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__38; +lean_inc(x_35); +x_255 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_255, 0, x_35); +lean_ctor_set(x_255, 1, x_253); +lean_ctor_set(x_255, 2, x_252); +lean_ctor_set(x_255, 3, x_254); +x_256 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__39; +lean_inc(x_35); +x_257 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_257, 0, x_35); +lean_ctor_set(x_257, 1, x_256); lean_inc(x_62); lean_inc(x_35); -x_259 = l_Lean_Syntax_node2(x_35, x_193, x_62, x_234); -lean_inc(x_35); -x_260 = l_Lean_Syntax_node1(x_35, x_52, x_259); -x_261 = lean_array_get_size(x_27); -x_262 = lean_usize_of_nat(x_261); -lean_dec(x_261); -x_263 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__18(x_38, x_262, x_18, x_27); -x_264 = lean_array_get_size(x_263); -x_265 = lean_usize_of_nat(x_264); -lean_dec(x_264); -x_266 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__17(x_265, x_18, x_263); -x_267 = l_Array_append___rarg(x_50, x_266); -lean_inc(x_35); -x_268 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_268, 0, x_35); -lean_ctor_set(x_268, 1, x_52); -lean_ctor_set(x_268, 2, x_267); -lean_inc(x_35); -x_269 = l_Lean_Syntax_node1(x_35, x_206, x_268); +x_258 = l_Lean_Syntax_node2(x_35, x_192, x_62, x_233); +lean_inc(x_35); +x_259 = l_Lean_Syntax_node1(x_35, x_52, x_258); +x_260 = lean_array_get_size(x_27); +x_261 = lean_usize_of_nat(x_260); +lean_dec(x_260); +x_262 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__18(x_38, x_261, x_18, x_27); +x_263 = lean_array_get_size(x_262); +x_264 = lean_usize_of_nat(x_263); +lean_dec(x_263); +x_265 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__17(x_264, x_18, x_262); +x_266 = l_Array_append___rarg(x_50, x_265); +lean_inc(x_35); +x_267 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_267, 0, x_35); +lean_ctor_set(x_267, 1, x_52); +lean_ctor_set(x_267, 2, x_266); +lean_inc(x_35); +x_268 = l_Lean_Syntax_node1(x_35, x_205, x_267); lean_inc_n(x_62, 2); lean_inc(x_35); -x_270 = l_Lean_Syntax_node6(x_35, x_208, x_192, x_62, x_62, x_260, x_197, x_269); -x_271 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__31; +x_269 = l_Lean_Syntax_node6(x_35, x_207, x_191, x_62, x_62, x_259, x_196, x_268); +x_270 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__31; lean_inc(x_35); -x_272 = l_Lean_Syntax_node3(x_35, x_271, x_256, x_258, x_270); +x_271 = l_Lean_Syntax_node3(x_35, x_270, x_255, x_257, x_269); lean_inc(x_35); -x_273 = l_Lean_Syntax_node1(x_35, x_245, x_272); +x_272 = l_Lean_Syntax_node1(x_35, x_244, x_271); lean_inc(x_62); lean_inc(x_35); -x_274 = l_Lean_Syntax_node2(x_35, x_227, x_273, x_62); +x_273 = l_Lean_Syntax_node2(x_35, x_226, x_272, x_62); lean_inc(x_35); -x_275 = l_Lean_Syntax_node3(x_35, x_52, x_228, x_251, x_274); -x_276 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__103; +x_274 = l_Lean_Syntax_node3(x_35, x_52, x_227, x_250, x_273); +x_275 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__103; lean_inc(x_35); -x_277 = l_Lean_Syntax_node1(x_35, x_276, x_275); -x_278 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__101; +x_276 = l_Lean_Syntax_node1(x_35, x_275, x_274); +x_277 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__101; lean_inc(x_35); -x_279 = l_Lean_Syntax_node2(x_35, x_278, x_224, x_277); +x_278 = l_Lean_Syntax_node2(x_35, x_277, x_223, x_276); lean_inc(x_140); lean_inc(x_62); lean_inc(x_35); -x_280 = l_Lean_Syntax_node5(x_35, x_212, x_165, x_222, x_62, x_140, x_279); +x_279 = l_Lean_Syntax_node5(x_35, x_211, x_165, x_221, x_62, x_140, x_278); lean_inc(x_35); -x_281 = l_Lean_Syntax_node1(x_35, x_214, x_280); +x_280 = l_Lean_Syntax_node1(x_35, x_213, x_279); lean_inc_n(x_62, 2); lean_inc(x_35); -x_282 = l_Lean_Syntax_node3(x_35, x_216, x_62, x_62, x_281); +x_281 = l_Lean_Syntax_node3(x_35, x_215, x_62, x_62, x_280); lean_inc(x_62); lean_inc(x_35); -x_283 = l_Lean_Syntax_node3(x_35, x_52, x_217, x_62, x_282); -x_284 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__80; +x_282 = l_Lean_Syntax_node3(x_35, x_52, x_216, x_62, x_281); +x_283 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__80; lean_inc(x_35); -x_285 = l_Lean_Syntax_node2(x_35, x_284, x_76, x_283); +x_284 = l_Lean_Syntax_node2(x_35, x_283, x_76, x_282); lean_inc(x_35); -x_286 = l_Lean_Syntax_node1(x_35, x_52, x_285); -x_287 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__66; +x_285 = l_Lean_Syntax_node1(x_35, x_52, x_284); +x_286 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__66; lean_inc(x_35); -x_288 = l_Lean_Syntax_node3(x_35, x_287, x_140, x_173, x_286); -x_289 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__47; +x_287 = l_Lean_Syntax_node3(x_35, x_286, x_140, x_173, x_285); +x_288 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__47; lean_inc_n(x_62, 3); lean_inc(x_35); -x_290 = l_Lean_Syntax_node8(x_35, x_289, x_123, x_125, x_62, x_62, x_138, x_288, x_62, x_62); +x_289 = l_Lean_Syntax_node8(x_35, x_288, x_123, x_125, x_62, x_62, x_138, x_287, x_62, x_62); lean_inc(x_35); -x_291 = l_Lean_Syntax_node2(x_35, x_106, x_121, x_290); -x_292 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__43; +x_290 = l_Lean_Syntax_node2(x_35, x_106, x_121, x_289); +x_291 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__43; lean_inc(x_35); -x_293 = l_Lean_Syntax_node3(x_35, x_292, x_113, x_115, x_291); -x_294 = l_Lean_Syntax_node2(x_35, x_52, x_107, x_293); -lean_ctor_set(x_57, 0, x_294); +x_292 = l_Lean_Syntax_node3(x_35, x_291, x_113, x_115, x_290); +x_293 = l_Lean_Syntax_node2(x_35, x_52, x_107, x_292); +lean_ctor_set(x_57, 0, x_293); return x_57; } 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; size_t 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; size_t x_436; lean_object* x_437; lean_object* x_438; size_t 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; size_t x_499; lean_object* x_500; lean_object* x_501; size_t 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; -x_295 = lean_ctor_get(x_57, 0); -x_296 = lean_ctor_get(x_57, 1); -lean_inc(x_296); +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; size_t 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; size_t x_434; lean_object* x_435; lean_object* x_436; size_t 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; size_t x_497; lean_object* x_498; lean_object* x_499; size_t 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; +x_294 = lean_ctor_get(x_57, 0); +x_295 = lean_ctor_get(x_57, 1); lean_inc(x_295); +lean_inc(x_294); lean_dec(x_57); -x_297 = lean_ctor_get(x_295, 0); -lean_inc(x_297); -lean_dec(x_295); -x_298 = lean_environment_main_module(x_297); -lean_inc(x_35); -x_299 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_299, 0, x_35); -lean_ctor_set(x_299, 1, x_52); -lean_ctor_set(x_299, 2, x_50); -x_300 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__7; -lean_inc_n(x_299, 6); -lean_inc(x_35); -x_301 = l_Lean_Syntax_node6(x_35, x_300, x_299, x_299, x_299, x_299, x_299, x_299); -x_302 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__2; -lean_inc(x_35); -x_303 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_303, 0, x_35); -lean_ctor_set(x_303, 1, x_302); -x_304 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__16; +x_296 = lean_ctor_get(x_294, 0); +lean_inc(x_296); +lean_dec(x_294); +x_297 = lean_environment_main_module(x_296); +lean_inc(x_35); +x_298 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_298, 0, x_35); +lean_ctor_set(x_298, 1, x_52); +lean_ctor_set(x_298, 2, x_50); +x_299 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__7; +lean_inc_n(x_298, 6); +lean_inc(x_35); +x_300 = l_Lean_Syntax_node6(x_35, x_299, x_298, x_298, x_298, x_298, x_298, x_298); +x_301 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__2; +lean_inc(x_35); +x_302 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_302, 0, x_35); +lean_ctor_set(x_302, 1, x_301); +x_303 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__16; lean_inc(x_56); +lean_inc(x_297); +x_304 = l_Lean_addMacroScope(x_297, x_303, x_56); +x_305 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__15; +lean_inc(x_35); +x_306 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_306, 0, x_35); +lean_ctor_set(x_306, 1, x_305); +lean_ctor_set(x_306, 2, x_304); +lean_ctor_set(x_306, 3, x_38); +x_307 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__13; lean_inc(x_298); -x_305 = l_Lean_addMacroScope(x_298, x_304, x_56); -x_306 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__15; -lean_inc(x_35); -x_307 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_307, 0, x_35); -lean_ctor_set(x_307, 1, x_306); -lean_ctor_set(x_307, 2, x_305); -lean_ctor_set(x_307, 3, x_38); -x_308 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__13; -lean_inc(x_299); -lean_inc(x_307); -lean_inc(x_35); -x_309 = l_Lean_Syntax_node2(x_35, x_308, x_307, x_299); -x_310 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__5; -lean_inc_n(x_299, 2); -lean_inc(x_35); -x_311 = l_Lean_Syntax_node2(x_35, x_310, x_299, x_299); -x_312 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__17; -lean_inc(x_35); -x_313 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_313, 0, x_35); -lean_ctor_set(x_313, 1, x_312); -lean_inc(x_313); -lean_inc(x_35); -x_314 = l_Lean_Syntax_node1(x_35, x_52, x_313); -x_315 = lean_array_get_size(x_23); -x_316 = lean_usize_of_nat(x_315); -lean_dec(x_315); -x_317 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__15(x_316, x_18, x_23); -x_318 = l_Array_append___rarg(x_50, x_317); -lean_inc(x_35); -x_319 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_319, 0, x_35); -lean_ctor_set(x_319, 1, x_52); -lean_ctor_set(x_319, 2, x_318); -x_320 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__22; -lean_inc(x_35); -x_321 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_321, 0, x_35); -lean_ctor_set(x_321, 1, x_320); -x_322 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__27; +lean_inc(x_306); +lean_inc(x_35); +x_308 = l_Lean_Syntax_node2(x_35, x_307, x_306, x_298); +x_309 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__5; +lean_inc_n(x_298, 2); +lean_inc(x_35); +x_310 = l_Lean_Syntax_node2(x_35, x_309, x_298, x_298); +x_311 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__17; +lean_inc(x_35); +x_312 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_312, 0, x_35); +lean_ctor_set(x_312, 1, x_311); +lean_inc(x_312); +lean_inc(x_35); +x_313 = l_Lean_Syntax_node1(x_35, x_52, x_312); +x_314 = lean_array_get_size(x_23); +x_315 = lean_usize_of_nat(x_314); +lean_dec(x_314); +x_316 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__15(x_315, x_18, x_23); +x_317 = l_Array_append___rarg(x_50, x_316); +lean_inc(x_35); +x_318 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_318, 0, x_35); +lean_ctor_set(x_318, 1, x_52); +lean_ctor_set(x_318, 2, x_317); +x_319 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__22; +lean_inc(x_35); +x_320 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_320, 0, x_35); +lean_ctor_set(x_320, 1, x_319); +x_321 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__27; lean_inc(x_56); +lean_inc(x_297); +x_322 = l_Lean_addMacroScope(x_297, x_321, x_56); +x_323 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__26; +x_324 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__32; +lean_inc(x_35); +x_325 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_325, 0, x_35); +lean_ctor_set(x_325, 1, x_323); +lean_ctor_set(x_325, 2, x_322); +lean_ctor_set(x_325, 3, x_324); +x_326 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__24; lean_inc(x_298); -x_323 = l_Lean_addMacroScope(x_298, x_322, x_56); -x_324 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__26; -x_325 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__32; -lean_inc(x_35); -x_326 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_326, 0, x_35); -lean_ctor_set(x_326, 1, x_324); -lean_ctor_set(x_326, 2, x_323); -lean_ctor_set(x_326, 3, x_325); -x_327 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__24; -lean_inc(x_299); -lean_inc(x_35); -x_328 = l_Lean_Syntax_node2(x_35, x_327, x_326, x_299); -x_329 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__33; -lean_inc(x_35); -x_330 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_330, 0, x_35); -lean_ctor_set(x_330, 1, x_329); -x_331 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__36; +lean_inc(x_35); +x_327 = l_Lean_Syntax_node2(x_35, x_326, x_325, x_298); +x_328 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__33; +lean_inc(x_35); +x_329 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_329, 0, x_35); +lean_ctor_set(x_329, 1, x_328); +x_330 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__36; lean_inc(x_56); +lean_inc(x_297); +x_331 = l_Lean_addMacroScope(x_297, x_330, x_56); +x_332 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__35; +x_333 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__41; +lean_inc(x_35); +x_334 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_334, 0, x_35); +lean_ctor_set(x_334, 1, x_332); +lean_ctor_set(x_334, 2, x_331); +lean_ctor_set(x_334, 3, x_333); lean_inc(x_298); -x_332 = l_Lean_addMacroScope(x_298, x_331, x_56); -x_333 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__35; -x_334 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__41; lean_inc(x_35); -x_335 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_335, 0, x_35); -lean_ctor_set(x_335, 1, x_333); -lean_ctor_set(x_335, 2, x_332); -lean_ctor_set(x_335, 3, x_334); -lean_inc(x_299); +x_335 = l_Lean_Syntax_node2(x_35, x_326, x_334, x_298); +lean_inc(x_329); lean_inc(x_35); -x_336 = l_Lean_Syntax_node2(x_35, x_327, x_335, x_299); -lean_inc(x_330); +x_336 = l_Lean_Syntax_node3(x_35, x_52, x_327, x_329, x_335); lean_inc(x_35); -x_337 = l_Lean_Syntax_node3(x_35, x_52, x_328, x_330, x_336); +x_337 = l_Lean_Syntax_node2(x_35, x_52, x_320, x_336); +x_338 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__21; lean_inc(x_35); -x_338 = l_Lean_Syntax_node2(x_35, x_52, x_321, x_337); -x_339 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__21; +x_339 = l_Lean_Syntax_node1(x_35, x_338, x_337); +x_340 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__3; +lean_inc(x_298); lean_inc(x_35); -x_340 = l_Lean_Syntax_node1(x_35, x_339, x_338); -x_341 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__3; -lean_inc(x_299); +x_341 = l_Lean_Syntax_node7(x_35, x_340, x_302, x_308, x_310, x_313, x_318, x_298, x_339); +x_342 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__5; lean_inc(x_35); -x_342 = l_Lean_Syntax_node7(x_35, x_341, x_303, x_309, x_311, x_314, x_319, x_299, x_340); -x_343 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__5; +x_343 = l_Lean_Syntax_node2(x_35, x_342, x_300, x_341); +x_344 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__44; lean_inc(x_35); -x_344 = l_Lean_Syntax_node2(x_35, x_343, x_301, x_342); -x_345 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__44; +x_345 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_345, 0, x_35); +lean_ctor_set(x_345, 1, x_344); +x_346 = l_Array_append___rarg(x_50, x_3); lean_inc(x_35); -x_346 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_346, 0, x_35); -lean_ctor_set(x_346, 1, x_345); -x_347 = l_Array_append___rarg(x_50, x_3); -lean_inc(x_35); -x_348 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_348, 0, x_35); -lean_ctor_set(x_348, 1, x_52); -lean_ctor_set(x_348, 2, x_347); -x_349 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__45; +x_347 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_347, 0, x_35); +lean_ctor_set(x_347, 1, x_52); +lean_ctor_set(x_347, 2, x_346); +x_348 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__45; lean_inc(x_35); -x_350 = l_Lean_Syntax_node2(x_35, x_349, x_346, x_348); -x_351 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__42; +x_349 = l_Lean_Syntax_node2(x_35, x_348, x_345, x_347); +x_350 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__42; lean_inc(x_35); -x_352 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_352, 0, x_35); -lean_ctor_set(x_352, 1, x_351); -x_353 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__4; +x_351 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_351, 0, x_35); +lean_ctor_set(x_351, 1, x_350); +x_352 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__4; lean_inc(x_35); -x_354 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_354, 0, x_35); -lean_ctor_set(x_354, 1, x_353); -x_355 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__5; +x_353 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_353, 0, x_35); +lean_ctor_set(x_353, 1, x_352); +x_354 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__5; lean_inc(x_35); -x_356 = l_Lean_Syntax_node1(x_35, x_355, x_354); +x_355 = l_Lean_Syntax_node1(x_35, x_354, x_353); lean_inc(x_35); -x_357 = l_Lean_Syntax_node1(x_35, x_52, x_356); -lean_inc_n(x_299, 5); +x_356 = l_Lean_Syntax_node1(x_35, x_52, x_355); +lean_inc_n(x_298, 5); lean_inc(x_35); -x_358 = l_Lean_Syntax_node6(x_35, x_300, x_299, x_299, x_299, x_299, x_299, x_357); -x_359 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__49; -lean_inc(x_299); +x_357 = l_Lean_Syntax_node6(x_35, x_299, x_298, x_298, x_298, x_298, x_298, x_356); +x_358 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__49; +lean_inc(x_298); lean_inc(x_35); -x_360 = l_Lean_Syntax_node1(x_35, x_359, x_299); -x_361 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__46; +x_359 = l_Lean_Syntax_node1(x_35, x_358, x_298); +x_360 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__46; lean_inc(x_35); -x_362 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_362, 0, x_35); -lean_ctor_set(x_362, 1, x_361); -x_363 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__8; +x_361 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_361, 0, x_35); +lean_ctor_set(x_361, 1, x_360); +x_362 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__8; lean_inc(x_35); -x_364 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_364, 0, x_35); -lean_ctor_set(x_364, 1, x_363); -x_365 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__53; -lean_inc(x_56); -lean_inc(x_298); -x_366 = l_Lean_addMacroScope(x_298, x_365, x_56); -x_367 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__52; -x_368 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__58; -lean_inc(x_35); -x_369 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_369, 0, x_35); -lean_ctor_set(x_369, 1, x_367); -lean_ctor_set(x_369, 2, x_366); -lean_ctor_set(x_369, 3, x_368); -lean_inc(x_35); -x_370 = l_Lean_Syntax_node1(x_35, x_52, x_55); -lean_inc(x_35); -x_371 = l_Lean_Syntax_node2(x_35, x_54, x_369, x_370); -x_372 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__7; -lean_inc(x_364); -lean_inc(x_35); -x_373 = l_Lean_Syntax_node2(x_35, x_372, x_364, x_371); -x_374 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__51; -lean_inc(x_373); -lean_inc(x_299); -lean_inc(x_35); -x_375 = l_Lean_Syntax_node2(x_35, x_374, x_299, x_373); -x_376 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__19; -lean_inc(x_35); -x_377 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_377, 0, x_35); -lean_ctor_set(x_377, 1, x_376); -x_378 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__69; -lean_inc(x_35); -x_379 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_379, 0, x_35); -lean_ctor_set(x_379, 1, x_378); -x_380 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__27; +x_363 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_363, 0, x_35); +lean_ctor_set(x_363, 1, x_362); +x_364 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__53; lean_inc(x_56); +lean_inc(x_297); +x_365 = l_Lean_addMacroScope(x_297, x_364, x_56); +x_366 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__52; +x_367 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__58; +lean_inc(x_35); +x_368 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_368, 0, x_35); +lean_ctor_set(x_368, 1, x_366); +lean_ctor_set(x_368, 2, x_365); +lean_ctor_set(x_368, 3, x_367); +lean_inc(x_35); +x_369 = l_Lean_Syntax_node1(x_35, x_52, x_55); +lean_inc(x_35); +x_370 = l_Lean_Syntax_node2(x_35, x_54, x_368, x_369); +x_371 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__7; +lean_inc(x_363); +lean_inc(x_35); +x_372 = l_Lean_Syntax_node2(x_35, x_371, x_363, x_370); +x_373 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__51; +lean_inc(x_372); lean_inc(x_298); -x_381 = l_Lean_addMacroScope(x_298, x_380, x_56); -x_382 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__26; -x_383 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__32; -lean_inc(x_35); -x_384 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_384, 0, x_35); -lean_ctor_set(x_384, 1, x_382); -lean_ctor_set(x_384, 2, x_381); -lean_ctor_set(x_384, 3, x_383); -x_385 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__15; -lean_inc(x_299); -lean_inc(x_35); -x_386 = l_Lean_Syntax_node2(x_35, x_385, x_384, x_299); -x_387 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__72; +lean_inc(x_35); +x_374 = l_Lean_Syntax_node2(x_35, x_373, x_298, x_372); +x_375 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__19; +lean_inc(x_35); +x_376 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_376, 0, x_35); +lean_ctor_set(x_376, 1, x_375); +x_377 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__69; +lean_inc(x_35); +x_378 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_378, 0, x_35); +lean_ctor_set(x_378, 1, x_377); +x_379 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__27; lean_inc(x_56); +lean_inc(x_297); +x_380 = l_Lean_addMacroScope(x_297, x_379, x_56); +x_381 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__26; +x_382 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__32; +lean_inc(x_35); +x_383 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_383, 0, x_35); +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 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__15; lean_inc(x_298); -x_388 = l_Lean_addMacroScope(x_298, x_387, x_56); -x_389 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__71; lean_inc(x_35); -x_390 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_390, 0, x_35); -lean_ctor_set(x_390, 1, x_389); -lean_ctor_set(x_390, 2, x_388); -lean_ctor_set(x_390, 3, x_38); -x_391 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__13; -lean_inc(x_390); -lean_inc(x_377); -lean_inc(x_35); -x_392 = l_Lean_Syntax_node3(x_35, x_391, x_386, x_377, x_390); -x_393 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__41; +x_385 = l_Lean_Syntax_node2(x_35, x_384, x_383, x_298); +x_386 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__72; lean_inc(x_56); -lean_inc(x_298); -x_394 = l_Lean_addMacroScope(x_298, x_393, x_56); -x_395 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__40; -x_396 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__44; -lean_inc(x_35); -x_397 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_397, 0, x_35); -lean_ctor_set(x_397, 1, x_395); -lean_ctor_set(x_397, 2, x_394); -lean_ctor_set(x_397, 3, x_396); -lean_inc(x_299); -lean_inc(x_35); -x_398 = l_Lean_Syntax_node2(x_35, x_385, x_397, x_299); -x_399 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__75; +lean_inc(x_297); +x_387 = l_Lean_addMacroScope(x_297, x_386, x_56); +x_388 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__71; +lean_inc(x_35); +x_389 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_389, 0, x_35); +lean_ctor_set(x_389, 1, x_388); +lean_ctor_set(x_389, 2, x_387); +lean_ctor_set(x_389, 3, x_38); +x_390 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__13; +lean_inc(x_389); +lean_inc(x_376); +lean_inc(x_35); +x_391 = l_Lean_Syntax_node3(x_35, x_390, x_385, x_376, x_389); +x_392 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__41; lean_inc(x_56); +lean_inc(x_297); +x_393 = l_Lean_addMacroScope(x_297, x_392, x_56); +x_394 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__40; +x_395 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__44; +lean_inc(x_35); +x_396 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_396, 0, x_35); +lean_ctor_set(x_396, 1, x_394); +lean_ctor_set(x_396, 2, x_393); +lean_ctor_set(x_396, 3, x_395); lean_inc(x_298); -x_400 = l_Lean_addMacroScope(x_298, x_399, x_56); -x_401 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__74; -lean_inc(x_35); -x_402 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_402, 0, x_35); -lean_ctor_set(x_402, 1, x_401); -lean_ctor_set(x_402, 2, x_400); -lean_ctor_set(x_402, 3, x_38); -lean_inc(x_402); -lean_inc(x_377); -lean_inc(x_35); -x_403 = l_Lean_Syntax_node3(x_35, x_391, x_398, x_377, x_402); -lean_inc(x_35); -x_404 = l_Lean_Syntax_node3(x_35, x_52, x_392, x_330, x_403); -x_405 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__77; -lean_inc(x_299); -lean_inc(x_35); -x_406 = l_Lean_Syntax_node1(x_35, x_405, x_299); -x_407 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__78; -lean_inc(x_35); -x_408 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_408, 0, x_35); -lean_ctor_set(x_408, 1, x_407); -x_409 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__68; -lean_inc_n(x_299, 2); -lean_inc(x_35); -x_410 = l_Lean_Syntax_node6(x_35, x_409, x_379, x_299, x_404, x_406, x_299, x_408); -x_411 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__8; +lean_inc(x_35); +x_397 = l_Lean_Syntax_node2(x_35, x_384, x_396, x_298); +x_398 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__75; lean_inc(x_56); +lean_inc(x_297); +x_399 = l_Lean_addMacroScope(x_297, x_398, x_56); +x_400 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__74; +lean_inc(x_35); +x_401 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_401, 0, x_35); +lean_ctor_set(x_401, 1, x_400); +lean_ctor_set(x_401, 2, x_399); +lean_ctor_set(x_401, 3, x_38); +lean_inc(x_401); +lean_inc(x_376); +lean_inc(x_35); +x_402 = l_Lean_Syntax_node3(x_35, x_390, x_397, x_376, x_401); +lean_inc(x_35); +x_403 = l_Lean_Syntax_node3(x_35, x_52, x_391, x_329, x_402); +x_404 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__77; lean_inc(x_298); -x_412 = l_Lean_addMacroScope(x_298, x_411, x_56); -x_413 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__7; -lean_inc(x_35); -x_414 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_414, 0, x_35); -lean_ctor_set(x_414, 1, x_413); -lean_ctor_set(x_414, 2, x_412); -lean_ctor_set(x_414, 3, x_38); -lean_inc(x_414); -lean_inc(x_35); -x_415 = l_Lean_Syntax_node1(x_35, x_52, x_414); -x_416 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__9; -lean_inc(x_35); -x_417 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_417, 0, x_35); -lean_ctor_set(x_417, 1, x_416); -x_418 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__17; +lean_inc(x_35); +x_405 = l_Lean_Syntax_node1(x_35, x_404, x_298); +x_406 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__78; +lean_inc(x_35); +x_407 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_407, 0, x_35); +lean_ctor_set(x_407, 1, x_406); +x_408 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__68; +lean_inc_n(x_298, 2); +lean_inc(x_35); +x_409 = l_Lean_Syntax_node6(x_35, x_408, x_378, x_298, x_403, x_405, x_298, x_407); +x_410 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__8; lean_inc(x_56); +lean_inc(x_297); +x_411 = l_Lean_addMacroScope(x_297, x_410, x_56); +x_412 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__7; +lean_inc(x_35); +x_413 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_413, 0, x_35); +lean_ctor_set(x_413, 1, x_412); +lean_ctor_set(x_413, 2, x_411); +lean_ctor_set(x_413, 3, x_38); +lean_inc(x_413); +lean_inc(x_35); +x_414 = l_Lean_Syntax_node1(x_35, x_52, x_413); +x_415 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__9; +lean_inc(x_35); +x_416 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_416, 0, x_35); +lean_ctor_set(x_416, 1, x_415); +x_417 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__17; +lean_inc(x_56); +lean_inc(x_297); +x_418 = l_Lean_addMacroScope(x_297, x_417, x_56); +x_419 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__16; +lean_inc(x_35); +x_420 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_420, 0, x_35); +lean_ctor_set(x_420, 1, x_419); +lean_ctor_set(x_420, 2, x_418); +lean_ctor_set(x_420, 3, x_38); +lean_inc(x_35); +x_421 = l_Lean_Syntax_node1(x_35, x_52, x_372); +x_422 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__14; +lean_inc(x_409); +lean_inc(x_376); lean_inc(x_298); -x_419 = l_Lean_addMacroScope(x_298, x_418, x_56); -x_420 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__16; lean_inc(x_35); -x_421 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_421, 0, x_35); -lean_ctor_set(x_421, 1, x_420); -lean_ctor_set(x_421, 2, x_419); -lean_ctor_set(x_421, 3, x_38); -lean_inc(x_299); +x_423 = l_Lean_Syntax_node5(x_35, x_422, x_420, x_298, x_421, x_376, x_409); +x_424 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__12; lean_inc(x_35); -x_422 = l_Lean_Syntax_node2(x_35, x_52, x_421, x_299); +x_425 = l_Lean_Syntax_node1(x_35, x_424, x_423); +x_426 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__18; lean_inc(x_35); -x_423 = l_Lean_Syntax_node1(x_35, x_52, x_373); -x_424 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__14; -lean_inc(x_410); -lean_inc(x_377); -lean_inc(x_35); -x_425 = l_Lean_Syntax_node4(x_35, x_424, x_422, x_423, x_377, x_410); -x_426 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__12; -lean_inc(x_35); -x_427 = l_Lean_Syntax_node1(x_35, x_426, x_425); -x_428 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__18; -lean_inc(x_35); -x_429 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_429, 0, x_35); -lean_ctor_set(x_429, 1, x_428); -x_430 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__21; -lean_inc(x_299); -lean_inc(x_35); -x_431 = l_Lean_Syntax_node2(x_35, x_430, x_299, x_414); -lean_inc(x_35); -x_432 = l_Lean_Syntax_node1(x_35, x_52, x_431); -x_433 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__22; -lean_inc(x_35); -x_434 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_434, 0, x_35); -lean_ctor_set(x_434, 1, x_433); -x_435 = lean_array_get_size(x_26); -x_436 = lean_usize_of_nat(x_435); -lean_dec(x_435); -x_437 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__16(x_38, x_436, x_18, x_26); -x_438 = lean_array_get_size(x_437); -x_439 = lean_usize_of_nat(x_438); -lean_dec(x_438); -x_440 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__17(x_439, x_18, x_437); -x_441 = l_Array_append___rarg(x_50, x_440); -lean_inc(x_35); -x_442 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_442, 0, x_35); -lean_ctor_set(x_442, 1, x_52); -lean_ctor_set(x_442, 2, x_441); -x_443 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__24; -lean_inc(x_35); -x_444 = l_Lean_Syntax_node1(x_35, x_443, x_442); -x_445 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__19; -lean_inc(x_434); -lean_inc_n(x_299, 2); -lean_inc(x_429); -lean_inc(x_35); -x_446 = l_Lean_Syntax_node6(x_35, x_445, x_429, x_299, x_299, x_432, x_434, x_444); -x_447 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__10; -lean_inc(x_299); +x_427 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_427, 0, x_35); +lean_ctor_set(x_427, 1, x_426); +x_428 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__21; +lean_inc(x_298); +lean_inc(x_35); +x_429 = l_Lean_Syntax_node2(x_35, x_428, x_298, x_413); +lean_inc(x_35); +x_430 = l_Lean_Syntax_node1(x_35, x_52, x_429); +x_431 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__22; +lean_inc(x_35); +x_432 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_432, 0, x_35); +lean_ctor_set(x_432, 1, x_431); +x_433 = lean_array_get_size(x_26); +x_434 = lean_usize_of_nat(x_433); +lean_dec(x_433); +x_435 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__16(x_38, x_434, x_18, x_26); +x_436 = lean_array_get_size(x_435); +x_437 = lean_usize_of_nat(x_436); +lean_dec(x_436); +x_438 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__17(x_437, x_18, x_435); +x_439 = l_Array_append___rarg(x_50, x_438); +lean_inc(x_35); +x_440 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_440, 0, x_35); +lean_ctor_set(x_440, 1, x_52); +lean_ctor_set(x_440, 2, x_439); +x_441 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__24; +lean_inc(x_35); +x_442 = l_Lean_Syntax_node1(x_35, x_441, x_440); +x_443 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__19; +lean_inc(x_432); +lean_inc_n(x_298, 2); lean_inc(x_427); -lean_inc(x_417); lean_inc(x_35); -x_448 = l_Lean_Syntax_node4(x_35, x_447, x_417, x_427, x_299, x_446); -x_449 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__86; -lean_inc(x_377); -lean_inc(x_299); +x_444 = l_Lean_Syntax_node6(x_35, x_443, x_427, x_298, x_298, x_430, x_432, x_442); +x_445 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__10; +lean_inc(x_298); +lean_inc(x_425); +lean_inc(x_416); +lean_inc(x_35); +x_446 = l_Lean_Syntax_node4(x_35, x_445, x_416, x_425, x_298, x_444); +x_447 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__86; +lean_inc(x_376); +lean_inc(x_298); lean_inc(x_35); -x_450 = l_Lean_Syntax_node5(x_35, x_449, x_390, x_415, x_299, x_377, x_448); -x_451 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__84; +x_448 = l_Lean_Syntax_node5(x_35, x_447, x_389, x_414, x_298, x_376, x_446); +x_449 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__84; lean_inc(x_35); -x_452 = l_Lean_Syntax_node1(x_35, x_451, x_450); -x_453 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__82; -lean_inc_n(x_299, 2); +x_450 = l_Lean_Syntax_node1(x_35, x_449, x_448); +x_451 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__82; +lean_inc_n(x_298, 2); lean_inc(x_35); -x_454 = l_Lean_Syntax_node3(x_35, x_453, x_299, x_299, x_452); -x_455 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__99; +x_452 = l_Lean_Syntax_node3(x_35, x_451, x_298, x_298, x_450); +x_453 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__99; lean_inc(x_56); -lean_inc(x_298); -x_456 = l_Lean_addMacroScope(x_298, x_455, x_56); -x_457 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__98; +lean_inc(x_297); +x_454 = l_Lean_addMacroScope(x_297, x_453, x_56); +x_455 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__98; lean_inc(x_35); -x_458 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_458, 0, x_35); -lean_ctor_set(x_458, 1, x_457); -lean_ctor_set(x_458, 2, x_456); -lean_ctor_set(x_458, 3, x_38); +x_456 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_456, 0, x_35); +lean_ctor_set(x_456, 1, x_455); +lean_ctor_set(x_456, 2, x_454); +lean_ctor_set(x_456, 3, x_38); lean_inc(x_35); -x_459 = l_Lean_Syntax_node1(x_35, x_52, x_458); -x_460 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__100; +x_457 = l_Lean_Syntax_node1(x_35, x_52, x_456); +x_458 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__100; lean_inc(x_35); -x_461 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_461, 0, x_35); -lean_ctor_set(x_461, 1, x_460); -x_462 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__26; +x_459 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_459, 0, x_35); +lean_ctor_set(x_459, 1, x_458); +x_460 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__26; lean_inc(x_35); -x_463 = l_Lean_Syntax_node2(x_35, x_462, x_417, x_427); -x_464 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__105; -lean_inc(x_299); +x_461 = l_Lean_Syntax_node2(x_35, x_460, x_416, x_425); +x_462 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__105; +lean_inc(x_298); lean_inc(x_35); -x_465 = l_Lean_Syntax_node2(x_35, x_464, x_463, x_299); -x_466 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__108; +x_463 = l_Lean_Syntax_node2(x_35, x_462, x_461, x_298); +x_464 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__108; lean_inc(x_35); -x_467 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_467, 0, x_35); -lean_ctor_set(x_467, 1, x_466); -x_468 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__29; +x_465 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_465, 0, x_35); +lean_ctor_set(x_465, 1, x_464); +x_466 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__29; lean_inc(x_56); -lean_inc(x_298); -x_469 = l_Lean_addMacroScope(x_298, x_468, x_56); -x_470 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__28; +lean_inc(x_297); +x_467 = l_Lean_addMacroScope(x_297, x_466, x_56); +x_468 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__28; lean_inc(x_35); -x_471 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_471, 0, x_35); -lean_ctor_set(x_471, 1, x_470); -lean_ctor_set(x_471, 2, x_469); -lean_ctor_set(x_471, 3, x_38); +x_469 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_469, 0, x_35); +lean_ctor_set(x_469, 1, x_468); +lean_ctor_set(x_469, 2, x_467); +lean_ctor_set(x_469, 3, x_38); lean_inc(x_35); -x_472 = l_Lean_Syntax_node2(x_35, x_372, x_364, x_307); +x_470 = l_Lean_Syntax_node2(x_35, x_371, x_363, x_306); lean_inc(x_35); -x_473 = l_Lean_Syntax_node1(x_35, x_52, x_472); -x_474 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__22; +x_471 = l_Lean_Syntax_node1(x_35, x_52, x_470); +x_472 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__22; lean_inc(x_35); -x_475 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_475, 0, x_35); -lean_ctor_set(x_475, 1, x_474); -x_476 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__115; +x_473 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_473, 0, x_35); +lean_ctor_set(x_473, 1, x_472); +x_474 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__115; lean_inc(x_56); +lean_inc(x_297); +x_475 = l_Lean_addMacroScope(x_297, x_474, x_56); +x_476 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__114; +x_477 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__118; +lean_inc(x_35); +x_478 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_478, 0, x_35); +lean_ctor_set(x_478, 1, x_476); +lean_ctor_set(x_478, 2, x_475); +lean_ctor_set(x_478, 3, x_477); +lean_inc(x_457); +lean_inc(x_35); +x_479 = l_Lean_Syntax_node2(x_35, x_54, x_478, x_457); +x_480 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__112; +lean_inc(x_35); +x_481 = l_Lean_Syntax_node1(x_35, x_480, x_479); +x_482 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__110; +lean_inc(x_469); +lean_inc(x_35); +x_483 = l_Lean_Syntax_node4(x_35, x_482, x_469, x_471, x_473, x_481); +x_484 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__107; +lean_inc(x_298); +lean_inc(x_35); +x_485 = l_Lean_Syntax_node3(x_35, x_484, x_465, x_298, x_483); lean_inc(x_298); -x_477 = l_Lean_addMacroScope(x_298, x_476, x_56); -x_478 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__114; -x_479 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__118; -lean_inc(x_35); -x_480 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_480, 0, x_35); -lean_ctor_set(x_480, 1, x_478); -lean_ctor_set(x_480, 2, x_477); -lean_ctor_set(x_480, 3, x_479); -lean_inc(x_459); -lean_inc(x_35); -x_481 = l_Lean_Syntax_node2(x_35, x_54, x_480, x_459); -x_482 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__112; -lean_inc(x_35); -x_483 = l_Lean_Syntax_node1(x_35, x_482, x_481); -x_484 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__110; -lean_inc(x_471); -lean_inc(x_35); -x_485 = l_Lean_Syntax_node4(x_35, x_484, x_471, x_473, x_475, x_483); -x_486 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__107; -lean_inc(x_299); -lean_inc(x_35); -x_487 = l_Lean_Syntax_node3(x_35, x_486, x_467, x_299, x_485); -lean_inc(x_299); -lean_inc(x_35); -x_488 = l_Lean_Syntax_node2(x_35, x_464, x_487, x_299); -x_489 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__34; -x_490 = l_Lean_addMacroScope(x_298, x_489, x_56); -x_491 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__33; -x_492 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__38; -lean_inc(x_35); -x_493 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_493, 0, x_35); -lean_ctor_set(x_493, 1, x_491); -lean_ctor_set(x_493, 2, x_490); -lean_ctor_set(x_493, 3, x_492); -x_494 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__39; lean_inc(x_35); -x_495 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_495, 0, x_35); -lean_ctor_set(x_495, 1, x_494); -lean_inc(x_299); +x_486 = l_Lean_Syntax_node2(x_35, x_462, x_485, x_298); +x_487 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__34; +x_488 = l_Lean_addMacroScope(x_297, x_487, x_56); +x_489 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__33; +x_490 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__38; lean_inc(x_35); -x_496 = l_Lean_Syntax_node2(x_35, x_430, x_299, x_471); +x_491 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_491, 0, x_35); +lean_ctor_set(x_491, 1, x_489); +lean_ctor_set(x_491, 2, x_488); +lean_ctor_set(x_491, 3, x_490); +x_492 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__39; lean_inc(x_35); -x_497 = l_Lean_Syntax_node1(x_35, x_52, x_496); -x_498 = lean_array_get_size(x_27); -x_499 = lean_usize_of_nat(x_498); -lean_dec(x_498); -x_500 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__18(x_38, x_499, x_18, x_27); -x_501 = lean_array_get_size(x_500); -x_502 = lean_usize_of_nat(x_501); -lean_dec(x_501); -x_503 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__17(x_502, x_18, x_500); -x_504 = l_Array_append___rarg(x_50, x_503); +x_493 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_493, 0, x_35); +lean_ctor_set(x_493, 1, x_492); +lean_inc(x_298); lean_inc(x_35); -x_505 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_505, 0, x_35); -lean_ctor_set(x_505, 1, x_52); -lean_ctor_set(x_505, 2, x_504); +x_494 = l_Lean_Syntax_node2(x_35, x_428, x_298, x_469); lean_inc(x_35); -x_506 = l_Lean_Syntax_node1(x_35, x_443, x_505); -lean_inc_n(x_299, 2); +x_495 = l_Lean_Syntax_node1(x_35, x_52, x_494); +x_496 = lean_array_get_size(x_27); +x_497 = lean_usize_of_nat(x_496); +lean_dec(x_496); +x_498 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__18(x_38, x_497, x_18, x_27); +x_499 = lean_array_get_size(x_498); +x_500 = lean_usize_of_nat(x_499); +lean_dec(x_499); +x_501 = l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__17(x_500, x_18, x_498); +x_502 = l_Array_append___rarg(x_50, x_501); lean_inc(x_35); -x_507 = l_Lean_Syntax_node6(x_35, x_445, x_429, x_299, x_299, x_497, x_434, x_506); -x_508 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__31; +x_503 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_503, 0, x_35); +lean_ctor_set(x_503, 1, x_52); +lean_ctor_set(x_503, 2, x_502); lean_inc(x_35); -x_509 = l_Lean_Syntax_node3(x_35, x_508, x_493, x_495, x_507); +x_504 = l_Lean_Syntax_node1(x_35, x_441, x_503); +lean_inc_n(x_298, 2); lean_inc(x_35); -x_510 = l_Lean_Syntax_node1(x_35, x_482, x_509); -lean_inc(x_299); +x_505 = l_Lean_Syntax_node6(x_35, x_443, x_427, x_298, x_298, x_495, x_432, x_504); +x_506 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__31; lean_inc(x_35); -x_511 = l_Lean_Syntax_node2(x_35, x_464, x_510, x_299); +x_507 = l_Lean_Syntax_node3(x_35, x_506, x_491, x_493, x_505); lean_inc(x_35); -x_512 = l_Lean_Syntax_node3(x_35, x_52, x_465, x_488, x_511); -x_513 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__103; +x_508 = l_Lean_Syntax_node1(x_35, x_480, x_507); +lean_inc(x_298); lean_inc(x_35); -x_514 = l_Lean_Syntax_node1(x_35, x_513, x_512); -x_515 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__101; +x_509 = l_Lean_Syntax_node2(x_35, x_462, x_508, x_298); lean_inc(x_35); -x_516 = l_Lean_Syntax_node2(x_35, x_515, x_461, x_514); -lean_inc(x_377); -lean_inc(x_299); +x_510 = l_Lean_Syntax_node3(x_35, x_52, x_463, x_486, x_509); +x_511 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__103; lean_inc(x_35); -x_517 = l_Lean_Syntax_node5(x_35, x_449, x_402, x_459, x_299, x_377, x_516); +x_512 = l_Lean_Syntax_node1(x_35, x_511, x_510); +x_513 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__101; lean_inc(x_35); -x_518 = l_Lean_Syntax_node1(x_35, x_451, x_517); -lean_inc_n(x_299, 2); +x_514 = l_Lean_Syntax_node2(x_35, x_513, x_459, x_512); +lean_inc(x_376); +lean_inc(x_298); lean_inc(x_35); -x_519 = l_Lean_Syntax_node3(x_35, x_453, x_299, x_299, x_518); -lean_inc(x_299); +x_515 = l_Lean_Syntax_node5(x_35, x_447, x_401, x_457, x_298, x_376, x_514); lean_inc(x_35); -x_520 = l_Lean_Syntax_node3(x_35, x_52, x_454, x_299, x_519); -x_521 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__80; +x_516 = l_Lean_Syntax_node1(x_35, x_449, x_515); +lean_inc_n(x_298, 2); lean_inc(x_35); -x_522 = l_Lean_Syntax_node2(x_35, x_521, x_313, x_520); +x_517 = l_Lean_Syntax_node3(x_35, x_451, x_298, x_298, x_516); +lean_inc(x_298); +lean_inc(x_35); +x_518 = l_Lean_Syntax_node3(x_35, x_52, x_452, x_298, x_517); +x_519 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__80; lean_inc(x_35); -x_523 = l_Lean_Syntax_node1(x_35, x_52, x_522); -x_524 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__66; +x_520 = l_Lean_Syntax_node2(x_35, x_519, x_312, x_518); lean_inc(x_35); -x_525 = l_Lean_Syntax_node3(x_35, x_524, x_377, x_410, x_523); -x_526 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__47; -lean_inc_n(x_299, 3); +x_521 = l_Lean_Syntax_node1(x_35, x_52, x_520); +x_522 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__66; lean_inc(x_35); -x_527 = l_Lean_Syntax_node8(x_35, x_526, x_360, x_362, x_299, x_299, x_375, x_525, x_299, x_299); +x_523 = l_Lean_Syntax_node3(x_35, x_522, x_376, x_409, x_521); +x_524 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__47; +lean_inc_n(x_298, 3); lean_inc(x_35); -x_528 = l_Lean_Syntax_node2(x_35, x_343, x_358, x_527); -x_529 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__43; +x_525 = l_Lean_Syntax_node8(x_35, x_524, x_359, x_361, x_298, x_298, x_374, x_523, x_298, x_298); lean_inc(x_35); -x_530 = l_Lean_Syntax_node3(x_35, x_529, x_350, x_352, x_528); -x_531 = l_Lean_Syntax_node2(x_35, x_52, x_344, x_530); -x_532 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_532, 0, x_531); -lean_ctor_set(x_532, 1, x_296); -return x_532; +x_526 = l_Lean_Syntax_node2(x_35, x_342, x_357, x_525); +x_527 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__43; +lean_inc(x_35); +x_528 = l_Lean_Syntax_node3(x_35, x_527, x_349, x_351, x_526); +x_529 = l_Lean_Syntax_node2(x_35, x_52, x_343, x_528); +x_530 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_530, 0, x_529); +lean_ctor_set(x_530, 1, x_295); +return x_530; } } else { -uint8_t x_533; +uint8_t x_531; lean_dec(x_27); lean_dec(x_26); lean_dec(x_23); @@ -7593,29 +7589,29 @@ lean_dec(x_10); lean_dec(x_9); lean_dec(x_3); lean_dec(x_1); -x_533 = !lean_is_exclusive(x_30); -if (x_533 == 0) +x_531 = !lean_is_exclusive(x_30); +if (x_531 == 0) { return x_30; } else { -lean_object* x_534; lean_object* x_535; lean_object* x_536; -x_534 = lean_ctor_get(x_30, 0); -x_535 = lean_ctor_get(x_30, 1); -lean_inc(x_535); -lean_inc(x_534); +lean_object* x_532; lean_object* x_533; lean_object* x_534; +x_532 = lean_ctor_get(x_30, 0); +x_533 = lean_ctor_get(x_30, 1); +lean_inc(x_533); +lean_inc(x_532); lean_dec(x_30); -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; +x_534 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_534, 0, x_532); +lean_ctor_set(x_534, 1, x_533); +return x_534; } } } else { -uint8_t x_537; +uint8_t x_535; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -7625,23 +7621,23 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_537 = !lean_is_exclusive(x_19); -if (x_537 == 0) +x_535 = !lean_is_exclusive(x_19); +if (x_535 == 0) { return x_19; } else { -lean_object* x_538; lean_object* x_539; lean_object* x_540; -x_538 = lean_ctor_get(x_19, 0); -x_539 = lean_ctor_get(x_19, 1); -lean_inc(x_539); -lean_inc(x_538); +lean_object* x_536; lean_object* x_537; lean_object* x_538; +x_536 = lean_ctor_get(x_19, 0); +x_537 = lean_ctor_get(x_19, 1); +lean_inc(x_537); +lean_inc(x_536); lean_dec(x_19); -x_540 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_540, 0, x_538); -lean_ctor_set(x_540, 1, x_539); -return x_540; +x_538 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_538, 0, x_536); +lean_ctor_set(x_538, 1, x_537); +return x_538; } } } @@ -8627,7 +8623,6 @@ lean_closure_set(x_9, 1, x_1); x_10 = lean_alloc_closure((void*)(l_Lean_Meta_forallTelescopeReducing___at_Lean_Elab_Deriving_mkInductArgNames___spec__2___rarg), 9, 2); lean_closure_set(x_10, 0, x_8); lean_closure_set(x_10, 1, x_9); -lean_inc(x_4); x_11 = l_Lean_Elab_Command_liftTermElabM___rarg(x_10, x_4, x_5, x_6); if (lean_obj_tag(x_11) == 0) { @@ -8983,7 +8978,7 @@ lean_dec(x_3); return x_7; } } -static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__1() { +static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__1() { _start: { lean_object* x_1; @@ -8991,7 +8986,7 @@ x_1 = lean_alloc_closure((void*)(l___private_Lean_Server_Rpc_Deriving_0__Lean_Se return x_1; } } -static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__2() { +static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -9001,27 +8996,27 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__3() { +static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__2; +x_1 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__2; x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__28; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__4() { +static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__3; +x_1 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__3; x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__29; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__5() { +static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__5() { _start: { lean_object* x_1; @@ -9029,17 +9024,17 @@ x_1 = lean_mk_string_from_bytes("initFn", 6); return x_1; } } -static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__6() { +static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__4; -x_2 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__5; +x_1 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__4; +x_2 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__7() { +static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__7() { _start: { lean_object* x_1; @@ -9047,37 +9042,37 @@ x_1 = lean_mk_string_from_bytes("_@", 2); return x_1; } } -static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__8() { +static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__6; -x_2 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__7; +x_1 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__6; +x_2 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__7; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__9() { +static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__8; +x_1 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__8; x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__10() { +static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__9; +x_1 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__9; x_2 = l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__28; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__11() { +static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__11() { _start: { lean_object* x_1; @@ -9085,27 +9080,27 @@ x_1 = lean_mk_string_from_bytes("Rpc", 3); return x_1; } } -static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__12() { +static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__10; -x_2 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__11; +x_1 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__10; +x_2 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__11; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__13() { +static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__12; +x_1 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__12; x_2 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__14() { +static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__14() { _start: { lean_object* x_1; @@ -9113,32 +9108,32 @@ x_1 = lean_mk_string_from_bytes("_hyg", 4); return x_1; } } -static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__15() { +static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__13; -x_2 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__14; +x_1 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__13; +x_2 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__14; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__16() { +static lean_object* _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__15; -x_2 = lean_unsigned_to_nat(5234u); +x_1 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__15; +x_2 = lean_unsigned_to_nat(5228u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228_(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_RpcEncodable_deriveStructureInstance___lambda__1___closed__54; -x_3 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__1; +x_3 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__1; x_4 = l_Lean_Elab_registerDerivingHandler(x_2, x_3, x_1); if (lean_obj_tag(x_4) == 0) { @@ -9148,7 +9143,7 @@ lean_inc(x_5); lean_dec(x_4); x_6 = l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___closed__3; x_7 = 0; -x_8 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__16; +x_8 = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__16; x_9 = l_Lean_registerTraceClass(x_6, x_7, x_8, x_5); return x_9; } @@ -9753,39 +9748,39 @@ l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInstance_ lean_mark_persistent(l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInstance___closed__1); l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInstance___closed__2 = _init_l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInstance___closed__2(); lean_mark_persistent(l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInstance___closed__2); -l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__1 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__1(); -lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__1); -l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__2 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__2(); -lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__2); -l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__3 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__3(); -lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__3); -l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__4 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__4(); -lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__4); -l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__5 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__5(); -lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__5); -l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__6 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__6(); -lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__6); -l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__7 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__7(); -lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__7); -l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__8 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__8(); -lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__8); -l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__9 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__9(); -lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__9); -l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__10 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__10(); -lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__10); -l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__11 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__11(); -lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__11); -l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__12 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__12(); -lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__12); -l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__13 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__13(); -lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__13); -l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__14 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__14(); -lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__14); -l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__15 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__15(); -lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__15); -l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__16 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__16(); -lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234____closed__16); -res = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5234_(lean_io_mk_world()); +l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__1 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__1(); +lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__1); +l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__2 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__2(); +lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__2); +l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__3 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__3(); +lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__3); +l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__4 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__4(); +lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__4); +l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__5 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__5(); +lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__5); +l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__6 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__6(); +lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__6); +l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__7 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__7(); +lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__7); +l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__8 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__8(); +lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__8); +l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__9 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__9(); +lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__9); +l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__10 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__10(); +lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__10); +l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__11 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__11(); +lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__11); +l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__12 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__12(); +lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__12); +l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__13 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__13(); +lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__13); +l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__14 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__14(); +lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__14); +l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__15 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__15(); +lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__15); +l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__16 = _init_l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__16(); +lean_mark_persistent(l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228____closed__16); +res = l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_5228_(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 5a08c2e277b6..1f58f4658897 100644 --- a/stage0/stdlib/Lean/Server/Snapshots.c +++ b/stage0/stdlib/Lean/Server/Snapshots.c @@ -13,7 +13,6 @@ #ifdef __cplusplus extern "C" { #endif -uint8_t l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(lean_object*, lean_object*); lean_object* l___private_Init_Util_0__outOfBounds___rarg(lean_object*); 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*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapMAux___at_Lean_Server_Snapshots_Snapshot_diagnostics___spec__2(lean_object*); @@ -83,6 +82,7 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_Snapshots_Snaps LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_Snapshots_Snapshot_diagnostics___spec__3(size_t, size_t, lean_object*); lean_object* l_IO_FS_Stream_ofBuffer(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_Snapshots_Snapshot_runCommandElabM___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Server_Snapshots_initFn____x40_Lean_Server_Snapshots___hyg_6____closed__1; 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_Snapshots_compileNextCmd___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); @@ -1711,7 +1711,7 @@ lean_closure_set(x_46, 0, x_45); lean_closure_set(x_46, 1, x_44); lean_closure_set(x_46, 2, x_27); x_47 = l_Lean_Server_Snapshots_compileNextCmd___closed__1; -x_48 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_12, x_47); +x_48 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_12, x_47); lean_dec(x_12); x_49 = l_IO_FS_withIsolatedStreams___at_Lean_Server_Snapshots_compileNextCmd___spec__1(x_46, x_48, x_35); if (lean_obj_tag(x_49) == 0) @@ -2025,7 +2025,7 @@ lean_closure_set(x_154, 0, x_153); lean_closure_set(x_154, 1, x_152); lean_closure_set(x_154, 2, x_135); x_155 = l_Lean_Server_Snapshots_compileNextCmd___closed__1; -x_156 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_119, x_155); +x_156 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_119, x_155); lean_dec(x_119); x_157 = l_IO_FS_withIsolatedStreams___at_Lean_Server_Snapshots_compileNextCmd___spec__1(x_154, x_156, x_143); if (lean_obj_tag(x_157) == 0) diff --git a/stage0/stdlib/Lean/Server/Watchdog.c b/stage0/stdlib/Lean/Server/Watchdog.c index e8d8914088a6..cec6f6a96e19 100644 --- a/stage0/stdlib/Lean/Server/Watchdog.c +++ b/stage0/stdlib/Lean/Server/Watchdog.c @@ -234,6 +234,7 @@ LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_parseParams___at_Lean_Server_Wat static lean_object* l_Lean_Server_Watchdog_tryWriteMessage___closed__60; lean_object* l_Lean_Server_References_removeWorkerRefs(lean_object*, lean_object*); static lean_object* l_Lean_Server_Watchdog_initAndRunWatchdog___closed__4; +lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210_(lean_object*); LEAN_EXPORT lean_object* l_IO_FS_Stream_writeLspResponse___at_Lean_Server_Watchdog_handleRequest___spec__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_forwardNotification(lean_object*); LEAN_EXPORT lean_object* lean_server_watchdog_main(lean_object*, lean_object*); @@ -359,7 +360,6 @@ lean_object* lean_get_prefix(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_Watchdog_terminateFileWorker___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_Watchdog_initAndRunWatchdog___closed__18; static lean_object* l___private_Lean_Server_Watchdog_0__Lean_Server_Watchdog_forwardMessages_loop___closed__10; -lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227_(lean_object*); static lean_object* l_Lean_Server_Watchdog_tryWriteMessage___closed__58; static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_Watchdog_initAndRunWatchdog___spec__2___closed__5; static lean_object* l_Lean_Server_Watchdog_tryWriteMessage___closed__8; @@ -4584,7 +4584,7 @@ lean_inc(x_133); lean_dec(x_132); x_134 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_134, 0, x_133); -x_135 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227_(x_134); +x_135 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210_(x_134); lean_dec(x_134); if (lean_obj_tag(x_135) == 0) { @@ -4625,7 +4625,7 @@ lean_inc(x_143); lean_dec(x_132); x_144 = lean_alloc_ctor(5, 1, 0); lean_ctor_set(x_144, 0, x_143); -x_145 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227_(x_144); +x_145 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210_(x_144); lean_dec(x_144); if (lean_obj_tag(x_145) == 0) { @@ -4687,7 +4687,7 @@ lean_inc(x_155); lean_dec(x_154); x_156 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_156, 0, x_155); -x_157 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227_(x_156); +x_157 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210_(x_156); lean_dec(x_156); if (lean_obj_tag(x_157) == 0) { @@ -4728,7 +4728,7 @@ lean_inc(x_165); lean_dec(x_154); x_166 = lean_alloc_ctor(5, 1, 0); lean_ctor_set(x_166, 0, x_165); -x_167 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1227_(x_166); +x_167 = l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_fromJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1210_(x_166); lean_dec(x_166); if (lean_obj_tag(x_167) == 0) { diff --git a/stage0/stdlib/Lean/SubExpr.c b/stage0/stdlib/Lean/SubExpr.c index 755ca9104b4e..660469764e26 100644 --- a/stage0/stdlib/Lean/SubExpr.c +++ b/stage0/stdlib/Lean/SubExpr.c @@ -16,7 +16,6 @@ extern "C" { static lean_object* l_Lean_SubExpr_Pos_push___closed__3; lean_object* l___private_Init_Util_0__outOfBounds___rarg(lean_object*); lean_object* l_List_join___rarg(lean_object*); -lean_object* l_Lean_Syntax_decodeNameLit(lean_object*); static lean_object* l___private_Lean_SubExpr_0__Lean_SubExpr_fromJsonGoalsLocation____x40_Lean_SubExpr___hyg_2016____closed__10; LEAN_EXPORT lean_object* l_Lean_SubExpr_instToJsonMVarId(lean_object*); LEAN_EXPORT lean_object* l_Lean_SubExpr_Pos_foldrM___at_Lean_SubExpr_Pos_all___spec__1(lean_object*, lean_object*, lean_object*); @@ -51,9 +50,11 @@ lean_object* lean_array_push(lean_object*, lean_object*); static lean_object* l_Lean_SubExpr_Pos_append___closed__1; uint8_t lean_usize_dec_eq(size_t, size_t); static lean_object* l___private_Lean_SubExpr_0__Lean_SubExpr_fromJsonGoalsLocation____x40_Lean_SubExpr___hyg_2016____closed__12; +uint8_t l_Lean_Name_isAnonymous(lean_object*); LEAN_EXPORT lean_object* l_Lean_SubExpr_Pos_pushLetBody(lean_object*); static lean_object* l___private_Lean_SubExpr_0__Lean_SubExpr_Pos_ofStringCoord___closed__2; LEAN_EXPORT lean_object* l___private_Lean_SubExpr_0__Lean_SubExpr_fromJsonGoalsLocation____x40_Lean_SubExpr___hyg_2016_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_SubExpr_instFromJsonFVarId___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_SubExpr_0__Lean_SubExpr_fromJsonGoalsLocation____x40_Lean_SubExpr___hyg_2016____spec__1(lean_object*, lean_object*); lean_object* l_Lean_Expr_bvar___override(lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); @@ -61,7 +62,6 @@ static lean_object* l_Lean_SubExpr_Pos_push___closed__1; static lean_object* l___private_Lean_SubExpr_0__Lean_SubExpr_fromJsonGoalsLocation____x40_Lean_SubExpr___hyg_2016____closed__13; static lean_object* l_Lean_instInhabitedSubExpr___closed__1; static lean_object* l___private_Lean_SubExpr_0__Lean_SubExpr_fromJsonGoalsLocation____x40_Lean_SubExpr___hyg_2016____closed__11; -static lean_object* l_Lean_SubExpr_instFromJsonFVarId___closed__5; LEAN_EXPORT uint8_t l_Lean_SubExpr_isRoot(lean_object*); LEAN_EXPORT lean_object* l_Lean_SubExpr_instToJsonGoalsLocation; static lean_object* l___private_Lean_SubExpr_0__Lean_SubExpr_fromJsonGoalsLocation____x40_Lean_SubExpr___hyg_2016____closed__8; @@ -190,6 +190,7 @@ LEAN_EXPORT lean_object* l___private_Lean_SubExpr_0__Lean_SubExpr_fromJsonGoalLo LEAN_EXPORT lean_object* l___private_Lean_SubExpr_0__Lean_SubExpr_fromJsonGoalLocation____x40_Lean_SubExpr___hyg_1589____lambda__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_SubExpr_Pos_append(lean_object*, lean_object*); static lean_object* l___private_Lean_SubExpr_0__Lean_SubExpr_fromJsonGoalsLocation____x40_Lean_SubExpr___hyg_2016____closed__16; +lean_object* l_String_toName(lean_object*); static lean_object* l_Lean_SubExpr_instToJsonGoalsLocation___closed__1; LEAN_EXPORT lean_object* l_Lean_SubExpr_mkRoot(lean_object*); static lean_object* l_Lean_SubExpr_Pos_toArray___closed__2; @@ -208,6 +209,7 @@ lean_object* lean_nat_mul(lean_object*, lean_object*); static lean_object* l___private_Lean_SubExpr_0__Lean_SubExpr_fromJsonGoalLocation____x40_Lean_SubExpr___hyg_1589____lambda__1___closed__1; static lean_object* l___private_Lean_SubExpr_0__Lean_SubExpr_fromJsonGoalsLocation____x40_Lean_SubExpr___hyg_2016____closed__5; LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_SubExpr_Pos_toString___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_SubExpr_instFromJsonFVarId___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_SubExpr_Pos_instReprPos___boxed(lean_object*, lean_object*); lean_object* l_List_reverse___rarg(lean_object*); static lean_object* l___private_Lean_SubExpr_0__Lean_SubExpr_Pos_ofStringCoord___closed__7; @@ -2232,23 +2234,24 @@ lean_ctor_set(x_4, 0, x_3); return x_4; } } -static lean_object* _init_l_Lean_SubExpr_instFromJsonFVarId___closed__1() { +LEAN_EXPORT lean_object* l_Lean_SubExpr_instFromJsonFVarId___lambda__1(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("[anonymous]", 11); -return x_1; +lean_object* x_3; +x_3 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_3, 0, x_1); +return x_3; } } -static lean_object* _init_l_Lean_SubExpr_instFromJsonFVarId___closed__2() { +static lean_object* _init_l_Lean_SubExpr_instFromJsonFVarId___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("`", 1); +x_1 = lean_mk_string_from_bytes("[anonymous]", 11); return x_1; } } -static lean_object* _init_l_Lean_SubExpr_instFromJsonFVarId___closed__3() { +static lean_object* _init_l_Lean_SubExpr_instFromJsonFVarId___closed__2() { _start: { lean_object* x_1; @@ -2256,7 +2259,7 @@ x_1 = lean_mk_string_from_bytes("expected a `Name`, got '", 24); return x_1; } } -static lean_object* _init_l_Lean_SubExpr_instFromJsonFVarId___closed__4() { +static lean_object* _init_l_Lean_SubExpr_instFromJsonFVarId___closed__3() { _start: { lean_object* x_1; @@ -2264,7 +2267,7 @@ x_1 = lean_mk_string_from_bytes("'", 1); return x_1; } } -static lean_object* _init_l_Lean_SubExpr_instFromJsonFVarId___closed__5() { +static lean_object* _init_l_Lean_SubExpr_instFromJsonFVarId___closed__4() { _start: { lean_object* x_1; lean_object* x_2; @@ -2311,97 +2314,97 @@ x_8 = l_Lean_SubExpr_instFromJsonFVarId___closed__1; x_9 = lean_string_dec_eq(x_7, x_8); if (x_9 == 0) { -lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = l_Lean_SubExpr_instFromJsonFVarId___closed__2; -x_11 = lean_string_append(x_10, x_7); -lean_dec(x_7); -x_12 = l_Lean_Syntax_decodeNameLit(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; -x_13 = lean_unsigned_to_nat(80u); -x_14 = l_Lean_Json_pretty(x_1, x_13); -x_15 = l_Lean_SubExpr_instFromJsonFVarId___closed__3; -x_16 = lean_string_append(x_15, x_14); -lean_dec(x_14); -x_17 = l_Lean_SubExpr_instFromJsonFVarId___closed__4; -x_18 = lean_string_append(x_16, x_17); -lean_ctor_set_tag(x_2, 0); -lean_ctor_set(x_2, 0, x_18); +lean_object* x_10; uint8_t x_11; +x_10 = l_String_toName(x_7); +x_11 = l_Lean_Name_isAnonymous(x_10); +if (x_11 == 0) +{ +lean_dec(x_1); +lean_ctor_set(x_2, 0, x_10); return x_2; } else { -lean_object* x_19; -lean_dec(x_1); -x_19 = lean_ctor_get(x_12, 0); -lean_inc(x_19); -lean_dec(x_12); -lean_ctor_set(x_2, 0, x_19); +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_dec(x_10); +x_12 = lean_unsigned_to_nat(80u); +x_13 = l_Lean_Json_pretty(x_1, x_12); +x_14 = l_Lean_SubExpr_instFromJsonFVarId___closed__2; +x_15 = lean_string_append(x_14, x_13); +lean_dec(x_13); +x_16 = l_Lean_SubExpr_instFromJsonFVarId___closed__3; +x_17 = lean_string_append(x_15, x_16); +lean_ctor_set_tag(x_2, 0); +lean_ctor_set(x_2, 0, x_17); return x_2; } } else { -lean_object* x_20; +lean_object* x_18; lean_free_object(x_2); lean_dec(x_7); lean_dec(x_1); -x_20 = l_Lean_SubExpr_instFromJsonFVarId___closed__5; -return x_20; +x_18 = l_Lean_SubExpr_instFromJsonFVarId___closed__4; +return x_18; } } else { -lean_object* x_21; lean_object* x_22; uint8_t x_23; -x_21 = lean_ctor_get(x_2, 0); -lean_inc(x_21); +lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_19 = lean_ctor_get(x_2, 0); +lean_inc(x_19); lean_dec(x_2); -x_22 = l_Lean_SubExpr_instFromJsonFVarId___closed__1; -x_23 = lean_string_dec_eq(x_21, x_22); -if (x_23 == 0) +x_20 = l_Lean_SubExpr_instFromJsonFVarId___closed__1; +x_21 = lean_string_dec_eq(x_19, x_20); +if (x_21 == 0) { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = l_Lean_SubExpr_instFromJsonFVarId___closed__2; -x_25 = lean_string_append(x_24, x_21); -lean_dec(x_21); -x_26 = l_Lean_Syntax_decodeNameLit(x_25); -if (lean_obj_tag(x_26) == 0) +lean_object* x_22; uint8_t x_23; +x_22 = l_String_toName(x_19); +x_23 = l_Lean_Name_isAnonymous(x_22); +if (x_23 == 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_unsigned_to_nat(80u); -x_28 = l_Lean_Json_pretty(x_1, x_27); -x_29 = l_Lean_SubExpr_instFromJsonFVarId___closed__3; -x_30 = lean_string_append(x_29, x_28); -lean_dec(x_28); -x_31 = l_Lean_SubExpr_instFromJsonFVarId___closed__4; -x_32 = lean_string_append(x_30, x_31); -x_33 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_33, 0, x_32); -return x_33; +lean_object* x_24; +lean_dec(x_1); +x_24 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_24, 0, x_22); +return x_24; } else { -lean_object* x_34; lean_object* x_35; -lean_dec(x_1); -x_34 = lean_ctor_get(x_26, 0); -lean_inc(x_34); +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +lean_dec(x_22); +x_25 = lean_unsigned_to_nat(80u); +x_26 = l_Lean_Json_pretty(x_1, x_25); +x_27 = l_Lean_SubExpr_instFromJsonFVarId___closed__2; +x_28 = lean_string_append(x_27, x_26); lean_dec(x_26); -x_35 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_35, 0, x_34); -return x_35; +x_29 = l_Lean_SubExpr_instFromJsonFVarId___closed__3; +x_30 = lean_string_append(x_28, x_29); +x_31 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_31, 0, x_30); +return x_31; } } else { -lean_object* x_36; -lean_dec(x_21); +lean_object* x_32; +lean_dec(x_19); lean_dec(x_1); -x_36 = l_Lean_SubExpr_instFromJsonFVarId___closed__5; -return x_36; +x_32 = l_Lean_SubExpr_instFromJsonFVarId___closed__4; +return x_32; +} +} } } } +LEAN_EXPORT lean_object* l_Lean_SubExpr_instFromJsonFVarId___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_SubExpr_instFromJsonFVarId___lambda__1(x_1, x_2); +lean_dec(x_2); +return x_3; } } LEAN_EXPORT lean_object* l_Lean_SubExpr_instFromJsonMVarId(lean_object* x_1) { @@ -2441,94 +2444,85 @@ x_8 = l_Lean_SubExpr_instFromJsonFVarId___closed__1; x_9 = lean_string_dec_eq(x_7, x_8); if (x_9 == 0) { -lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = l_Lean_SubExpr_instFromJsonFVarId___closed__2; -x_11 = lean_string_append(x_10, x_7); -lean_dec(x_7); -x_12 = l_Lean_Syntax_decodeNameLit(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; -x_13 = lean_unsigned_to_nat(80u); -x_14 = l_Lean_Json_pretty(x_1, x_13); -x_15 = l_Lean_SubExpr_instFromJsonFVarId___closed__3; -x_16 = lean_string_append(x_15, x_14); -lean_dec(x_14); -x_17 = l_Lean_SubExpr_instFromJsonFVarId___closed__4; -x_18 = lean_string_append(x_16, x_17); -lean_ctor_set_tag(x_2, 0); -lean_ctor_set(x_2, 0, x_18); +lean_object* x_10; uint8_t x_11; +x_10 = l_String_toName(x_7); +x_11 = l_Lean_Name_isAnonymous(x_10); +if (x_11 == 0) +{ +lean_dec(x_1); +lean_ctor_set(x_2, 0, x_10); return x_2; } else { -lean_object* x_19; -lean_dec(x_1); -x_19 = lean_ctor_get(x_12, 0); -lean_inc(x_19); -lean_dec(x_12); -lean_ctor_set(x_2, 0, x_19); +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_dec(x_10); +x_12 = lean_unsigned_to_nat(80u); +x_13 = l_Lean_Json_pretty(x_1, x_12); +x_14 = l_Lean_SubExpr_instFromJsonFVarId___closed__2; +x_15 = lean_string_append(x_14, x_13); +lean_dec(x_13); +x_16 = l_Lean_SubExpr_instFromJsonFVarId___closed__3; +x_17 = lean_string_append(x_15, x_16); +lean_ctor_set_tag(x_2, 0); +lean_ctor_set(x_2, 0, x_17); return x_2; } } else { -lean_object* x_20; +lean_object* x_18; lean_free_object(x_2); lean_dec(x_7); lean_dec(x_1); -x_20 = l_Lean_SubExpr_instFromJsonFVarId___closed__5; -return x_20; +x_18 = l_Lean_SubExpr_instFromJsonFVarId___closed__4; +return x_18; } } else { -lean_object* x_21; lean_object* x_22; uint8_t x_23; -x_21 = lean_ctor_get(x_2, 0); -lean_inc(x_21); +lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_19 = lean_ctor_get(x_2, 0); +lean_inc(x_19); lean_dec(x_2); -x_22 = l_Lean_SubExpr_instFromJsonFVarId___closed__1; -x_23 = lean_string_dec_eq(x_21, x_22); -if (x_23 == 0) +x_20 = l_Lean_SubExpr_instFromJsonFVarId___closed__1; +x_21 = lean_string_dec_eq(x_19, x_20); +if (x_21 == 0) { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = l_Lean_SubExpr_instFromJsonFVarId___closed__2; -x_25 = lean_string_append(x_24, x_21); -lean_dec(x_21); -x_26 = l_Lean_Syntax_decodeNameLit(x_25); -if (lean_obj_tag(x_26) == 0) +lean_object* x_22; uint8_t x_23; +x_22 = l_String_toName(x_19); +x_23 = l_Lean_Name_isAnonymous(x_22); +if (x_23 == 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_unsigned_to_nat(80u); -x_28 = l_Lean_Json_pretty(x_1, x_27); -x_29 = l_Lean_SubExpr_instFromJsonFVarId___closed__3; -x_30 = lean_string_append(x_29, x_28); -lean_dec(x_28); -x_31 = l_Lean_SubExpr_instFromJsonFVarId___closed__4; -x_32 = lean_string_append(x_30, x_31); -x_33 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_33, 0, x_32); -return x_33; +lean_object* x_24; +lean_dec(x_1); +x_24 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_24, 0, x_22); +return x_24; } else { -lean_object* x_34; lean_object* x_35; -lean_dec(x_1); -x_34 = lean_ctor_get(x_26, 0); -lean_inc(x_34); +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +lean_dec(x_22); +x_25 = lean_unsigned_to_nat(80u); +x_26 = l_Lean_Json_pretty(x_1, x_25); +x_27 = l_Lean_SubExpr_instFromJsonFVarId___closed__2; +x_28 = lean_string_append(x_27, x_26); lean_dec(x_26); -x_35 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_35, 0, x_34); -return x_35; +x_29 = l_Lean_SubExpr_instFromJsonFVarId___closed__3; +x_30 = lean_string_append(x_28, x_29); +x_31 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_31, 0, x_30); +return x_31; } } else { -lean_object* x_36; -lean_dec(x_21); +lean_object* x_32; +lean_dec(x_19); lean_dec(x_1); -x_36 = l_Lean_SubExpr_instFromJsonFVarId___closed__5; -return x_36; +x_32 = l_Lean_SubExpr_instFromJsonFVarId___closed__4; +return x_32; } } } @@ -2609,305 +2603,27 @@ return x_13; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_77; uint8_t x_78; +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_77; lean_object* x_117; uint8_t x_118; x_14 = lean_ctor_get(x_6, 0); lean_inc(x_14); lean_dec(x_6); x_15 = lean_array_get_size(x_14); -x_77 = lean_unsigned_to_nat(0u); -x_78 = lean_nat_dec_lt(x_77, x_15); -if (x_78 == 0) -{ -lean_object* x_79; lean_object* x_80; lean_object* x_81; -x_79 = l_Lean_instInhabitedJson; -x_80 = l___private_Init_Util_0__outOfBounds___rarg(x_79); -x_81 = l_Lean_Json_getStr_x3f(x_80); -if (lean_obj_tag(x_81) == 0) -{ -uint8_t x_82; -lean_dec(x_80); -lean_dec(x_15); -lean_dec(x_14); -x_82 = !lean_is_exclusive(x_81); -if (x_82 == 0) -{ -lean_object* x_83; lean_object* x_84; -x_83 = l___private_Lean_SubExpr_0__Lean_SubExpr_fromJsonGoalLocation____x40_Lean_SubExpr___hyg_1589____lambda__2___closed__2; -x_84 = l_Except_orElseLazy___rarg(x_81, x_83); -return x_84; -} -else -{ -lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; -x_85 = lean_ctor_get(x_81, 0); -lean_inc(x_85); -lean_dec(x_81); -x_86 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_86, 0, x_85); -x_87 = l___private_Lean_SubExpr_0__Lean_SubExpr_fromJsonGoalLocation____x40_Lean_SubExpr___hyg_1589____lambda__2___closed__2; -x_88 = l_Except_orElseLazy___rarg(x_86, x_87); -return x_88; -} -} -else -{ -uint8_t x_89; -x_89 = !lean_is_exclusive(x_81); -if (x_89 == 0) -{ -lean_object* x_90; lean_object* x_91; uint8_t x_92; -x_90 = lean_ctor_get(x_81, 0); -x_91 = l_Lean_SubExpr_instFromJsonFVarId___closed__1; -x_92 = lean_string_dec_eq(x_90, x_91); -if (x_92 == 0) -{ -lean_object* x_93; lean_object* x_94; lean_object* x_95; -x_93 = l_Lean_SubExpr_instFromJsonFVarId___closed__2; -x_94 = lean_string_append(x_93, x_90); -lean_dec(x_90); -x_95 = l_Lean_Syntax_decodeNameLit(x_94); -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_dec(x_15); -lean_dec(x_14); -x_96 = lean_unsigned_to_nat(80u); -x_97 = l_Lean_Json_pretty(x_80, x_96); -x_98 = l_Lean_SubExpr_instFromJsonFVarId___closed__3; -x_99 = lean_string_append(x_98, x_97); -lean_dec(x_97); -x_100 = l_Lean_SubExpr_instFromJsonFVarId___closed__4; -x_101 = lean_string_append(x_99, x_100); -lean_ctor_set_tag(x_81, 0); -lean_ctor_set(x_81, 0, x_101); -x_102 = l___private_Lean_SubExpr_0__Lean_SubExpr_fromJsonGoalLocation____x40_Lean_SubExpr___hyg_1589____lambda__2___closed__2; -x_103 = l_Except_orElseLazy___rarg(x_81, x_102); -return x_103; -} -else -{ -lean_object* x_104; -lean_free_object(x_81); -lean_dec(x_80); -x_104 = lean_ctor_get(x_95, 0); -lean_inc(x_104); -lean_dec(x_95); -x_16 = x_104; -goto block_76; -} -} -else -{ -lean_object* x_105; -lean_free_object(x_81); -lean_dec(x_90); -lean_dec(x_80); -x_105 = lean_box(0); -x_16 = x_105; -goto block_76; -} -} -else -{ -lean_object* x_106; lean_object* x_107; uint8_t x_108; -x_106 = lean_ctor_get(x_81, 0); -lean_inc(x_106); -lean_dec(x_81); -x_107 = l_Lean_SubExpr_instFromJsonFVarId___closed__1; -x_108 = lean_string_dec_eq(x_106, x_107); -if (x_108 == 0) -{ -lean_object* x_109; lean_object* x_110; lean_object* x_111; -x_109 = l_Lean_SubExpr_instFromJsonFVarId___closed__2; -x_110 = lean_string_append(x_109, x_106); -lean_dec(x_106); -x_111 = l_Lean_Syntax_decodeNameLit(x_110); -if (lean_obj_tag(x_111) == 0) +x_117 = lean_unsigned_to_nat(0u); +x_118 = lean_nat_dec_lt(x_117, x_15); +if (x_118 == 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; lean_object* x_120; -lean_dec(x_15); -lean_dec(x_14); -x_112 = lean_unsigned_to_nat(80u); -x_113 = l_Lean_Json_pretty(x_80, x_112); -x_114 = l_Lean_SubExpr_instFromJsonFVarId___closed__3; -x_115 = lean_string_append(x_114, x_113); -lean_dec(x_113); -x_116 = l_Lean_SubExpr_instFromJsonFVarId___closed__4; -x_117 = lean_string_append(x_115, x_116); -x_118 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_118, 0, x_117); -x_119 = l___private_Lean_SubExpr_0__Lean_SubExpr_fromJsonGoalLocation____x40_Lean_SubExpr___hyg_1589____lambda__2___closed__2; -x_120 = l_Except_orElseLazy___rarg(x_118, x_119); -return x_120; +lean_object* x_119; lean_object* x_120; +x_119 = l_Lean_instInhabitedJson; +x_120 = l___private_Init_Util_0__outOfBounds___rarg(x_119); +x_77 = x_120; +goto block_116; } else { lean_object* x_121; -lean_dec(x_80); -x_121 = lean_ctor_get(x_111, 0); -lean_inc(x_121); -lean_dec(x_111); -x_16 = x_121; -goto block_76; -} -} -else -{ -lean_object* x_122; -lean_dec(x_106); -lean_dec(x_80); -x_122 = lean_box(0); -x_16 = x_122; -goto block_76; -} -} -} -} -else -{ -lean_object* x_123; lean_object* x_124; -x_123 = lean_array_fget(x_14, x_77); -x_124 = l_Lean_Json_getStr_x3f(x_123); -if (lean_obj_tag(x_124) == 0) -{ -uint8_t x_125; -lean_dec(x_123); -lean_dec(x_15); -lean_dec(x_14); -x_125 = !lean_is_exclusive(x_124); -if (x_125 == 0) -{ -lean_object* x_126; lean_object* x_127; -x_126 = l___private_Lean_SubExpr_0__Lean_SubExpr_fromJsonGoalLocation____x40_Lean_SubExpr___hyg_1589____lambda__2___closed__2; -x_127 = l_Except_orElseLazy___rarg(x_124, x_126); -return x_127; -} -else -{ -lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; -x_128 = lean_ctor_get(x_124, 0); -lean_inc(x_128); -lean_dec(x_124); -x_129 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_129, 0, x_128); -x_130 = l___private_Lean_SubExpr_0__Lean_SubExpr_fromJsonGoalLocation____x40_Lean_SubExpr___hyg_1589____lambda__2___closed__2; -x_131 = l_Except_orElseLazy___rarg(x_129, x_130); -return x_131; -} -} -else -{ -uint8_t x_132; -x_132 = !lean_is_exclusive(x_124); -if (x_132 == 0) -{ -lean_object* x_133; lean_object* x_134; uint8_t x_135; -x_133 = lean_ctor_get(x_124, 0); -x_134 = l_Lean_SubExpr_instFromJsonFVarId___closed__1; -x_135 = lean_string_dec_eq(x_133, x_134); -if (x_135 == 0) -{ -lean_object* x_136; lean_object* x_137; lean_object* x_138; -x_136 = l_Lean_SubExpr_instFromJsonFVarId___closed__2; -x_137 = lean_string_append(x_136, x_133); -lean_dec(x_133); -x_138 = l_Lean_Syntax_decodeNameLit(x_137); -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; lean_object* x_144; lean_object* x_145; lean_object* x_146; -lean_dec(x_15); -lean_dec(x_14); -x_139 = lean_unsigned_to_nat(80u); -x_140 = l_Lean_Json_pretty(x_123, x_139); -x_141 = l_Lean_SubExpr_instFromJsonFVarId___closed__3; -x_142 = lean_string_append(x_141, x_140); -lean_dec(x_140); -x_143 = l_Lean_SubExpr_instFromJsonFVarId___closed__4; -x_144 = lean_string_append(x_142, x_143); -lean_ctor_set_tag(x_124, 0); -lean_ctor_set(x_124, 0, x_144); -x_145 = l___private_Lean_SubExpr_0__Lean_SubExpr_fromJsonGoalLocation____x40_Lean_SubExpr___hyg_1589____lambda__2___closed__2; -x_146 = l_Except_orElseLazy___rarg(x_124, x_145); -return x_146; -} -else -{ -lean_object* x_147; -lean_free_object(x_124); -lean_dec(x_123); -x_147 = lean_ctor_get(x_138, 0); -lean_inc(x_147); -lean_dec(x_138); -x_16 = x_147; -goto block_76; -} -} -else -{ -lean_object* x_148; -lean_free_object(x_124); -lean_dec(x_133); -lean_dec(x_123); -x_148 = lean_box(0); -x_16 = x_148; -goto block_76; -} -} -else -{ -lean_object* x_149; lean_object* x_150; uint8_t x_151; -x_149 = lean_ctor_get(x_124, 0); -lean_inc(x_149); -lean_dec(x_124); -x_150 = l_Lean_SubExpr_instFromJsonFVarId___closed__1; -x_151 = lean_string_dec_eq(x_149, x_150); -if (x_151 == 0) -{ -lean_object* x_152; lean_object* x_153; lean_object* x_154; -x_152 = l_Lean_SubExpr_instFromJsonFVarId___closed__2; -x_153 = lean_string_append(x_152, x_149); -lean_dec(x_149); -x_154 = l_Lean_Syntax_decodeNameLit(x_153); -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; lean_object* x_161; lean_object* x_162; lean_object* x_163; -lean_dec(x_15); -lean_dec(x_14); -x_155 = lean_unsigned_to_nat(80u); -x_156 = l_Lean_Json_pretty(x_123, x_155); -x_157 = l_Lean_SubExpr_instFromJsonFVarId___closed__3; -x_158 = lean_string_append(x_157, x_156); -lean_dec(x_156); -x_159 = l_Lean_SubExpr_instFromJsonFVarId___closed__4; -x_160 = lean_string_append(x_158, x_159); -x_161 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_161, 0, x_160); -x_162 = l___private_Lean_SubExpr_0__Lean_SubExpr_fromJsonGoalLocation____x40_Lean_SubExpr___hyg_1589____lambda__2___closed__2; -x_163 = l_Except_orElseLazy___rarg(x_161, x_162); -return x_163; -} -else -{ -lean_object* x_164; -lean_dec(x_123); -x_164 = lean_ctor_get(x_154, 0); -lean_inc(x_164); -lean_dec(x_154); -x_16 = x_164; -goto block_76; -} -} -else -{ -lean_object* x_165; -lean_dec(x_149); -lean_dec(x_123); -x_165 = lean_box(0); -x_16 = x_165; -goto block_76; -} -} -} +x_121 = lean_array_fget(x_14, x_117); +x_77 = x_121; +goto block_116; } block_76: { @@ -3113,342 +2829,207 @@ return x_75; } } } -} -} -} -static lean_object* _init_l___private_Lean_SubExpr_0__Lean_SubExpr_fromJsonGoalLocation____x40_Lean_SubExpr___hyg_1589____lambda__3___closed__1() { -_start: +block_116: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("hypType", 7); -return x_1; -} -} -LEAN_EXPORT lean_object* l___private_Lean_SubExpr_0__Lean_SubExpr_fromJsonGoalLocation____x40_Lean_SubExpr___hyg_1589____lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +lean_object* x_78; +x_78 = l_Lean_Json_getStr_x3f(x_77); +if (lean_obj_tag(x_78) == 0) { -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_SubExpr_0__Lean_SubExpr_fromJsonGoalLocation____x40_Lean_SubExpr___hyg_1589____lambda__3___closed__1; -x_5 = lean_unsigned_to_nat(2u); -x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_SubExpr_0__Lean_SubExpr_fromJsonGoalLocation____x40_Lean_SubExpr___hyg_1589____lambda__2___boxed), 3, 2); -lean_closure_set(x_7, 0, x_1); -lean_closure_set(x_7, 1, x_2); -if (lean_obj_tag(x_6) == 0) -{ -uint8_t x_8; -x_8 = !lean_is_exclusive(x_6); -if (x_8 == 0) +uint8_t x_79; +lean_dec(x_77); +lean_dec(x_15); +lean_dec(x_14); +x_79 = !lean_is_exclusive(x_78); +if (x_79 == 0) { -lean_object* x_9; -x_9 = l_Except_orElseLazy___rarg(x_6, x_7); -return x_9; +lean_object* x_80; lean_object* x_81; +x_80 = l___private_Lean_SubExpr_0__Lean_SubExpr_fromJsonGoalLocation____x40_Lean_SubExpr___hyg_1589____lambda__2___closed__2; +x_81 = l_Except_orElseLazy___rarg(x_78, x_80); +return x_81; } else { -lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = lean_ctor_get(x_6, 0); -lean_inc(x_10); -lean_dec(x_6); -x_11 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_11, 0, x_10); -x_12 = l_Except_orElseLazy___rarg(x_11, x_7); -return x_12; +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; +x_82 = lean_ctor_get(x_78, 0); +lean_inc(x_82); +lean_dec(x_78); +x_83 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_83, 0, x_82); +x_84 = l___private_Lean_SubExpr_0__Lean_SubExpr_fromJsonGoalLocation____x40_Lean_SubExpr___hyg_1589____lambda__2___closed__2; +x_85 = l_Except_orElseLazy___rarg(x_83, x_84); +return x_85; } } else { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_64; uint8_t x_65; -x_13 = lean_ctor_get(x_6, 0); -lean_inc(x_13); -lean_dec(x_6); -x_14 = lean_array_get_size(x_13); -x_64 = lean_unsigned_to_nat(0u); -x_65 = lean_nat_dec_lt(x_64, x_14); -if (x_65 == 0) -{ -lean_object* x_66; lean_object* x_67; lean_object* x_68; -x_66 = l_Lean_instInhabitedJson; -x_67 = l___private_Init_Util_0__outOfBounds___rarg(x_66); -x_68 = l_Lean_Json_getStr_x3f(x_67); -if (lean_obj_tag(x_68) == 0) -{ -uint8_t x_69; -lean_dec(x_67); -lean_dec(x_14); -lean_dec(x_13); -x_69 = !lean_is_exclusive(x_68); -if (x_69 == 0) +uint8_t x_86; +x_86 = !lean_is_exclusive(x_78); +if (x_86 == 0) { -lean_object* x_70; -x_70 = l_Except_orElseLazy___rarg(x_68, x_7); -return x_70; -} -else +lean_object* x_87; lean_object* x_88; uint8_t x_89; +x_87 = lean_ctor_get(x_78, 0); +x_88 = l_Lean_SubExpr_instFromJsonFVarId___closed__1; +x_89 = lean_string_dec_eq(x_87, x_88); +if (x_89 == 0) { -lean_object* x_71; lean_object* x_72; lean_object* x_73; -x_71 = lean_ctor_get(x_68, 0); -lean_inc(x_71); -lean_dec(x_68); -x_72 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_72, 0, x_71); -x_73 = l_Except_orElseLazy___rarg(x_72, x_7); -return x_73; -} +lean_object* x_90; uint8_t x_91; +x_90 = l_String_toName(x_87); +x_91 = l_Lean_Name_isAnonymous(x_90); +if (x_91 == 0) +{ +lean_free_object(x_78); +lean_dec(x_77); +x_16 = x_90; +goto block_76; } else { -uint8_t x_74; -x_74 = !lean_is_exclusive(x_68); -if (x_74 == 0) -{ -lean_object* x_75; lean_object* x_76; uint8_t x_77; -x_75 = lean_ctor_get(x_68, 0); -x_76 = l_Lean_SubExpr_instFromJsonFVarId___closed__1; -x_77 = lean_string_dec_eq(x_75, x_76); -if (x_77 == 0) -{ -lean_object* x_78; lean_object* x_79; lean_object* x_80; -x_78 = l_Lean_SubExpr_instFromJsonFVarId___closed__2; -x_79 = lean_string_append(x_78, x_75); -lean_dec(x_75); -x_80 = l_Lean_Syntax_decodeNameLit(x_79); -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; lean_object* x_86; lean_object* x_87; +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_dec(x_90); +lean_dec(x_15); lean_dec(x_14); -lean_dec(x_13); -x_81 = lean_unsigned_to_nat(80u); -x_82 = l_Lean_Json_pretty(x_67, x_81); -x_83 = l_Lean_SubExpr_instFromJsonFVarId___closed__3; -x_84 = lean_string_append(x_83, x_82); -lean_dec(x_82); -x_85 = l_Lean_SubExpr_instFromJsonFVarId___closed__4; -x_86 = lean_string_append(x_84, x_85); -lean_ctor_set_tag(x_68, 0); -lean_ctor_set(x_68, 0, x_86); -x_87 = l_Except_orElseLazy___rarg(x_68, x_7); -return x_87; -} -else -{ -lean_object* x_88; -lean_free_object(x_68); -lean_dec(x_67); -x_88 = lean_ctor_get(x_80, 0); -lean_inc(x_88); -lean_dec(x_80); -x_15 = x_88; -goto block_63; +x_92 = lean_unsigned_to_nat(80u); +x_93 = l_Lean_Json_pretty(x_77, x_92); +x_94 = l_Lean_SubExpr_instFromJsonFVarId___closed__2; +x_95 = lean_string_append(x_94, x_93); +lean_dec(x_93); +x_96 = l_Lean_SubExpr_instFromJsonFVarId___closed__3; +x_97 = lean_string_append(x_95, x_96); +lean_ctor_set_tag(x_78, 0); +lean_ctor_set(x_78, 0, x_97); +x_98 = l___private_Lean_SubExpr_0__Lean_SubExpr_fromJsonGoalLocation____x40_Lean_SubExpr___hyg_1589____lambda__2___closed__2; +x_99 = l_Except_orElseLazy___rarg(x_78, x_98); +return x_99; } } else { -lean_object* x_89; -lean_free_object(x_68); -lean_dec(x_75); -lean_dec(x_67); -x_89 = lean_box(0); -x_15 = x_89; -goto block_63; +lean_object* x_100; +lean_free_object(x_78); +lean_dec(x_87); +lean_dec(x_77); +x_100 = lean_box(0); +x_16 = x_100; +goto block_76; } } else { -lean_object* x_90; lean_object* x_91; uint8_t x_92; -x_90 = lean_ctor_get(x_68, 0); -lean_inc(x_90); -lean_dec(x_68); -x_91 = l_Lean_SubExpr_instFromJsonFVarId___closed__1; -x_92 = lean_string_dec_eq(x_90, x_91); -if (x_92 == 0) +lean_object* x_101; lean_object* x_102; uint8_t x_103; +x_101 = lean_ctor_get(x_78, 0); +lean_inc(x_101); +lean_dec(x_78); +x_102 = l_Lean_SubExpr_instFromJsonFVarId___closed__1; +x_103 = lean_string_dec_eq(x_101, x_102); +if (x_103 == 0) { -lean_object* x_93; lean_object* x_94; lean_object* x_95; -x_93 = l_Lean_SubExpr_instFromJsonFVarId___closed__2; -x_94 = lean_string_append(x_93, x_90); -lean_dec(x_90); -x_95 = l_Lean_Syntax_decodeNameLit(x_94); -if (lean_obj_tag(x_95) == 0) +lean_object* x_104; uint8_t x_105; +x_104 = l_String_toName(x_101); +x_105 = l_Lean_Name_isAnonymous(x_104); +if (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_dec(x_14); -lean_dec(x_13); -x_96 = lean_unsigned_to_nat(80u); -x_97 = l_Lean_Json_pretty(x_67, x_96); -x_98 = l_Lean_SubExpr_instFromJsonFVarId___closed__3; -x_99 = lean_string_append(x_98, x_97); -lean_dec(x_97); -x_100 = l_Lean_SubExpr_instFromJsonFVarId___closed__4; -x_101 = lean_string_append(x_99, x_100); -x_102 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_102, 0, x_101); -x_103 = l_Except_orElseLazy___rarg(x_102, x_7); -return x_103; +lean_dec(x_77); +x_16 = x_104; +goto block_76; } else { -lean_object* x_104; -lean_dec(x_67); -x_104 = lean_ctor_get(x_95, 0); -lean_inc(x_104); -lean_dec(x_95); -x_15 = x_104; -goto block_63; +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_dec(x_104); +lean_dec(x_15); +lean_dec(x_14); +x_106 = lean_unsigned_to_nat(80u); +x_107 = l_Lean_Json_pretty(x_77, x_106); +x_108 = l_Lean_SubExpr_instFromJsonFVarId___closed__2; +x_109 = lean_string_append(x_108, x_107); +lean_dec(x_107); +x_110 = l_Lean_SubExpr_instFromJsonFVarId___closed__3; +x_111 = lean_string_append(x_109, x_110); +x_112 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_112, 0, x_111); +x_113 = l___private_Lean_SubExpr_0__Lean_SubExpr_fromJsonGoalLocation____x40_Lean_SubExpr___hyg_1589____lambda__2___closed__2; +x_114 = l_Except_orElseLazy___rarg(x_112, x_113); +return x_114; } } else { -lean_object* x_105; -lean_dec(x_90); -lean_dec(x_67); -x_105 = lean_box(0); -x_15 = x_105; -goto block_63; +lean_object* x_115; +lean_dec(x_101); +lean_dec(x_77); +x_115 = lean_box(0); +x_16 = x_115; +goto block_76; } } } } -else -{ -lean_object* x_106; lean_object* x_107; -x_106 = lean_array_fget(x_13, x_64); -x_107 = l_Lean_Json_getStr_x3f(x_106); -if (lean_obj_tag(x_107) == 0) -{ -uint8_t x_108; -lean_dec(x_106); -lean_dec(x_14); -lean_dec(x_13); -x_108 = !lean_is_exclusive(x_107); -if (x_108 == 0) -{ -lean_object* x_109; -x_109 = l_Except_orElseLazy___rarg(x_107, x_7); -return x_109; } -else -{ -lean_object* x_110; lean_object* x_111; lean_object* x_112; -x_110 = lean_ctor_get(x_107, 0); -lean_inc(x_110); -lean_dec(x_107); -x_111 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_111, 0, x_110); -x_112 = l_Except_orElseLazy___rarg(x_111, x_7); -return x_112; } } -else +static lean_object* _init_l___private_Lean_SubExpr_0__Lean_SubExpr_fromJsonGoalLocation____x40_Lean_SubExpr___hyg_1589____lambda__3___closed__1() { +_start: { -uint8_t x_113; -x_113 = !lean_is_exclusive(x_107); -if (x_113 == 0) +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("hypType", 7); +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Lean_SubExpr_0__Lean_SubExpr_fromJsonGoalLocation____x40_Lean_SubExpr___hyg_1589____lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: { -lean_object* x_114; lean_object* x_115; uint8_t x_116; -x_114 = lean_ctor_get(x_107, 0); -x_115 = l_Lean_SubExpr_instFromJsonFVarId___closed__1; -x_116 = lean_string_dec_eq(x_114, x_115); -if (x_116 == 0) +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_4 = l___private_Lean_SubExpr_0__Lean_SubExpr_fromJsonGoalLocation____x40_Lean_SubExpr___hyg_1589____lambda__3___closed__1; +x_5 = lean_unsigned_to_nat(2u); +x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_SubExpr_0__Lean_SubExpr_fromJsonGoalLocation____x40_Lean_SubExpr___hyg_1589____lambda__2___boxed), 3, 2); +lean_closure_set(x_7, 0, x_1); +lean_closure_set(x_7, 1, x_2); +if (lean_obj_tag(x_6) == 0) { -lean_object* x_117; lean_object* x_118; lean_object* x_119; -x_117 = l_Lean_SubExpr_instFromJsonFVarId___closed__2; -x_118 = lean_string_append(x_117, x_114); -lean_dec(x_114); -x_119 = l_Lean_Syntax_decodeNameLit(x_118); -if (lean_obj_tag(x_119) == 0) +uint8_t x_8; +x_8 = !lean_is_exclusive(x_6); +if (x_8 == 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_dec(x_14); -lean_dec(x_13); -x_120 = lean_unsigned_to_nat(80u); -x_121 = l_Lean_Json_pretty(x_106, x_120); -x_122 = l_Lean_SubExpr_instFromJsonFVarId___closed__3; -x_123 = lean_string_append(x_122, x_121); -lean_dec(x_121); -x_124 = l_Lean_SubExpr_instFromJsonFVarId___closed__4; -x_125 = lean_string_append(x_123, x_124); -lean_ctor_set_tag(x_107, 0); -lean_ctor_set(x_107, 0, x_125); -x_126 = l_Except_orElseLazy___rarg(x_107, x_7); -return x_126; -} -else -{ -lean_object* x_127; -lean_free_object(x_107); -lean_dec(x_106); -x_127 = lean_ctor_get(x_119, 0); -lean_inc(x_127); -lean_dec(x_119); -x_15 = x_127; -goto block_63; -} +lean_object* x_9; +x_9 = l_Except_orElseLazy___rarg(x_6, x_7); +return x_9; } else { -lean_object* x_128; -lean_free_object(x_107); -lean_dec(x_114); -lean_dec(x_106); -x_128 = lean_box(0); -x_15 = x_128; -goto block_63; +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = lean_ctor_get(x_6, 0); +lean_inc(x_10); +lean_dec(x_6); +x_11 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_11, 0, x_10); +x_12 = l_Except_orElseLazy___rarg(x_11, x_7); +return x_12; } } else { -lean_object* x_129; lean_object* x_130; uint8_t x_131; -x_129 = lean_ctor_get(x_107, 0); -lean_inc(x_129); -lean_dec(x_107); -x_130 = l_Lean_SubExpr_instFromJsonFVarId___closed__1; -x_131 = lean_string_dec_eq(x_129, x_130); -if (x_131 == 0) -{ -lean_object* x_132; lean_object* x_133; lean_object* x_134; -x_132 = l_Lean_SubExpr_instFromJsonFVarId___closed__2; -x_133 = lean_string_append(x_132, x_129); -lean_dec(x_129); -x_134 = l_Lean_Syntax_decodeNameLit(x_133); -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; lean_object* x_140; lean_object* x_141; lean_object* x_142; -lean_dec(x_14); -lean_dec(x_13); -x_135 = lean_unsigned_to_nat(80u); -x_136 = l_Lean_Json_pretty(x_106, x_135); -x_137 = l_Lean_SubExpr_instFromJsonFVarId___closed__3; -x_138 = lean_string_append(x_137, x_136); -lean_dec(x_136); -x_139 = l_Lean_SubExpr_instFromJsonFVarId___closed__4; -x_140 = lean_string_append(x_138, x_139); -x_141 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_141, 0, x_140); -x_142 = l_Except_orElseLazy___rarg(x_141, x_7); -return x_142; -} -else -{ -lean_object* x_143; -lean_dec(x_106); -x_143 = lean_ctor_get(x_134, 0); -lean_inc(x_143); -lean_dec(x_134); -x_15 = x_143; -goto block_63; -} +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_64; lean_object* x_100; uint8_t x_101; +x_13 = lean_ctor_get(x_6, 0); +lean_inc(x_13); +lean_dec(x_6); +x_14 = lean_array_get_size(x_13); +x_100 = lean_unsigned_to_nat(0u); +x_101 = lean_nat_dec_lt(x_100, x_14); +if (x_101 == 0) +{ +lean_object* x_102; lean_object* x_103; +x_102 = l_Lean_instInhabitedJson; +x_103 = l___private_Init_Util_0__outOfBounds___rarg(x_102); +x_64 = x_103; +goto block_99; } else { -lean_object* x_144; -lean_dec(x_129); -lean_dec(x_106); -x_144 = lean_box(0); -x_15 = x_144; -goto block_63; -} -} -} +lean_object* x_104; +x_104 = lean_array_fget(x_13, x_100); +x_64 = x_104; +goto block_99; } block_63: { @@ -3599,46 +3180,177 @@ return x_51; } else { -lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_52 = lean_ctor_get(x_49, 0); -lean_inc(x_52); -lean_dec(x_49); -x_53 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_53, 0, x_52); -x_54 = l_Except_orElseLazy___rarg(x_53, x_7); -return x_54; +lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_52 = lean_ctor_get(x_49, 0); +lean_inc(x_52); +lean_dec(x_49); +x_53 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_53, 0, x_52); +x_54 = l_Except_orElseLazy___rarg(x_53, x_7); +return x_54; +} +} +else +{ +uint8_t x_55; +x_55 = !lean_is_exclusive(x_49); +if (x_55 == 0) +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_56 = lean_ctor_get(x_49, 0); +x_57 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_57, 0, x_15); +lean_ctor_set(x_57, 1, x_56); +lean_ctor_set(x_49, 0, x_57); +x_58 = l_Except_orElseLazy___rarg(x_49, x_7); +return x_58; +} +else +{ +lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_59 = lean_ctor_get(x_49, 0); +lean_inc(x_59); +lean_dec(x_49); +x_60 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_60, 0, x_15); +lean_ctor_set(x_60, 1, x_59); +x_61 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_61, 0, x_60); +x_62 = l_Except_orElseLazy___rarg(x_61, x_7); +return x_62; +} +} +} +} +} +block_99: +{ +lean_object* x_65; +x_65 = l_Lean_Json_getStr_x3f(x_64); +if (lean_obj_tag(x_65) == 0) +{ +uint8_t x_66; +lean_dec(x_64); +lean_dec(x_14); +lean_dec(x_13); +x_66 = !lean_is_exclusive(x_65); +if (x_66 == 0) +{ +lean_object* x_67; +x_67 = l_Except_orElseLazy___rarg(x_65, x_7); +return x_67; +} +else +{ +lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_68 = lean_ctor_get(x_65, 0); +lean_inc(x_68); +lean_dec(x_65); +x_69 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_69, 0, x_68); +x_70 = l_Except_orElseLazy___rarg(x_69, x_7); +return x_70; +} +} +else +{ +uint8_t x_71; +x_71 = !lean_is_exclusive(x_65); +if (x_71 == 0) +{ +lean_object* x_72; lean_object* x_73; uint8_t x_74; +x_72 = lean_ctor_get(x_65, 0); +x_73 = l_Lean_SubExpr_instFromJsonFVarId___closed__1; +x_74 = lean_string_dec_eq(x_72, x_73); +if (x_74 == 0) +{ +lean_object* x_75; uint8_t x_76; +x_75 = l_String_toName(x_72); +x_76 = l_Lean_Name_isAnonymous(x_75); +if (x_76 == 0) +{ +lean_free_object(x_65); +lean_dec(x_64); +x_15 = x_75; +goto block_63; +} +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_dec(x_75); +lean_dec(x_14); +lean_dec(x_13); +x_77 = lean_unsigned_to_nat(80u); +x_78 = l_Lean_Json_pretty(x_64, x_77); +x_79 = l_Lean_SubExpr_instFromJsonFVarId___closed__2; +x_80 = lean_string_append(x_79, x_78); +lean_dec(x_78); +x_81 = l_Lean_SubExpr_instFromJsonFVarId___closed__3; +x_82 = lean_string_append(x_80, x_81); +lean_ctor_set_tag(x_65, 0); +lean_ctor_set(x_65, 0, x_82); +x_83 = l_Except_orElseLazy___rarg(x_65, x_7); +return x_83; +} +} +else +{ +lean_object* x_84; +lean_free_object(x_65); +lean_dec(x_72); +lean_dec(x_64); +x_84 = lean_box(0); +x_15 = x_84; +goto block_63; } } else { -uint8_t x_55; -x_55 = !lean_is_exclusive(x_49); -if (x_55 == 0) +lean_object* x_85; lean_object* x_86; uint8_t x_87; +x_85 = lean_ctor_get(x_65, 0); +lean_inc(x_85); +lean_dec(x_65); +x_86 = l_Lean_SubExpr_instFromJsonFVarId___closed__1; +x_87 = lean_string_dec_eq(x_85, x_86); +if (x_87 == 0) { -lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_56 = lean_ctor_get(x_49, 0); -x_57 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_57, 0, x_15); -lean_ctor_set(x_57, 1, x_56); -lean_ctor_set(x_49, 0, x_57); -x_58 = l_Except_orElseLazy___rarg(x_49, x_7); -return x_58; +lean_object* x_88; uint8_t x_89; +x_88 = l_String_toName(x_85); +x_89 = l_Lean_Name_isAnonymous(x_88); +if (x_89 == 0) +{ +lean_dec(x_64); +x_15 = x_88; +goto block_63; } else { -lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_59 = lean_ctor_get(x_49, 0); -lean_inc(x_59); -lean_dec(x_49); -x_60 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_60, 0, x_15); -lean_ctor_set(x_60, 1, x_59); -x_61 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_61, 0, x_60); -x_62 = l_Except_orElseLazy___rarg(x_61, x_7); -return x_62; +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_dec(x_88); +lean_dec(x_14); +lean_dec(x_13); +x_90 = lean_unsigned_to_nat(80u); +x_91 = l_Lean_Json_pretty(x_64, x_90); +x_92 = l_Lean_SubExpr_instFromJsonFVarId___closed__2; +x_93 = lean_string_append(x_92, x_91); +lean_dec(x_91); +x_94 = l_Lean_SubExpr_instFromJsonFVarId___closed__3; +x_95 = lean_string_append(x_93, x_94); +x_96 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_96, 0, x_95); +x_97 = l_Except_orElseLazy___rarg(x_96, x_7); +return x_97; } } +else +{ +lean_object* x_98; +lean_dec(x_85); +lean_dec(x_64); +x_98 = lean_box(0); +x_15 = x_98; +goto block_63; +} } } } @@ -3676,7 +3388,7 @@ return x_2; LEAN_EXPORT lean_object* l___private_Lean_SubExpr_0__Lean_SubExpr_fromJsonGoalLocation____x40_Lean_SubExpr___hyg_1589____lambda__4(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_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; x_4 = l___private_Lean_SubExpr_0__Lean_SubExpr_fromJsonGoalLocation____x40_Lean_SubExpr___hyg_1589____lambda__4___closed__1; x_5 = lean_unsigned_to_nat(1u); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); @@ -3685,320 +3397,183 @@ lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) { -uint8_t x_8; -x_8 = !lean_is_exclusive(x_6); -if (x_8 == 0) +uint8_t x_51; +x_51 = !lean_is_exclusive(x_6); +if (x_51 == 0) { -lean_object* x_9; -x_9 = l_Except_orElseLazy___rarg(x_6, x_7); -return x_9; +lean_object* x_52; +x_52 = l_Except_orElseLazy___rarg(x_6, x_7); +return x_52; } else { -lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = lean_ctor_get(x_6, 0); -lean_inc(x_10); +lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_53 = lean_ctor_get(x_6, 0); +lean_inc(x_53); lean_dec(x_6); -x_11 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_11, 0, x_10); -x_12 = l_Except_orElseLazy___rarg(x_11, x_7); -return x_12; +x_54 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_54, 0, x_53); +x_55 = l_Except_orElseLazy___rarg(x_54, x_7); +return x_55; } } else { -lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_13 = lean_ctor_get(x_6, 0); -lean_inc(x_13); +lean_object* x_56; lean_object* x_57; lean_object* x_58; uint8_t x_59; +x_56 = lean_ctor_get(x_6, 0); +lean_inc(x_56); lean_dec(x_6); -x_14 = lean_array_get_size(x_13); -x_15 = lean_unsigned_to_nat(0u); -x_16 = lean_nat_dec_lt(x_15, x_14); -lean_dec(x_14); -if (x_16 == 0) -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; -lean_dec(x_13); -x_17 = l_Lean_instInhabitedJson; -x_18 = l___private_Init_Util_0__outOfBounds___rarg(x_17); -x_19 = l_Lean_Json_getStr_x3f(x_18); -if (lean_obj_tag(x_19) == 0) -{ -uint8_t x_20; -lean_dec(x_18); -x_20 = !lean_is_exclusive(x_19); -if (x_20 == 0) -{ -lean_object* x_21; -x_21 = l_Except_orElseLazy___rarg(x_19, x_7); -return x_21; -} -else -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_19, 0); -lean_inc(x_22); -lean_dec(x_19); -x_23 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_23, 0, x_22); -x_24 = l_Except_orElseLazy___rarg(x_23, x_7); -return x_24; -} -} -else -{ -uint8_t x_25; -x_25 = !lean_is_exclusive(x_19); -if (x_25 == 0) -{ -lean_object* x_26; lean_object* x_27; uint8_t x_28; -x_26 = lean_ctor_get(x_19, 0); -x_27 = l_Lean_SubExpr_instFromJsonFVarId___closed__1; -x_28 = lean_string_dec_eq(x_26, x_27); -if (x_28 == 0) -{ -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = l_Lean_SubExpr_instFromJsonFVarId___closed__2; -x_30 = lean_string_append(x_29, x_26); -lean_dec(x_26); -x_31 = l_Lean_Syntax_decodeNameLit(x_30); -if (lean_obj_tag(x_31) == 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_unsigned_to_nat(80u); -x_33 = l_Lean_Json_pretty(x_18, x_32); -x_34 = l_Lean_SubExpr_instFromJsonFVarId___closed__3; -x_35 = lean_string_append(x_34, x_33); -lean_dec(x_33); -x_36 = l_Lean_SubExpr_instFromJsonFVarId___closed__4; -x_37 = lean_string_append(x_35, x_36); -lean_ctor_set_tag(x_19, 0); -lean_ctor_set(x_19, 0, x_37); -x_38 = l_Except_orElseLazy___rarg(x_19, x_7); -return x_38; -} -else +x_57 = lean_array_get_size(x_56); +x_58 = lean_unsigned_to_nat(0u); +x_59 = lean_nat_dec_lt(x_58, x_57); +lean_dec(x_57); +if (x_59 == 0) { -lean_object* x_39; lean_object* x_40; lean_object* x_41; -lean_dec(x_18); -x_39 = lean_ctor_get(x_31, 0); -lean_inc(x_39); -lean_dec(x_31); -x_40 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_40, 0, x_39); -lean_ctor_set(x_19, 0, x_40); -x_41 = l_Except_orElseLazy___rarg(x_19, x_7); -return x_41; -} +lean_object* x_60; lean_object* x_61; +lean_dec(x_56); +x_60 = l_Lean_instInhabitedJson; +x_61 = l___private_Init_Util_0__outOfBounds___rarg(x_60); +x_8 = x_61; +goto block_50; } else { -lean_object* x_42; lean_object* x_43; -lean_free_object(x_19); -lean_dec(x_26); -lean_dec(x_18); -x_42 = l___private_Lean_SubExpr_0__Lean_SubExpr_fromJsonGoalLocation____x40_Lean_SubExpr___hyg_1589____lambda__4___closed__3; -x_43 = l_Except_orElseLazy___rarg(x_42, x_7); -return x_43; +lean_object* x_62; +x_62 = lean_array_fget(x_56, x_58); +lean_dec(x_56); +x_8 = x_62; +goto block_50; } } -else +block_50: { -lean_object* x_44; lean_object* x_45; uint8_t x_46; -x_44 = lean_ctor_get(x_19, 0); -lean_inc(x_44); -lean_dec(x_19); -x_45 = l_Lean_SubExpr_instFromJsonFVarId___closed__1; -x_46 = lean_string_dec_eq(x_44, x_45); -if (x_46 == 0) -{ -lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_47 = l_Lean_SubExpr_instFromJsonFVarId___closed__2; -x_48 = lean_string_append(x_47, x_44); -lean_dec(x_44); -x_49 = l_Lean_Syntax_decodeNameLit(x_48); -if (lean_obj_tag(x_49) == 0) +lean_object* x_9; +x_9 = l_Lean_Json_getStr_x3f(x_8); +if (lean_obj_tag(x_9) == 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; -x_50 = lean_unsigned_to_nat(80u); -x_51 = l_Lean_Json_pretty(x_18, x_50); -x_52 = l_Lean_SubExpr_instFromJsonFVarId___closed__3; -x_53 = lean_string_append(x_52, x_51); -lean_dec(x_51); -x_54 = l_Lean_SubExpr_instFromJsonFVarId___closed__4; -x_55 = lean_string_append(x_53, x_54); -x_56 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_56, 0, x_55); -x_57 = l_Except_orElseLazy___rarg(x_56, x_7); -return x_57; -} -else +uint8_t x_10; +lean_dec(x_8); +x_10 = !lean_is_exclusive(x_9); +if (x_10 == 0) { -lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; -lean_dec(x_18); -x_58 = lean_ctor_get(x_49, 0); -lean_inc(x_58); -lean_dec(x_49); -x_59 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_59, 0, x_58); -x_60 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_60, 0, x_59); -x_61 = l_Except_orElseLazy___rarg(x_60, x_7); -return x_61; -} +lean_object* x_11; +x_11 = l_Except_orElseLazy___rarg(x_9, x_7); +return x_11; } else { -lean_object* x_62; lean_object* x_63; -lean_dec(x_44); -lean_dec(x_18); -x_62 = l___private_Lean_SubExpr_0__Lean_SubExpr_fromJsonGoalLocation____x40_Lean_SubExpr___hyg_1589____lambda__4___closed__3; -x_63 = l_Except_orElseLazy___rarg(x_62, x_7); -return x_63; -} -} -} +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_9, 0); +lean_inc(x_12); +lean_dec(x_9); +x_13 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_13, 0, x_12); +x_14 = l_Except_orElseLazy___rarg(x_13, x_7); +return x_14; } -else -{ -lean_object* x_64; lean_object* x_65; -x_64 = lean_array_fget(x_13, x_15); -lean_dec(x_13); -x_65 = l_Lean_Json_getStr_x3f(x_64); -if (lean_obj_tag(x_65) == 0) -{ -uint8_t x_66; -lean_dec(x_64); -x_66 = !lean_is_exclusive(x_65); -if (x_66 == 0) -{ -lean_object* x_67; -x_67 = l_Except_orElseLazy___rarg(x_65, x_7); -return x_67; } else { -lean_object* x_68; lean_object* x_69; lean_object* x_70; -x_68 = lean_ctor_get(x_65, 0); -lean_inc(x_68); -lean_dec(x_65); -x_69 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_69, 0, x_68); -x_70 = l_Except_orElseLazy___rarg(x_69, x_7); -return x_70; -} -} -else +uint8_t x_15; +x_15 = !lean_is_exclusive(x_9); +if (x_15 == 0) { -uint8_t x_71; -x_71 = !lean_is_exclusive(x_65); -if (x_71 == 0) +lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_16 = lean_ctor_get(x_9, 0); +x_17 = l_Lean_SubExpr_instFromJsonFVarId___closed__1; +x_18 = lean_string_dec_eq(x_16, x_17); +if (x_18 == 0) { -lean_object* x_72; lean_object* x_73; uint8_t x_74; -x_72 = lean_ctor_get(x_65, 0); -x_73 = l_Lean_SubExpr_instFromJsonFVarId___closed__1; -x_74 = lean_string_dec_eq(x_72, x_73); -if (x_74 == 0) +lean_object* x_19; uint8_t x_20; +x_19 = l_String_toName(x_16); +x_20 = l_Lean_Name_isAnonymous(x_19); +if (x_20 == 0) { -lean_object* x_75; lean_object* x_76; lean_object* x_77; -x_75 = l_Lean_SubExpr_instFromJsonFVarId___closed__2; -x_76 = lean_string_append(x_75, x_72); -lean_dec(x_72); -x_77 = l_Lean_Syntax_decodeNameLit(x_76); -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; lean_object* x_83; lean_object* x_84; -x_78 = lean_unsigned_to_nat(80u); -x_79 = l_Lean_Json_pretty(x_64, x_78); -x_80 = l_Lean_SubExpr_instFromJsonFVarId___closed__3; -x_81 = lean_string_append(x_80, x_79); -lean_dec(x_79); -x_82 = l_Lean_SubExpr_instFromJsonFVarId___closed__4; -x_83 = lean_string_append(x_81, x_82); -lean_ctor_set_tag(x_65, 0); -lean_ctor_set(x_65, 0, x_83); -x_84 = l_Except_orElseLazy___rarg(x_65, x_7); -return x_84; +lean_object* x_21; lean_object* x_22; +lean_dec(x_8); +x_21 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_9, 0, x_21); +x_22 = l_Except_orElseLazy___rarg(x_9, x_7); +return x_22; } else { -lean_object* x_85; lean_object* x_86; lean_object* x_87; -lean_dec(x_64); -x_85 = lean_ctor_get(x_77, 0); -lean_inc(x_85); -lean_dec(x_77); -x_86 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_86, 0, x_85); -lean_ctor_set(x_65, 0, x_86); -x_87 = l_Except_orElseLazy___rarg(x_65, x_7); -return x_87; +lean_object* x_23; lean_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_19); +x_23 = lean_unsigned_to_nat(80u); +x_24 = l_Lean_Json_pretty(x_8, x_23); +x_25 = l_Lean_SubExpr_instFromJsonFVarId___closed__2; +x_26 = lean_string_append(x_25, x_24); +lean_dec(x_24); +x_27 = l_Lean_SubExpr_instFromJsonFVarId___closed__3; +x_28 = lean_string_append(x_26, x_27); +lean_ctor_set_tag(x_9, 0); +lean_ctor_set(x_9, 0, x_28); +x_29 = l_Except_orElseLazy___rarg(x_9, x_7); +return x_29; } } else { -lean_object* x_88; lean_object* x_89; -lean_free_object(x_65); -lean_dec(x_72); -lean_dec(x_64); -x_88 = l___private_Lean_SubExpr_0__Lean_SubExpr_fromJsonGoalLocation____x40_Lean_SubExpr___hyg_1589____lambda__4___closed__3; -x_89 = l_Except_orElseLazy___rarg(x_88, x_7); -return x_89; +lean_object* x_30; lean_object* x_31; +lean_free_object(x_9); +lean_dec(x_16); +lean_dec(x_8); +x_30 = l___private_Lean_SubExpr_0__Lean_SubExpr_fromJsonGoalLocation____x40_Lean_SubExpr___hyg_1589____lambda__4___closed__3; +x_31 = l_Except_orElseLazy___rarg(x_30, x_7); +return x_31; } } else { -lean_object* x_90; lean_object* x_91; uint8_t x_92; -x_90 = lean_ctor_get(x_65, 0); -lean_inc(x_90); -lean_dec(x_65); -x_91 = l_Lean_SubExpr_instFromJsonFVarId___closed__1; -x_92 = lean_string_dec_eq(x_90, x_91); -if (x_92 == 0) +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_33 = l_Lean_SubExpr_instFromJsonFVarId___closed__1; +x_34 = lean_string_dec_eq(x_32, x_33); +if (x_34 == 0) { -lean_object* x_93; lean_object* x_94; lean_object* x_95; -x_93 = l_Lean_SubExpr_instFromJsonFVarId___closed__2; -x_94 = lean_string_append(x_93, x_90); -lean_dec(x_90); -x_95 = l_Lean_Syntax_decodeNameLit(x_94); -if (lean_obj_tag(x_95) == 0) +lean_object* x_35; uint8_t x_36; +x_35 = l_String_toName(x_32); +x_36 = l_Lean_Name_isAnonymous(x_35); +if (x_36 == 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; -x_96 = lean_unsigned_to_nat(80u); -x_97 = l_Lean_Json_pretty(x_64, x_96); -x_98 = l_Lean_SubExpr_instFromJsonFVarId___closed__3; -x_99 = lean_string_append(x_98, x_97); -lean_dec(x_97); -x_100 = l_Lean_SubExpr_instFromJsonFVarId___closed__4; -x_101 = lean_string_append(x_99, x_100); -x_102 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_102, 0, x_101); -x_103 = l_Except_orElseLazy___rarg(x_102, x_7); -return x_103; +lean_object* x_37; lean_object* x_38; lean_object* x_39; +lean_dec(x_8); +x_37 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_37, 0, x_35); +x_38 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_38, 0, x_37); +x_39 = l_Except_orElseLazy___rarg(x_38, x_7); +return x_39; } else { -lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; -lean_dec(x_64); -x_104 = lean_ctor_get(x_95, 0); -lean_inc(x_104); -lean_dec(x_95); -x_105 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_105, 0, x_104); -x_106 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_106, 0, x_105); -x_107 = l_Except_orElseLazy___rarg(x_106, x_7); -return x_107; +lean_object* x_40; lean_object* x_41; 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_35); +x_40 = lean_unsigned_to_nat(80u); +x_41 = l_Lean_Json_pretty(x_8, x_40); +x_42 = l_Lean_SubExpr_instFromJsonFVarId___closed__2; +x_43 = lean_string_append(x_42, x_41); +lean_dec(x_41); +x_44 = l_Lean_SubExpr_instFromJsonFVarId___closed__3; +x_45 = lean_string_append(x_43, x_44); +x_46 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_46, 0, x_45); +x_47 = l_Except_orElseLazy___rarg(x_46, x_7); +return x_47; } } else { -lean_object* x_108; lean_object* x_109; -lean_dec(x_90); -lean_dec(x_64); -x_108 = l___private_Lean_SubExpr_0__Lean_SubExpr_fromJsonGoalLocation____x40_Lean_SubExpr___hyg_1589____lambda__4___closed__3; -x_109 = l_Except_orElseLazy___rarg(x_108, x_7); -return x_109; -} +lean_object* x_48; lean_object* x_49; +lean_dec(x_32); +lean_dec(x_8); +x_48 = l___private_Lean_SubExpr_0__Lean_SubExpr_fromJsonGoalLocation____x40_Lean_SubExpr___hyg_1589____lambda__4___closed__3; +x_49 = l_Except_orElseLazy___rarg(x_48, x_7); +return x_49; } } } @@ -4463,94 +4038,85 @@ x_10 = l_Lean_SubExpr_instFromJsonFVarId___closed__1; x_11 = lean_string_dec_eq(x_9, x_10); if (x_11 == 0) { -lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_12 = l_Lean_SubExpr_instFromJsonFVarId___closed__2; -x_13 = lean_string_append(x_12, x_9); -lean_dec(x_9); -x_14 = l_Lean_Syntax_decodeNameLit(x_13); -if (lean_obj_tag(x_14) == 0) +lean_object* x_12; uint8_t x_13; +x_12 = l_String_toName(x_9); +x_13 = l_Lean_Name_isAnonymous(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; -x_15 = lean_unsigned_to_nat(80u); -x_16 = l_Lean_Json_pretty(x_3, x_15); -x_17 = l_Lean_SubExpr_instFromJsonFVarId___closed__3; -x_18 = lean_string_append(x_17, x_16); -lean_dec(x_16); -x_19 = l_Lean_SubExpr_instFromJsonFVarId___closed__4; -x_20 = lean_string_append(x_18, x_19); -lean_ctor_set_tag(x_4, 0); -lean_ctor_set(x_4, 0, x_20); +lean_dec(x_3); +lean_ctor_set(x_4, 0, x_12); return x_4; } else { -lean_object* x_21; -lean_dec(x_3); -x_21 = lean_ctor_get(x_14, 0); -lean_inc(x_21); -lean_dec(x_14); -lean_ctor_set(x_4, 0, x_21); +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_12); +x_14 = lean_unsigned_to_nat(80u); +x_15 = l_Lean_Json_pretty(x_3, x_14); +x_16 = l_Lean_SubExpr_instFromJsonFVarId___closed__2; +x_17 = lean_string_append(x_16, x_15); +lean_dec(x_15); +x_18 = l_Lean_SubExpr_instFromJsonFVarId___closed__3; +x_19 = lean_string_append(x_17, x_18); +lean_ctor_set_tag(x_4, 0); +lean_ctor_set(x_4, 0, x_19); return x_4; } } else { -lean_object* x_22; +lean_object* x_20; lean_free_object(x_4); lean_dec(x_9); lean_dec(x_3); -x_22 = l_Lean_SubExpr_instFromJsonFVarId___closed__5; -return x_22; +x_20 = l_Lean_SubExpr_instFromJsonFVarId___closed__4; +return x_20; } } else { -lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_23 = lean_ctor_get(x_4, 0); -lean_inc(x_23); +lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_21 = lean_ctor_get(x_4, 0); +lean_inc(x_21); lean_dec(x_4); -x_24 = l_Lean_SubExpr_instFromJsonFVarId___closed__1; -x_25 = lean_string_dec_eq(x_23, x_24); -if (x_25 == 0) +x_22 = l_Lean_SubExpr_instFromJsonFVarId___closed__1; +x_23 = lean_string_dec_eq(x_21, x_22); +if (x_23 == 0) { -lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_26 = l_Lean_SubExpr_instFromJsonFVarId___closed__2; -x_27 = lean_string_append(x_26, x_23); -lean_dec(x_23); -x_28 = l_Lean_Syntax_decodeNameLit(x_27); -if (lean_obj_tag(x_28) == 0) +lean_object* x_24; uint8_t x_25; +x_24 = l_String_toName(x_21); +x_25 = l_Lean_Name_isAnonymous(x_24); +if (x_25 == 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; -x_29 = lean_unsigned_to_nat(80u); -x_30 = l_Lean_Json_pretty(x_3, x_29); -x_31 = l_Lean_SubExpr_instFromJsonFVarId___closed__3; -x_32 = lean_string_append(x_31, x_30); -lean_dec(x_30); -x_33 = l_Lean_SubExpr_instFromJsonFVarId___closed__4; -x_34 = lean_string_append(x_32, x_33); -x_35 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_35, 0, x_34); -return x_35; +lean_object* x_26; +lean_dec(x_3); +x_26 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_26, 0, x_24); +return x_26; } else { -lean_object* x_36; lean_object* x_37; -lean_dec(x_3); -x_36 = lean_ctor_get(x_28, 0); -lean_inc(x_36); +lean_object* x_27; lean_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_24); +x_27 = lean_unsigned_to_nat(80u); +x_28 = l_Lean_Json_pretty(x_3, x_27); +x_29 = l_Lean_SubExpr_instFromJsonFVarId___closed__2; +x_30 = lean_string_append(x_29, x_28); lean_dec(x_28); -x_37 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_37, 0, x_36); -return x_37; +x_31 = l_Lean_SubExpr_instFromJsonFVarId___closed__3; +x_32 = lean_string_append(x_30, x_31); +x_33 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_33, 0, x_32); +return x_33; } } else { -lean_object* x_38; -lean_dec(x_23); +lean_object* x_34; +lean_dec(x_21); lean_dec(x_3); -x_38 = l_Lean_SubExpr_instFromJsonFVarId___closed__5; -return x_38; +x_34 = l_Lean_SubExpr_instFromJsonFVarId___closed__4; +return x_34; } } } @@ -5135,8 +4701,6 @@ l_Lean_SubExpr_instFromJsonFVarId___closed__3 = _init_l_Lean_SubExpr_instFromJso lean_mark_persistent(l_Lean_SubExpr_instFromJsonFVarId___closed__3); l_Lean_SubExpr_instFromJsonFVarId___closed__4 = _init_l_Lean_SubExpr_instFromJsonFVarId___closed__4(); lean_mark_persistent(l_Lean_SubExpr_instFromJsonFVarId___closed__4); -l_Lean_SubExpr_instFromJsonFVarId___closed__5 = _init_l_Lean_SubExpr_instFromJsonFVarId___closed__5(); -lean_mark_persistent(l_Lean_SubExpr_instFromJsonFVarId___closed__5); l___private_Lean_SubExpr_0__Lean_SubExpr_fromJsonGoalLocation____x40_Lean_SubExpr___hyg_1589____lambda__1___closed__1 = _init_l___private_Lean_SubExpr_0__Lean_SubExpr_fromJsonGoalLocation____x40_Lean_SubExpr___hyg_1589____lambda__1___closed__1(); lean_mark_persistent(l___private_Lean_SubExpr_0__Lean_SubExpr_fromJsonGoalLocation____x40_Lean_SubExpr___hyg_1589____lambda__1___closed__1); l___private_Lean_SubExpr_0__Lean_SubExpr_fromJsonGoalLocation____x40_Lean_SubExpr___hyg_1589____lambda__1___closed__2 = _init_l___private_Lean_SubExpr_0__Lean_SubExpr_fromJsonGoalLocation____x40_Lean_SubExpr___hyg_1589____lambda__1___closed__2(); diff --git a/stage0/stdlib/Lean/Util/PPExt.c b/stage0/stdlib/Lean/Util/PPExt.c index db8db3a2eb9f..4fa3d7c2b813 100644 --- a/stage0/stdlib/Lean/Util/PPExt.c +++ b/stage0/stdlib/Lean/Util/PPExt.c @@ -15,9 +15,7 @@ extern "C" { #endif static lean_object* l_Lean_PPContext_lctx___default___closed__1; LEAN_EXPORT lean_object* l_Lean_instInhabitedPPFns___lambda__1(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_Lean_initFn____x40_Lean_Util_PPExt___hyg_90_(lean_object*); -lean_object* l_Lean_Option_register___at_Std_Format_initFn____x40_Lean_Data_Format___hyg_30____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_ppExprWithInfos___closed__2; static lean_object* l_Lean_ppTerm___closed__5; LEAN_EXPORT lean_object* l_Lean_PPContext_openDecls___default; @@ -33,7 +31,6 @@ static lean_object* l_Lean_initFn____x40_Lean_Util_PPExt___hyg_6____closed__3; static lean_object* l_Lean_initFn____x40_Lean_Util_PPExt___hyg_90____closed__4; lean_object* l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg(lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Util_PPExt___hyg_90____closed__5; -lean_object* l_Lean_Option_get___at_Std_Format_pretty_x27___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ppExt; static uint32_t l_Lean_instInhabitedPPFns___lambda__1___closed__1; static lean_object* l_Lean_PPContext_mctx___default___closed__6; @@ -64,6 +61,7 @@ static lean_object* l_Lean_initFn____x40_Lean_Util_PPExt___hyg_452____closed__1; LEAN_EXPORT lean_object* l_Lean_ppExprWithInfos(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PPContext_mctx___default; static lean_object* l_Lean_ppTerm___closed__4; +lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instInhabitedPPFns; LEAN_EXPORT lean_object* l_Lean_PPContext_opts___default; lean_object* lean_st_ref_get(lean_object*, lean_object*); @@ -77,11 +75,11 @@ LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Util_PPExt___hyg_452____lambd LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Util_PPExt___hyg_452____lambda__3(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Util_PPExt___hyg_6____closed__4; static lean_object* l_Lean_initFn____x40_Lean_Util_PPExt___hyg_130____closed__1; +uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_instInhabitedPPFns___lambda__1___closed__2; static lean_object* l_Lean_initFn____x40_Lean_Util_PPExt___hyg_452____closed__4; static lean_object* l_Lean_ppExprWithInfos___closed__5; LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Util_PPExt___hyg_557____lambda__1(lean_object*); -lean_object* l_Lean_Option_register___at_Std_Format_initFn____x40_Lean_Data_Format___hyg_95____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Util_PPExt___hyg_452____closed__2; static lean_object* l_Lean_ppExprWithInfos___closed__1; LEAN_EXPORT lean_object* l_Lean_PPContext_lctx___default; @@ -101,6 +99,7 @@ static lean_object* l_Lean_initFn____x40_Lean_Util_PPExt___hyg_50____closed__2; static lean_object* l_Lean_ppExprWithInfos___closed__9; static lean_object* l_Lean_PPContext_lctx___default___closed__3; static lean_object* l_Lean_ppTerm___closed__2; +lean_object* l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); uint32_t lean_uint32_of_nat(lean_object*); static lean_object* l_Lean_ppExprWithInfos___closed__4; lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); @@ -121,6 +120,7 @@ static lean_object* l_Lean_initFn____x40_Lean_Util_PPExt___hyg_452____lambda__2_ static lean_object* l_Lean_initFn____x40_Lean_Util_PPExt___hyg_6____closed__5; lean_object* lean_expr_dbg_to_string(lean_object*); static lean_object* l_Lean_PPContext_mctx___default___closed__2; +lean_object* l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_68____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Util_PPExt___hyg_50____closed__5; LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Util_PPExt___hyg_50_(lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Util_PPExt___hyg_50____closed__1; @@ -200,7 +200,7 @@ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_2 = l_Lean_initFn____x40_Lean_Util_PPExt___hyg_6____closed__3; x_3 = l_Lean_initFn____x40_Lean_Util_PPExt___hyg_6____closed__5; x_4 = l_Lean_initFn____x40_Lean_Util_PPExt___hyg_6____closed__7; -x_5 = l_Lean_Option_register___at_Std_Format_initFn____x40_Lean_Data_Format___hyg_95____spec__1(x_2, x_3, x_4, x_1); +x_5 = l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__1(x_2, x_3, x_4, x_1); return x_5; } } @@ -265,7 +265,7 @@ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_2 = l_Lean_initFn____x40_Lean_Util_PPExt___hyg_50____closed__2; x_3 = l_Lean_initFn____x40_Lean_Util_PPExt___hyg_50____closed__4; x_4 = l_Lean_initFn____x40_Lean_Util_PPExt___hyg_50____closed__5; -x_5 = l_Lean_Option_register___at_Std_Format_initFn____x40_Lean_Data_Format___hyg_95____spec__1(x_2, x_3, x_4, x_1); +x_5 = l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__1(x_2, x_3, x_4, x_1); return x_5; } } @@ -329,7 +329,7 @@ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_2 = l_Lean_initFn____x40_Lean_Util_PPExt___hyg_90____closed__2; x_3 = l_Lean_initFn____x40_Lean_Util_PPExt___hyg_90____closed__4; x_4 = l_Lean_initFn____x40_Lean_Util_PPExt___hyg_90____closed__5; -x_5 = l_Lean_Option_register___at_Std_Format_initFn____x40_Lean_Data_Format___hyg_30____spec__1(x_2, x_3, x_4, x_1); +x_5 = l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_68____spec__1(x_2, x_3, x_4, x_1); return x_5; } } @@ -392,7 +392,7 @@ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_2 = l_Lean_initFn____x40_Lean_Util_PPExt___hyg_130____closed__2; x_3 = l_Lean_initFn____x40_Lean_Util_PPExt___hyg_130____closed__4; x_4 = l_Lean_initFn____x40_Lean_Util_PPExt___hyg_130____closed__5; -x_5 = l_Lean_Option_register___at_Std_Format_initFn____x40_Lean_Data_Format___hyg_95____spec__1(x_2, x_3, x_4, x_1); +x_5 = l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__1(x_2, x_3, x_4, x_1); return x_5; } } @@ -693,7 +693,7 @@ LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Util_PPExt___hyg_452____lambd 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; x_4 = lean_ctor_get(x_1, 3); x_5 = l_Lean_initFn____x40_Lean_Util_PPExt___hyg_452____lambda__2___closed__1; -x_6 = l_Lean_Option_get___at_Std_Format_pretty_x27___spec__1(x_4, x_5); +x_6 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_4, x_5); x_7 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_7, 0, x_6); x_8 = 0; @@ -992,7 +992,7 @@ lean_dec(x_5); x_7 = lean_ctor_get(x_1, 3); lean_inc(x_7); x_8 = l_Lean_ppExprWithInfos___closed__1; -x_9 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_7, x_8); +x_9 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_7, 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; @@ -1022,7 +1022,7 @@ if (x_16 == 0) lean_object* x_17; lean_object* x_18; uint8_t x_19; x_17 = lean_ctor_get(x_15, 0); x_18 = l_Lean_ppExprWithInfos___closed__3; -x_19 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_7, x_18); +x_19 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_7, x_18); lean_dec(x_7); if (x_19 == 0) { @@ -1084,7 +1084,7 @@ lean_inc(x_38); lean_inc(x_37); lean_dec(x_15); x_39 = l_Lean_ppExprWithInfos___closed__3; -x_40 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_7, x_39); +x_40 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_7, x_39); lean_dec(x_7); if (x_40 == 0) { @@ -1212,7 +1212,7 @@ lean_object* x_4; lean_object* x_5; uint8_t x_6; x_4 = lean_ctor_get(x_1, 3); lean_inc(x_4); x_5 = l_Lean_ppExprWithInfos___closed__1; -x_6 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_4, x_5); +x_6 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_4, x_5); 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; @@ -1242,7 +1242,7 @@ if (x_13 == 0) lean_object* x_14; lean_object* x_15; uint8_t x_16; x_14 = lean_ctor_get(x_12, 0); x_15 = l_Lean_ppExprWithInfos___closed__3; -x_16 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_4, x_15); +x_16 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_4, x_15); if (x_16 == 0) { lean_object* x_17; @@ -1277,11 +1277,11 @@ x_27 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_27, 0, x_25); lean_ctor_set(x_27, 1, x_26); x_28 = l_Lean_initFn____x40_Lean_Util_PPExt___hyg_452____lambda__2___closed__1; -x_29 = l_Lean_Option_get___at_Std_Format_pretty_x27___spec__1(x_4, x_28); +x_29 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_4, x_28); x_30 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_30, 0, x_29); x_31 = l_Lean_ppTerm___closed__5; -x_32 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_4, x_31); +x_32 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_4, x_31); lean_dec(x_4); x_33 = lean_unsigned_to_nat(0u); x_34 = l_Lean_Syntax_formatStxAux(x_30, x_32, x_33, x_2); @@ -1306,7 +1306,7 @@ lean_inc(x_38); lean_inc(x_37); lean_dec(x_12); x_39 = l_Lean_ppExprWithInfos___closed__3; -x_40 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_4, x_39); +x_40 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_4, x_39); if (x_40 == 0) { lean_object* x_41; lean_object* x_42; @@ -1342,11 +1342,11 @@ x_52 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_52, 0, x_50); lean_ctor_set(x_52, 1, x_51); x_53 = l_Lean_initFn____x40_Lean_Util_PPExt___hyg_452____lambda__2___closed__1; -x_54 = l_Lean_Option_get___at_Std_Format_pretty_x27___spec__1(x_4, x_53); +x_54 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_4, x_53); x_55 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_55, 0, x_54); x_56 = l_Lean_ppTerm___closed__5; -x_57 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_4, x_56); +x_57 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_4, x_56); lean_dec(x_4); x_58 = lean_unsigned_to_nat(0u); x_59 = l_Lean_Syntax_formatStxAux(x_55, x_57, x_58, x_2); @@ -1370,11 +1370,11 @@ else lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_dec(x_1); x_63 = l_Lean_initFn____x40_Lean_Util_PPExt___hyg_452____lambda__2___closed__1; -x_64 = l_Lean_Option_get___at_Std_Format_pretty_x27___spec__1(x_4, x_63); +x_64 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_4, x_63); x_65 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_65, 0, x_64); x_66 = l_Lean_ppTerm___closed__5; -x_67 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_4, x_66); +x_67 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_4, x_66); lean_dec(x_4); x_68 = lean_unsigned_to_nat(0u); x_69 = l_Lean_Syntax_formatStxAux(x_65, x_67, x_68, x_2); diff --git a/stage0/stdlib/Lean/Util/Profile.c b/stage0/stdlib/Lean/Util/Profile.c index e2444bce4faa..7f523b5db263 100644 --- a/stage0/stdlib/Lean/Util/Profile.c +++ b/stage0/stdlib/Lean/Util/Profile.c @@ -14,16 +14,518 @@ extern "C" { #endif lean_object* lean_profileit(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_profiler; +static lean_object* l_Lean_initFn____x40_Lean_Util_Profile___hyg_6____closed__2; LEAN_EXPORT lean_object* l_Lean_profileitIOUnsafe___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +double lean_float_div(double, double); +static lean_object* l_Lean_profiler_threshold_getSecs___closed__1; +static lean_object* l_Lean_initFn____x40_Lean_Util_Profile___hyg_6____closed__4; LEAN_EXPORT lean_object* l_Lean_profileitIOUnsafe(lean_object*, lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Util_Profile___hyg_6____closed__6; +lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Util_Profile___hyg_68____closed__2; +static double l_Lean_profiler_threshold_getSecs___closed__2; +LEAN_EXPORT lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT uint8_t lean_get_profiler(lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Util_Profile___hyg_6____closed__3; +lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_profileitM___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Util_Profile___hyg_6_(lean_object*); LEAN_EXPORT lean_object* l_Lean_profileitIOUnsafe___rarg___lambda__1(lean_object*, lean_object*); +lean_object* lean_register_option(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Util_Profile___hyg_68_(lean_object*); +double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); +LEAN_EXPORT double lean_get_profiler_threshold(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Util_Profile___hyg_68____closed__1; +lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Util_Profile___hyg_68____closed__3; LEAN_EXPORT lean_object* l_Lean_profileitM(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_profileitM___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_profileitIOUnsafe___rarg___lambda__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_68____spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Util_Profile_0__Lean_get__profiler___boxed(lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Util_Profile___hyg_68____closed__5; +static lean_object* l___private_Lean_Util_Profile_0__Lean_get__profiler___closed__1; LEAN_EXPORT lean_object* l_Lean_profileit___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_KVMap_findCore(lean_object*, lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Util_Profile___hyg_6____closed__5; +static lean_object* l_Lean_initFn____x40_Lean_Util_Profile___hyg_68____closed__4; LEAN_EXPORT lean_object* l_Lean_profileitIOUnsafe___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_profileitM___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_68____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_profiler_threshold; +LEAN_EXPORT lean_object* l_Lean_profiler_threshold_getSecs___boxed(lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Util_Profile___hyg_6____closed__1; +LEAN_EXPORT lean_object* l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__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; uint8_t x_9; lean_object* x_10; lean_object* x_11; +x_5 = lean_ctor_get(x_2, 0); +x_6 = lean_ctor_get(x_2, 1); +x_7 = lean_ctor_get(x_2, 2); +x_8 = lean_alloc_ctor(1, 0, 1); +x_9 = lean_unbox(x_5); +lean_ctor_set_uint8(x_8, 0, x_9); +lean_inc(x_7); +lean_inc(x_6); +x_10 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_10, 0, x_3); +lean_ctor_set(x_10, 1, x_8); +lean_ctor_set(x_10, 2, x_6); +lean_ctor_set(x_10, 3, x_7); +lean_inc(x_1); +x_11 = lean_register_option(x_1, x_10, x_4); +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; +x_13 = lean_ctor_get(x_11, 0); +lean_dec(x_13); +lean_inc(x_5); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_1); +lean_ctor_set(x_14, 1, x_5); +lean_ctor_set(x_11, 0, x_14); +return x_11; +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_11, 1); +lean_inc(x_15); +lean_dec(x_11); +lean_inc(x_5); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_1); +lean_ctor_set(x_16, 1, x_5); +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; +} +} +else +{ +uint8_t x_18; +lean_dec(x_1); +x_18 = !lean_is_exclusive(x_11); +if (x_18 == 0) +{ +return x_11; +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_11, 0); +x_20 = lean_ctor_get(x_11, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_11); +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; +} +} +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Util_Profile___hyg_6____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("profiler", 8); +return x_1; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Util_Profile___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_initFn____x40_Lean_Util_Profile___hyg_6____closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Util_Profile___hyg_6____closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("show execution times of various Lean components", 47); +return x_1; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Util_Profile___hyg_6____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_initFn____x40_Lean_Util_Profile___hyg_6____closed__1; +x_3 = l_Lean_initFn____x40_Lean_Util_Profile___hyg_6____closed__3; +x_4 = lean_box(x_1); +x_5 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_5, 0, x_4); +lean_ctor_set(x_5, 1, x_2); +lean_ctor_set(x_5, 2, x_3); +return x_5; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Util_Profile___hyg_6____closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Lean", 4); +return x_1; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Util_Profile___hyg_6____closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_initFn____x40_Lean_Util_Profile___hyg_6____closed__5; +x_2 = l_Lean_initFn____x40_Lean_Util_Profile___hyg_6____closed__1; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Util_Profile___hyg_6_(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_initFn____x40_Lean_Util_Profile___hyg_6____closed__2; +x_3 = l_Lean_initFn____x40_Lean_Util_Profile___hyg_6____closed__4; +x_4 = l_Lean_initFn____x40_Lean_Util_Profile___hyg_6____closed__6; +x_5 = l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__1(x_2, x_3, x_4, x_1); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__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_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__1(x_1, x_2, x_3, x_4); +lean_dec(x_2); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_68____spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_5 = lean_ctor_get(x_2, 0); +x_6 = lean_ctor_get(x_2, 1); +x_7 = lean_ctor_get(x_2, 2); +lean_inc(x_5); +x_8 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_8, 0, x_5); +lean_inc(x_7); +lean_inc(x_6); +x_9 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_9, 0, x_3); +lean_ctor_set(x_9, 1, x_8); +lean_ctor_set(x_9, 2, x_6); +lean_ctor_set(x_9, 3, x_7); +lean_inc(x_1); +x_10 = lean_register_option(x_1, x_9, x_4); +if (lean_obj_tag(x_10) == 0) +{ +uint8_t x_11; +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; +x_12 = lean_ctor_get(x_10, 0); +lean_dec(x_12); +lean_inc(x_5); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_1); +lean_ctor_set(x_13, 1, x_5); +lean_ctor_set(x_10, 0, x_13); +return x_10; +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_10, 1); +lean_inc(x_14); +lean_dec(x_10); +lean_inc(x_5); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_1); +lean_ctor_set(x_15, 1, x_5); +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 +{ +uint8_t x_17; +lean_dec(x_1); +x_17 = !lean_is_exclusive(x_10); +if (x_17 == 0) +{ +return x_10; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_10, 0); +x_19 = lean_ctor_get(x_10, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_10); +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; +} +} +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Util_Profile___hyg_68____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("threshold", 9); +return x_1; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Util_Profile___hyg_68____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_initFn____x40_Lean_Util_Profile___hyg_6____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Util_Profile___hyg_68____closed__1; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Util_Profile___hyg_68____closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("threshold in milliseconds, profiling times under threshold will not be reported individually", 92); +return x_1; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Util_Profile___hyg_68____closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = lean_unsigned_to_nat(100u); +x_2 = l_Lean_initFn____x40_Lean_Util_Profile___hyg_6____closed__1; +x_3 = l_Lean_initFn____x40_Lean_Util_Profile___hyg_68____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_initFn____x40_Lean_Util_Profile___hyg_68____closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_initFn____x40_Lean_Util_Profile___hyg_6____closed__5; +x_2 = l_Lean_initFn____x40_Lean_Util_Profile___hyg_6____closed__1; +x_3 = l_Lean_initFn____x40_Lean_Util_Profile___hyg_68____closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Util_Profile___hyg_68_(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_initFn____x40_Lean_Util_Profile___hyg_68____closed__2; +x_3 = l_Lean_initFn____x40_Lean_Util_Profile___hyg_68____closed__4; +x_4 = l_Lean_initFn____x40_Lean_Util_Profile___hyg_68____closed__5; +x_5 = l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_68____spec__1(x_2, x_3, x_4, x_1); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_68____spec__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_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_68____spec__1(x_1, x_2, x_3, x_4); +lean_dec(x_2); +return x_5; +} +} +LEAN_EXPORT uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(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_2, 0); +x_4 = lean_ctor_get(x_2, 1); +x_5 = l_Lean_KVMap_findCore(x_1, x_3); +if (lean_obj_tag(x_5) == 0) +{ +uint8_t x_6; +x_6 = lean_unbox(x_4); +return x_6; +} +else +{ +lean_object* x_7; +x_7 = lean_ctor_get(x_5, 0); +lean_inc(x_7); +lean_dec(x_5); +if (lean_obj_tag(x_7) == 1) +{ +uint8_t x_8; +x_8 = lean_ctor_get_uint8(x_7, 0); +lean_dec(x_7); +return x_8; +} +else +{ +uint8_t x_9; +lean_dec(x_7); +x_9 = lean_unbox(x_4); +return x_9; +} +} +} +} +static lean_object* _init_l___private_Lean_Util_Profile_0__Lean_get__profiler___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_profiler; +return x_1; +} +} +LEAN_EXPORT uint8_t lean_get_profiler(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; +x_2 = l___private_Lean_Util_Profile_0__Lean_get__profiler___closed__1; +x_3 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(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___private_Lean_Util_Profile_0__Lean_get__profiler___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_get_profiler(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(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_2, 0); +x_4 = lean_ctor_get(x_2, 1); +x_5 = l_Lean_KVMap_findCore(x_1, x_3); +if (lean_obj_tag(x_5) == 0) +{ +lean_inc(x_4); +return x_4; +} +else +{ +lean_object* x_6; +x_6 = lean_ctor_get(x_5, 0); +lean_inc(x_6); +lean_dec(x_5); +if (lean_obj_tag(x_6) == 3) +{ +lean_object* x_7; +x_7 = lean_ctor_get(x_6, 0); +lean_inc(x_7); +lean_dec(x_6); +return x_7; +} +else +{ +lean_dec(x_6); +lean_inc(x_4); +return x_4; +} +} +} +} +static lean_object* _init_l_Lean_profiler_threshold_getSecs___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_profiler_threshold; +return x_1; +} +} +static double _init_l_Lean_profiler_threshold_getSecs___closed__2() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; double x_4; +x_1 = lean_unsigned_to_nat(1000u); +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 double lean_get_profiler_threshold(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; uint8_t x_4; lean_object* x_5; double x_6; double x_7; double x_8; +x_2 = l_Lean_profiler_threshold_getSecs___closed__1; +x_3 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_2); +lean_dec(x_1); +x_4 = 0; +x_5 = lean_unsigned_to_nat(0u); +x_6 = l_Float_ofScientific(x_3, x_4, x_5); +lean_dec(x_3); +x_7 = l_Lean_profiler_threshold_getSecs___closed__2; +x_8 = lean_float_div(x_6, x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_profiler_threshold_getSecs___boxed(lean_object* x_1) { +_start: +{ +double x_2; lean_object* x_3; +x_2 = lean_get_profiler_threshold(x_1); +x_3 = lean_box_float(x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Lean_profileit___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { @@ -171,6 +673,43 @@ lean_dec_ref(res); res = initialize_Lean_Data_Options(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +l_Lean_initFn____x40_Lean_Util_Profile___hyg_6____closed__1 = _init_l_Lean_initFn____x40_Lean_Util_Profile___hyg_6____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Util_Profile___hyg_6____closed__1); +l_Lean_initFn____x40_Lean_Util_Profile___hyg_6____closed__2 = _init_l_Lean_initFn____x40_Lean_Util_Profile___hyg_6____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Util_Profile___hyg_6____closed__2); +l_Lean_initFn____x40_Lean_Util_Profile___hyg_6____closed__3 = _init_l_Lean_initFn____x40_Lean_Util_Profile___hyg_6____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Util_Profile___hyg_6____closed__3); +l_Lean_initFn____x40_Lean_Util_Profile___hyg_6____closed__4 = _init_l_Lean_initFn____x40_Lean_Util_Profile___hyg_6____closed__4(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Util_Profile___hyg_6____closed__4); +l_Lean_initFn____x40_Lean_Util_Profile___hyg_6____closed__5 = _init_l_Lean_initFn____x40_Lean_Util_Profile___hyg_6____closed__5(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Util_Profile___hyg_6____closed__5); +l_Lean_initFn____x40_Lean_Util_Profile___hyg_6____closed__6 = _init_l_Lean_initFn____x40_Lean_Util_Profile___hyg_6____closed__6(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Util_Profile___hyg_6____closed__6); +if (builtin) {res = l_Lean_initFn____x40_Lean_Util_Profile___hyg_6_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +l_Lean_profiler = lean_io_result_get_value(res); +lean_mark_persistent(l_Lean_profiler); +lean_dec_ref(res); +}l_Lean_initFn____x40_Lean_Util_Profile___hyg_68____closed__1 = _init_l_Lean_initFn____x40_Lean_Util_Profile___hyg_68____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Util_Profile___hyg_68____closed__1); +l_Lean_initFn____x40_Lean_Util_Profile___hyg_68____closed__2 = _init_l_Lean_initFn____x40_Lean_Util_Profile___hyg_68____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Util_Profile___hyg_68____closed__2); +l_Lean_initFn____x40_Lean_Util_Profile___hyg_68____closed__3 = _init_l_Lean_initFn____x40_Lean_Util_Profile___hyg_68____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Util_Profile___hyg_68____closed__3); +l_Lean_initFn____x40_Lean_Util_Profile___hyg_68____closed__4 = _init_l_Lean_initFn____x40_Lean_Util_Profile___hyg_68____closed__4(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Util_Profile___hyg_68____closed__4); +l_Lean_initFn____x40_Lean_Util_Profile___hyg_68____closed__5 = _init_l_Lean_initFn____x40_Lean_Util_Profile___hyg_68____closed__5(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Util_Profile___hyg_68____closed__5); +if (builtin) {res = l_Lean_initFn____x40_Lean_Util_Profile___hyg_68_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +l_Lean_profiler_threshold = lean_io_result_get_value(res); +lean_mark_persistent(l_Lean_profiler_threshold); +lean_dec_ref(res); +}l___private_Lean_Util_Profile_0__Lean_get__profiler___closed__1 = _init_l___private_Lean_Util_Profile_0__Lean_get__profiler___closed__1(); +lean_mark_persistent(l___private_Lean_Util_Profile_0__Lean_get__profiler___closed__1); +l_Lean_profiler_threshold_getSecs___closed__1 = _init_l_Lean_profiler_threshold_getSecs___closed__1(); +lean_mark_persistent(l_Lean_profiler_threshold_getSecs___closed__1); +l_Lean_profiler_threshold_getSecs___closed__2 = _init_l_Lean_profiler_threshold_getSecs___closed__2(); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lean/Util/Trace.c b/stage0/stdlib/Lean/Util/Trace.c index fcb73e3b2f88..fd629360b264 100644 --- a/stage0/stdlib/Lean/Util/Trace.c +++ b/stage0/stdlib/Lean/Util/Trace.c @@ -23,34 +23,40 @@ LEAN_EXPORT lean_object* l_Lean_addTrace___rarg___lambda__1(lean_object*, lean_o static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__20; LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___rarg(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___lambda__1___closed__28; +extern lean_object* l_Lean_profiler; lean_object* lean_mk_empty_array_with_capacity(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_printTraces___spec__5___rarg___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___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__6; static lean_object* l_Array_forInUnsafe_loop___at_Lean_printTraces___spec__4___rarg___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_printTraces___spec__5___rarg___lambda__2(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___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_Lean_withTraceNode___rarg___lambda__6(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_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, uint8_t, double, lean_object*); static lean_object* l_Lean_withTraceNode_x27___rarg___closed__1; -LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___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*, uint8_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_exceptOptionEmoji(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___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_initFn____x40_Lean_Util_Trace___hyg_1382____closed__5; +LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___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_PersistentArray_forIn___at_Lean_printTraces___spec__1(lean_object*); -static lean_object* l___auto____x40_Lean_Util_Trace___hyg_1952____closed__2; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__12___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_shouldEnableNestedTrace___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___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*, uint8_t, lean_object*); static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__18; +LEAN_EXPORT lean_object* l_Lean_shouldProfile(lean_object*); lean_object* l_Lean_PersistentArray_toArray___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_instMonadTrace___rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___auto____x40_Lean_Util_Trace___hyg_1952_; +static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1478____closed__3; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_printTraces___spec__4___rarg___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___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__33; LEAN_EXPORT lean_object* l_Lean_withTraceNode___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*, uint8_t); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_printTraces___spec__5(lean_object*); +static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2356____closed__2; LEAN_EXPORT lean_object* l_Lean_instExceptToEmojiOption(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___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*, uint8_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___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*, uint8_t, lean_object*, uint8_t, lean_object*); double lean_float_div(double, double); -static lean_object* l___auto____x40_Lean_Util_Trace___hyg_1952____closed__24; LEAN_EXPORT lean_object* l_Lean_withTraceNode_x27___rarg___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); -static lean_object* l___auto____x40_Lean_Util_Trace___hyg_1952____closed__13; +LEAN_EXPORT lean_object* l_Lean_shouldEnableNestedTrace___rarg(lean_object*, lean_object*, lean_object*, lean_object*, double); 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___closed__7; LEAN_EXPORT lean_object* l_Lean_exceptBoolEmoji(lean_object*); @@ -60,20 +66,23 @@ LEAN_EXPORT lean_object* l_Lean_modifyTraces___rarg(lean_object*, lean_object*); static lean_object* l_Lean_TraceState_traces___default___closed__2; LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___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_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__8(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_object*, uint8_t); static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__44; static lean_object* l_Lean_TraceState_traces___default___closed__3; +LEAN_EXPORT double l_Lean_trace_profiler_threshold_getSecs(lean_object*); lean_object* l_Lean_Syntax_getId(lean_object*); static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__8; LEAN_EXPORT lean_object* l_Lean_withSeconds(lean_object*, lean_object*); lean_object* l_Lean_PersistentArray_push___rarg(lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_trace_profiler_threshold_getSecs___boxed(lean_object*); lean_object* l_Lean_mkHashSetImp___rarg(lean_object*); -static lean_object* l___auto____x40_Lean_Util_Trace___hyg_1952____closed__16; static lean_object* l_Lean_registerTraceClass___closed__1; -LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___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*, uint8_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, uint8_t, double, lean_object*); static double l_Lean_withSeconds___rarg___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_doElemTrace_x5b___x5d____; LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2356____closed__7; LEAN_EXPORT lean_object* l_Lean_getTraces(lean_object*); static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__22; LEAN_EXPORT lean_object* l_Lean_exceptBoolEmoji___rarg(lean_object*); @@ -84,58 +93,66 @@ static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doEle LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__4(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*); static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__41; LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2356____closed__13; static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__9; -static lean_object* l___auto____x40_Lean_Util_Trace___hyg_1952____closed__17; -LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__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_withTraceNode___rarg___lambda__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*); 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_withTraceNodeBefore___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*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_shouldProfile___rarg___lambda__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__5___boxed(lean_object**); LEAN_EXPORT lean_object* l_Lean_setTraceState___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_printTraces___spec__4___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___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__39; LEAN_EXPORT lean_object* l_Lean_traceM(lean_object*); static lean_object* l_Lean_crossEmoji___closed__1; +uint8_t lean_float_decLt(double, double); static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__19; static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__23; -LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___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_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___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_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_printTraces___spec__3___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1478____closed__5; LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_printTraces___spec__1___rarg___lambda__2(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_instInhabitedTraceElem___closed__2; static lean_object* l_Lean_withSeconds___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_withTraceNode_x27___rarg___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___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__16; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); -static lean_object* l___auto____x40_Lean_Util_Trace___hyg_1952____closed__3; static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__15; LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Util_Trace___hyg_1952____closed__28; +static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__3; lean_object* l_Lean_Syntax_node5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); static lean_object* l_Lean_instInhabitedTraceState___closed__4; static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__5; LEAN_EXPORT lean_object* l_Lean_addRawTrace(lean_object*); +static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2356____closed__25; lean_object* l_Lean_stringToMessageData(lean_object*); LEAN_EXPORT lean_object* l_Lean_addTrace___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_TraceState_traces___default; LEAN_EXPORT uint8_t l___private_Lean_Util_Trace_0__Lean_checkTraceOption(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instInhabitedTraceState; -static lean_object* l_Lean_withOptProfile___rarg___closed__1; +static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__1; lean_object* l_Lean_Exception_toMessageData(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_printTraces___spec__1___rarg___lambda__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__13; -LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___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*, uint8_t, uint8_t); +LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__6(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_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_resetTraceState___rarg(lean_object*); -static lean_object* l___auto____x40_Lean_Util_Trace___hyg_1952____closed__7; +LEAN_EXPORT lean_object* l___auto____x40_Lean_Util_Trace___hyg_2356_; LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___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*); static lean_object* l_Lean_instInhabitedTraceElem___closed__1; +static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2356____closed__23; 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__2; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_printTraces___spec__4___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__13___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_shouldProfile___rarg___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_resetTraceState(lean_object*); static lean_object* l_Lean_instInhabitedTraceState___closed__2; static lean_object* l_Lean_instExceptToEmojiOption___closed__1; +lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addRawTrace___rarg___lambda__2(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__7(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_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode(lean_object*); LEAN_EXPORT lean_object* l_Lean_exceptBoolEmoji___rarg___boxed(lean_object*); @@ -144,21 +161,18 @@ static lean_object* l___private_Lean_Util_Trace_0__Lean_checkTraceOption___close LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_printTraces___spec__4(lean_object*); static lean_object* l_Lean_bombEmoji___closed__1; size_t lean_usize_of_nat(lean_object*); -static lean_object* l___auto____x40_Lean_Util_Trace___hyg_1952____closed__1; uint8_t l_Lean_HashSetImp_contains___at_Lean_NameHashSet_contains___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_printTraces___spec__3___rarg___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_withTraceNode___rarg___lambda__5___closed__5; -static lean_object* l___auto____x40_Lean_Util_Trace___hyg_1952____closed__10; -static lean_object* l___auto____x40_Lean_Util_Trace___hyg_1952____closed__19; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__13(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_object*); static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__29; static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__2; -LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___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*, uint8_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, lean_object*); lean_object* lean_st_ref_take(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_withTraceNode___rarg___lambda__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, lean_object*); static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__3; static lean_object* l_Lean_resetTraceState___rarg___closed__1; -static lean_object* l___auto____x40_Lean_Util_Trace___hyg_1952____closed__20; -static lean_object* l_Lean_withOptProfile___rarg___closed__2; LEAN_EXPORT lean_object* l_Lean_setTraceState___rarg___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_printTraces___spec__3___rarg___lambda__3(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___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1(lean_object*, lean_object*, lean_object*); @@ -166,20 +180,25 @@ lean_object* l_Lean_SourceInfo_fromRef(lean_object*, uint8_t); static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__52; lean_object* l_Lean_Syntax_node6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_printTraces___spec__2(lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1478____closed__2; +lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(lean_object*, lean_object*); lean_object* l_Lean_HashSetImp_insert___at_Lean_NameHashSet_insert___spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, lean_object*); static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__56; LEAN_EXPORT lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__4; static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__10; LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_printTraces___spec__2___rarg___lambda__2(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___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* l_Lean_Syntax_getKind(lean_object*); LEAN_EXPORT lean_object* l_Lean_resetTraceState___rarg___lambda__1___boxed(lean_object*); +lean_object* l_Lean_PersistentArray_append___rarg(lean_object*, lean_object*); static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__14; LEAN_EXPORT lean_object* l_Lean_printTraces(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_printTraces___spec__3___rarg(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___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__10; LEAN_EXPORT lean_object* l_Lean_setTraceState___rarg___lambda__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__5(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_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, lean_object*); LEAN_EXPORT lean_object* l_Lean_traceM___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_quoteNameMk(lean_object*); LEAN_EXPORT lean_object* l_Lean_bombEmoji; @@ -187,12 +206,11 @@ LEAN_EXPORT lean_object* l_Lean_checkEmoji; lean_object* lean_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getTraces___rarg___lambda__1(lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Util_Trace___hyg_1952____closed__26; static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__12; uint8_t l_List_isEmpty___rarg(lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__1; -LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__8(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_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_printTraces___rarg___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withSeconds___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Util_Trace_0__Lean_checkTraceOption_go(lean_object*, lean_object*, lean_object*); @@ -200,52 +218,56 @@ static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doEle lean_object* l_Lean_Syntax_node3(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__50; lean_object* l_Lean_Name_append(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNode___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*, uint8_t, uint8_t); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_printTraces___spec__5___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___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__51; 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__13; +static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2356____closed__3; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_printTraces___spec__3___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor(lean_object*); -LEAN_EXPORT lean_object* l_Lean_withOptProfile___rarg___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___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*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_printTraces___spec__4___rarg(lean_object*, lean_object*, 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__15; -static lean_object* l___auto____x40_Lean_Util_Trace___hyg_1952____closed__21; LEAN_EXPORT lean_object* l_Lean_instInhabitedTraceElem; -static lean_object* l___auto____x40_Lean_Util_Trace___hyg_1952____closed__9; static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__11; +static double l_Lean_trace_profiler_threshold_getSecs___closed__2; static lean_object* l_Lean_withTraceNode___rarg___lambda__5___closed__1; +LEAN_EXPORT lean_object* l_Lean_shouldEnableNestedTrace___rarg___lambda__1___boxed(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; lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withSeconds___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Util_Trace___hyg_1952____closed__14; +static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2356____closed__9; +LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___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_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___rarg___lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); -lean_object* l_Lean_getBoolOption___rarg(lean_object*, lean_object*, lean_object*, uint8_t); 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*); static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__46; lean_object* l_Lean_Syntax_node2(lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_trace_profiler_threshold; lean_object* l_Lean_Syntax_getArg(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_instExceptToEmojiBool___closed__1; +static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2356____closed__20; LEAN_EXPORT lean_object* l_Lean_withTraceNode_x27___rarg___lambda__2(lean_object*); static lean_object* l_Lean_withTraceNode___rarg___lambda__5___closed__6; -static lean_object* l___auto____x40_Lean_Util_Trace___hyg_1952____closed__6; +static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2356____closed__5; static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__9; static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__8; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_printTraces___spec__5___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withSeconds___rarg___lambda__2(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_Lean_shouldEnableNestedTrace___rarg___lambda__2(lean_object*, double, lean_object*, lean_object*, uint8_t); 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_withOptProfile___rarg___lambda__3(lean_object*, lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___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__1; LEAN_EXPORT lean_object* l_Lean_trace(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___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__6; -static lean_object* l___auto____x40_Lean_Util_Trace___hyg_1952____closed__5; +static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2356____closed__21; lean_object* lean_register_option(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_inheritedTraceOptions; LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_printTraces___spec__1___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*); @@ -254,7 +276,6 @@ LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__1(lean_obje LEAN_EXPORT lean_object* l_Lean_traceM___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_Array_forInUnsafe_loop___at_Lean_printTraces___spec__3___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_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Util_Trace___hyg_1952____closed__22; static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__5; 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__49; @@ -265,18 +286,23 @@ LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg(lean_object*, lean_ob LEAN_EXPORT lean_object* l_Lean_crossEmoji; static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__21; 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_withTraceNode___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*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___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*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_registerTraceClass___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withOptProfile(lean_object*, lean_object*); static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__18; +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1478_(lean_object*); +static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2356____closed__4; 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*); +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382_(lean_object*); static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__24; static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__4; +LEAN_EXPORT lean_object* l_Lean_shouldEnableNestedTrace___rarg___lambda__1(lean_object*, uint8_t, double, lean_object*); lean_object* l_MonadExcept_ofExcept___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___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_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1478____closed__4; LEAN_EXPORT lean_object* l_Lean_printTraces___rarg___lambda__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__11; +static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__6; LEAN_EXPORT lean_object* l_Lean_getTraces___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode_x27___rarg___lambda__2___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore(lean_object*, lean_object*); @@ -286,40 +312,46 @@ lean_object* l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(lean_object*, LEAN_EXPORT lean_object* l_Lean_withSeconds___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__9; static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__32; -LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___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_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Util_Trace___hyg_1952____closed__18; +static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2356____closed__24; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___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_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2356____closed__14; static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__6; static lean_object* l___private_Lean_Util_Trace_0__Lean_checkTraceOption___closed__1; LEAN_EXPORT lean_object* l_Lean_withTraceNode___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___auto____x40_Lean_Util_Trace___hyg_1952____closed__12; lean_object* lean_float_to_string(double); static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_79____closed__1; static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__55; +static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2356____closed__19; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Util_Trace___hyg_1952____closed__23; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_printTraces___spec__5___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__14; static lean_object* l_Lean_withTraceNode___rarg___lambda__5___closed__4; -LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__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_Lean_withTraceNodeBefore___rarg___lambda__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*); static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__20; -LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___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*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2356____closed__16; lean_object* l_Lean_Syntax_node1(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_isTracingEnabledFor___rarg(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_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___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*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__2(lean_object*, 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_withTraceNodeBefore___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*); +static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2356____closed__22; lean_object* lean_nat_sub(lean_object*, lean_object*); +lean_object* l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__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__30; -static lean_object* l_Lean_withOptProfile___rarg___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1(lean_object*, 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_Lean_isTracingEnabledFor___rarg___lambda__1(lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__12(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_instMonadTrace(lean_object*, lean_object*); lean_object* l_IO_println___at_Lean_instEval__1___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_doElemTrace_x5b___x5d_______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_shouldEnableNestedTrace(lean_object*); +static lean_object* l_Lean_shouldProfile___rarg___lambda__1___closed__2; static lean_object* l_Lean_isTracingEnabledFor___rarg___lambda__1___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___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___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__23; @@ -327,8 +359,8 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_printTraces___spec__ LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore(lean_object*); lean_object* lean_erase_macro_scopes(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_checkTraceOption___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Util_Trace___hyg_1952____closed__15; -LEAN_EXPORT lean_object* l_Lean_withOptProfile___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_trace_profiler_threshold_getSecs___closed__1; +static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2356____closed__8; lean_object* l_String_intercalate(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withSeconds___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_IO_monoNanosNow___boxed(lean_object*); @@ -337,38 +369,47 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_printTraces___spec__ LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_printTraces___spec__1___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_exceptOptionEmoji___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___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_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2356____closed__27; LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_checkTraceOption_go___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1478____closed__1; size_t lean_usize_add(size_t, size_t); +static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2356____closed__11; static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__54; static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__7; static lean_object* l_Lean_instInhabitedTraceState___closed__1; static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__13; +LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___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_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__2___closed__2; lean_object* l_Lean_KVMap_findCore(lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_printTraces___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_printTraces___spec__3___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_trace_profiler; LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_79_(lean_object*); static size_t l_Lean_instInhabitedTraceState___closed__3; lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2356____closed__12; static lean_object* l_Lean_withTraceNode___rarg___lambda__5___closed__2; +LEAN_EXPORT lean_object* l_Lean_shouldProfile___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addRawTrace___rarg___lambda__1(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr4(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__7; +LEAN_EXPORT lean_object* l_Lean_shouldEnableNestedTrace___rarg___lambda__2___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___closed__8; static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__40; static lean_object* l_Lean_withTraceNode___rarg___lambda__5___closed__3; LEAN_EXPORT lean_object* l_Lean_addTrace___rarg___lambda__3(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__4; lean_object* lean_string_append(lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2356____closed__15; +static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2356____closed__17; lean_object* lean_array_get_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_instExceptToEmojiBool(lean_object*); -static lean_object* l___auto____x40_Lean_Util_Trace___hyg_1952____closed__25; -static lean_object* l_Lean_withOptProfile___rarg___lambda__3___closed__2; static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__11; -static lean_object* l___auto____x40_Lean_Util_Trace___hyg_1952____closed__11; +static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2356____closed__18; uint8_t lean_usize_dec_lt(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_addRawTrace___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2356____closed__10; static lean_object* l_Lean_registerTraceClass___closed__2; LEAN_EXPORT lean_object* l_Lean_withTraceNode_x27(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode(lean_object*, lean_object*); @@ -377,33 +418,34 @@ static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doEle LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_printTraces___spec__3(lean_object*); LEAN_EXPORT lean_object* l_Lean_printTraces___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_isTracingEnabledFor___rarg___closed__1; -static lean_object* l___auto____x40_Lean_Util_Trace___hyg_1952____closed__4; LEAN_EXPORT lean_object* l_Lean_addTrace(lean_object*); static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__19; +lean_object* l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_68____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_printTraces___spec__2___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__42; static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__35; lean_object* l_String_toSubstring_x27(lean_object*); -LEAN_EXPORT lean_object* l_Lean_withOptProfile___rarg___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_printTraces___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__2; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); +static lean_object* l_Lean_shouldProfile___rarg___lambda__1___closed__1; static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__34; static lean_object* l_Lean_doElemTrace_x5b___x5d_______closed__17; LEAN_EXPORT lean_object* l_Lean_trace___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Util_Trace___hyg_1952____closed__8; +static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2356____closed__6; LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___rarg___lambda__1___boxed(lean_object*, lean_object*, 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*); 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__12; lean_object* l_Lean_MessageData_format(lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Util_Trace___hyg_1952____closed__27; +static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2356____closed__26; +static lean_object* l___auto____x40_Lean_Util_Trace___hyg_2356____closed__1; static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__45; 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___lambda__1___closed__17; static lean_object* l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__21; static lean_object* l_Lean_checkEmoji___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_printTraces___spec__2___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withOptProfile___rarg___lambda__2(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNodeCore___spec__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__38; LEAN_EXPORT lean_object* l_Lean_modifyTraces(lean_object*); @@ -2575,154 +2617,370 @@ lean_dec(x_2); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_withOptProfile___rarg___lambda__1(lean_object* x_1) { +static lean_object* _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__1() { _start: { -lean_object* x_2; lean_object* x_3; -x_2 = lean_box(0); -x_3 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("profiler", 8); +return x_1; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Util_Trace_0__Lean_checkTraceOption___closed__1; +x_2 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__1; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_withOptProfile___rarg___lambda__2(lean_object* x_1) { +static lean_object* _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__3() { _start: { -uint8_t x_2; -x_2 = !lean_is_exclusive(x_1); -if (x_2 == 0) +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("activate nested traces with execution time over threshold", 57); +return x_1; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__4() { +_start: { -lean_object* x_3; lean_object* x_4; -x_3 = lean_ctor_get(x_1, 1); -x_4 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_4, 0, x_3); -lean_ctor_set(x_1, 1, x_4); +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_initFn____x40_Lean_Util_Trace___hyg_1382____closed__1; +x_3 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__3; +x_4 = lean_box(x_1); +x_5 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_5, 0, x_4); +lean_ctor_set(x_5, 1, x_2); +lean_ctor_set(x_5, 2, x_3); +return x_5; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Lean", 4); return x_1; } -else +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__6() { +_start: { -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_5 = lean_ctor_get(x_1, 0); -x_6 = lean_ctor_get(x_1, 1); -lean_inc(x_6); -lean_inc(x_5); -lean_dec(x_1); -x_7 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_7, 0, x_6); -x_8 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_8, 0, x_5); -lean_ctor_set(x_8, 1, x_7); -return x_8; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__5; +x_2 = l___private_Lean_Util_Trace_0__Lean_checkTraceOption___closed__1; +x_3 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} } +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382_(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_initFn____x40_Lean_Util_Trace___hyg_1382____closed__2; +x_3 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__4; +x_4 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__6; +x_5 = l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__1(x_2, x_3, x_4, x_1); +return x_5; } } -static lean_object* _init_l_Lean_withOptProfile___rarg___lambda__3___closed__1() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1478____closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_withOptProfile___rarg___lambda__1), 1, 0); +x_1 = lean_mk_string_from_bytes("threshold", 9); return x_1; } } -static lean_object* _init_l_Lean_withOptProfile___rarg___lambda__3___closed__2() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1478____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Util_Trace_0__Lean_checkTraceOption___closed__1; +x_2 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__1; +x_3 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1478____closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1478____closed__3() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_withOptProfile___rarg___lambda__2), 1, 0); +x_1 = lean_mk_string_from_bytes("threshold in milliseconds, traces below threshold will not be activated", 71); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_withOptProfile___rarg___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4) { +static lean_object* _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1478____closed__4() { _start: { -if (x_4 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = lean_unsigned_to_nat(10u); +x_2 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__1; +x_3 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1478____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_initFn____x40_Lean_Util_Trace___hyg_1478____closed__5() { +_start: { -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__5; +x_2 = l___private_Lean_Util_Trace_0__Lean_checkTraceOption___closed__1; +x_3 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__1; +x_4 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1478____closed__1; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Util_Trace___hyg_1478_(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_initFn____x40_Lean_Util_Trace___hyg_1478____closed__2; +x_3 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1478____closed__4; +x_4 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1478____closed__5; +x_5 = l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_68____spec__1(x_2, x_3, x_4, x_1); +return x_5; +} +} +static lean_object* _init_l_Lean_trace_profiler_threshold_getSecs___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_trace_profiler_threshold; +return x_1; +} +} +static double _init_l_Lean_trace_profiler_threshold_getSecs___closed__2() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; double x_4; +x_1 = lean_unsigned_to_nat(1000u); +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 double l_Lean_trace_profiler_threshold_getSecs(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; uint8_t x_4; lean_object* x_5; double x_6; double x_7; double x_8; +x_2 = l_Lean_trace_profiler_threshold_getSecs___closed__1; +x_3 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_2); +x_4 = 0; +x_5 = lean_unsigned_to_nat(0u); +x_6 = l_Float_ofScientific(x_3, x_4, x_5); +lean_dec(x_3); +x_7 = l_Lean_trace_profiler_threshold_getSecs___closed__2; +x_8 = lean_float_div(x_6, x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_trace_profiler_threshold_getSecs___boxed(lean_object* x_1) { +_start: +{ +double x_2; lean_object* x_3; +x_2 = l_Lean_trace_profiler_threshold_getSecs(x_1); +lean_dec(x_1); +x_3 = lean_box_float(x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_shouldProfile___rarg___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_profiler; +return x_1; +} +} +static lean_object* _init_l_Lean_shouldProfile___rarg___lambda__1___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_trace_profiler; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_shouldProfile___rarg___lambda__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +lean_dec(x_1); +x_4 = lean_ctor_get(x_3, 1); +lean_inc(x_4); lean_dec(x_3); +x_5 = l_Lean_shouldProfile___rarg___lambda__1___closed__1; +x_6 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_2, x_5); +if (x_6 == 0) +{ +lean_object* x_7; uint8_t x_8; lean_object* x_9; lean_object* x_10; +x_7 = l_Lean_shouldProfile___rarg___lambda__1___closed__2; +x_8 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_2, x_7); +x_9 = lean_box(x_8); +x_10 = lean_apply_2(x_4, lean_box(0), x_9); +return x_10; +} +else +{ +uint8_t x_11; lean_object* x_12; lean_object* x_13; +x_11 = 1; +x_12 = lean_box(x_11); +x_13 = lean_apply_2(x_4, lean_box(0), x_12); +return x_13; +} +} +} +LEAN_EXPORT lean_object* l_Lean_shouldProfile___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_ctor_get(x_1, 1); +lean_inc(x_3); +x_4 = lean_alloc_closure((void*)(l_Lean_shouldProfile___rarg___lambda__1___boxed), 2, 1); +lean_closure_set(x_4, 0, x_1); +x_5 = lean_apply_4(x_3, lean_box(0), lean_box(0), x_2, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_shouldProfile(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_shouldProfile___rarg), 2, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_shouldProfile___rarg___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_shouldProfile___rarg___lambda__1(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_shouldEnableNestedTrace___rarg___lambda__1(lean_object* x_1, uint8_t x_2, double x_3, lean_object* x_4) { +_start: +{ +if (x_2 == 0) +{ +lean_object* x_5; lean_object* x_6; double x_7; uint8_t x_8; lean_object* x_9; lean_object* x_10; x_5 = lean_ctor_get(x_1, 0); lean_inc(x_5); lean_dec(x_1); -x_6 = lean_ctor_get(x_5, 0); +x_6 = lean_ctor_get(x_5, 1); lean_inc(x_6); lean_dec(x_5); -x_7 = lean_ctor_get(x_6, 0); -lean_inc(x_7); -lean_dec(x_6); -x_8 = l_Lean_withOptProfile___rarg___lambda__3___closed__1; -x_9 = lean_apply_4(x_7, lean_box(0), lean_box(0), x_8, x_2); -return x_9; +x_7 = l_Lean_trace_profiler_threshold_getSecs(x_4); +x_8 = lean_float_decLt(x_3, x_7); +x_9 = lean_box(x_8); +x_10 = lean_apply_2(x_6, lean_box(0), x_9); +return x_10; } else { -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_1, 0); -lean_inc(x_10); -x_11 = lean_ctor_get(x_10, 0); +lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; +x_11 = lean_ctor_get(x_1, 0); lean_inc(x_11); -lean_dec(x_10); -x_12 = lean_ctor_get(x_11, 0); +lean_dec(x_1); +x_12 = lean_ctor_get(x_11, 1); lean_inc(x_12); lean_dec(x_11); -x_13 = l_Lean_withSeconds___rarg(x_1, x_3, x_2); -x_14 = l_Lean_withOptProfile___rarg___lambda__3___closed__2; -x_15 = lean_apply_4(x_12, lean_box(0), lean_box(0), x_14, x_13); +x_13 = 1; +x_14 = lean_box(x_13); +x_15 = lean_apply_2(x_12, lean_box(0), x_14); return x_15; } } } -static lean_object* _init_l_Lean_withOptProfile___rarg___closed__1() { +LEAN_EXPORT lean_object* l_Lean_shouldEnableNestedTrace___rarg___lambda__2(lean_object* x_1, double x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("profiler", 8); -return x_1; -} -} -static lean_object* _init_l_Lean_withOptProfile___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_withOptProfile___rarg___closed__1; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_6 = lean_box(x_5); +x_7 = lean_box_float(x_2); +x_8 = lean_alloc_closure((void*)(l_Lean_shouldEnableNestedTrace___rarg___lambda__1___boxed), 4, 3); +lean_closure_set(x_8, 0, x_1); +lean_closure_set(x_8, 1, x_6); +lean_closure_set(x_8, 2, x_7); +x_9 = lean_apply_4(x_3, lean_box(0), lean_box(0), x_4, x_8); +return x_9; } } -LEAN_EXPORT lean_object* l_Lean_withOptProfile___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_shouldEnableNestedTrace___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, double x_5) { _start: { -lean_object* x_5; lean_object* x_6; uint8_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_5 = lean_ctor_get(x_1, 1); -lean_inc(x_5); -x_6 = l_Lean_withOptProfile___rarg___closed__2; -x_7 = 0; +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_6 = lean_ctor_get(x_1, 1); +lean_inc(x_6); +lean_inc(x_3); lean_inc(x_1); -x_8 = l_Lean_getBoolOption___rarg(x_1, x_3, x_6, x_7); -x_9 = lean_alloc_closure((void*)(l_Lean_withOptProfile___rarg___lambda__3___boxed), 4, 3); +x_7 = l_Lean_isTracingEnabledFor___rarg(x_1, x_3, x_2, x_4); +x_8 = lean_box_float(x_5); +lean_inc(x_6); +x_9 = lean_alloc_closure((void*)(l_Lean_shouldEnableNestedTrace___rarg___lambda__2___boxed), 5, 4); lean_closure_set(x_9, 0, x_1); -lean_closure_set(x_9, 1, x_4); -lean_closure_set(x_9, 2, x_2); -x_10 = lean_apply_4(x_5, lean_box(0), lean_box(0), x_8, x_9); +lean_closure_set(x_9, 1, x_8); +lean_closure_set(x_9, 2, x_6); +lean_closure_set(x_9, 3, x_3); +x_10 = lean_apply_4(x_6, lean_box(0), lean_box(0), x_7, x_9); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_withOptProfile(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_shouldEnableNestedTrace(lean_object* x_1) { _start: { -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Lean_withOptProfile___rarg), 4, 0); -return x_3; +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_shouldEnableNestedTrace___rarg___boxed), 5, 0); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_withOptProfile___rarg___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_shouldEnableNestedTrace___rarg___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -uint8_t x_5; lean_object* x_6; -x_5 = lean_unbox(x_4); +uint8_t x_5; double x_6; lean_object* x_7; +x_5 = lean_unbox(x_2); +lean_dec(x_2); +x_6 = lean_unbox_float(x_3); +lean_dec(x_3); +x_7 = l_Lean_shouldEnableNestedTrace___rarg___lambda__1(x_1, x_5, x_6, x_4); lean_dec(x_4); -x_6 = l_Lean_withOptProfile___rarg___lambda__3(x_1, x_2, x_3, x_5); -return x_6; +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_shouldEnableNestedTrace___rarg___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +double x_6; uint8_t x_7; lean_object* x_8; +x_6 = lean_unbox_float(x_2); +lean_dec(x_2); +x_7 = lean_unbox(x_5); +lean_dec(x_5); +x_8 = l_Lean_shouldEnableNestedTrace___rarg___lambda__2(x_1, x_6, x_3, x_4, x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_shouldEnableNestedTrace___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +double x_6; lean_object* x_7; +x_6 = lean_unbox_float(x_5); +lean_dec(x_5); +x_7 = l_Lean_shouldEnableNestedTrace___rarg(x_1, x_2, x_3, x_4, x_6); +return x_7; } } LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__1(lean_object* x_1, lean_object* x_2) { @@ -2830,126 +3088,329 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___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, uint8_t 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_withTraceNode___rarg___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, uint8_t x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, uint8_t x_13, double x_14, lean_object* x_15) { _start: { -if (lean_obj_tag(x_12) == 0) +lean_object* x_16; uint8_t x_17; +x_16 = l_Lean_shouldProfile___rarg___lambda__1___closed__1; +x_17 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_12, x_16); +if (x_17 == 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; -x_14 = lean_ctor_get(x_1, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_14, 1); -lean_inc(x_15); -lean_dec(x_14); -x_16 = lean_box(0); -x_17 = lean_apply_2(x_15, lean_box(0), x_16); -x_18 = lean_box(x_8); +if (x_13 == 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; +x_18 = lean_ctor_get(x_1, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_18, 1); +lean_inc(x_19); +lean_dec(x_18); +x_20 = lean_box(0); +x_21 = lean_apply_2(x_19, lean_box(0), x_20); +x_22 = lean_box(x_8); lean_inc(x_11); -x_19 = lean_alloc_closure((void*)(l_Lean_withTraceNode___rarg___lambda__4___boxed), 13, 12); -lean_closure_set(x_19, 0, x_1); -lean_closure_set(x_19, 1, x_2); -lean_closure_set(x_19, 2, x_3); -lean_closure_set(x_19, 3, x_4); -lean_closure_set(x_19, 4, x_5); -lean_closure_set(x_19, 5, x_6); -lean_closure_set(x_19, 6, x_7); -lean_closure_set(x_19, 7, x_18); -lean_closure_set(x_19, 8, x_9); -lean_closure_set(x_19, 9, x_10); -lean_closure_set(x_19, 10, x_11); -lean_closure_set(x_19, 11, x_13); -x_20 = lean_apply_4(x_11, lean_box(0), lean_box(0), x_17, x_19); -return x_20; +x_23 = lean_alloc_closure((void*)(l_Lean_withTraceNode___rarg___lambda__4___boxed), 13, 12); +lean_closure_set(x_23, 0, x_1); +lean_closure_set(x_23, 1, x_2); +lean_closure_set(x_23, 2, x_3); +lean_closure_set(x_23, 3, x_4); +lean_closure_set(x_23, 4, x_5); +lean_closure_set(x_23, 5, x_6); +lean_closure_set(x_23, 6, x_7); +lean_closure_set(x_23, 7, x_22); +lean_closure_set(x_23, 8, x_9); +lean_closure_set(x_23, 9, x_10); +lean_closure_set(x_23, 10, x_11); +lean_closure_set(x_23, 11, x_15); +x_24 = lean_apply_4(x_11, lean_box(0), lean_box(0), x_21, x_23); +return x_24; } else { -lean_object* x_21; double 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_21 = lean_ctor_get(x_12, 0); -lean_inc(x_21); -lean_dec(x_12); -x_22 = lean_unbox_float(x_21); -lean_dec(x_21); -x_23 = lean_float_to_string(x_22); -x_24 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_24, 0, x_23); -x_25 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_25, 0, x_24); -x_26 = l_Lean_withTraceNode___rarg___lambda__5___closed__2; -x_27 = lean_alloc_ctor(7, 2, 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; +x_25 = lean_float_to_string(x_14); +x_26 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_26, 0, x_25); +x_27 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_27, 0, x_26); -lean_ctor_set(x_27, 1, x_25); -x_28 = l_Lean_withTraceNode___rarg___lambda__5___closed__4; +x_28 = l_Lean_withTraceNode___rarg___lambda__5___closed__2; x_29 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_29, 0, x_27); -lean_ctor_set(x_29, 1, x_28); -x_30 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_30, 0, x_29); -lean_ctor_set(x_30, 1, x_13); -x_31 = l_Lean_withTraceNode___rarg___lambda__5___closed__6; +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_27); +x_30 = l_Lean_withTraceNode___rarg___lambda__5___closed__4; +x_31 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_30); x_32 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_32, 0, x_30); -lean_ctor_set(x_32, 1, x_31); -x_33 = lean_ctor_get(x_1, 0); -lean_inc(x_33); -x_34 = lean_ctor_get(x_33, 1); -lean_inc(x_34); -lean_dec(x_33); -x_35 = lean_box(0); -x_36 = lean_apply_2(x_34, lean_box(0), x_35); -x_37 = lean_box(x_8); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_15); +x_33 = l_Lean_withTraceNode___rarg___lambda__5___closed__6; +x_34 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +x_35 = lean_ctor_get(x_1, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_35, 1); +lean_inc(x_36); +lean_dec(x_35); +x_37 = lean_box(0); +x_38 = lean_apply_2(x_36, lean_box(0), x_37); +x_39 = lean_box(x_8); lean_inc(x_11); -x_38 = lean_alloc_closure((void*)(l_Lean_withTraceNode___rarg___lambda__4___boxed), 13, 12); -lean_closure_set(x_38, 0, x_1); -lean_closure_set(x_38, 1, x_2); -lean_closure_set(x_38, 2, x_3); -lean_closure_set(x_38, 3, x_4); -lean_closure_set(x_38, 4, x_5); -lean_closure_set(x_38, 5, x_6); -lean_closure_set(x_38, 6, x_7); -lean_closure_set(x_38, 7, x_37); -lean_closure_set(x_38, 8, x_9); -lean_closure_set(x_38, 9, x_10); -lean_closure_set(x_38, 10, x_11); -lean_closure_set(x_38, 11, x_32); -x_39 = lean_apply_4(x_11, lean_box(0), lean_box(0), x_36, x_38); -return x_39; +x_40 = lean_alloc_closure((void*)(l_Lean_withTraceNode___rarg___lambda__4___boxed), 13, 12); +lean_closure_set(x_40, 0, x_1); +lean_closure_set(x_40, 1, x_2); +lean_closure_set(x_40, 2, x_3); +lean_closure_set(x_40, 3, x_4); +lean_closure_set(x_40, 4, x_5); +lean_closure_set(x_40, 5, x_6); +lean_closure_set(x_40, 6, x_7); +lean_closure_set(x_40, 7, x_39); +lean_closure_set(x_40, 8, x_9); +lean_closure_set(x_40, 9, x_10); +lean_closure_set(x_40, 10, x_11); +lean_closure_set(x_40, 11, x_34); +x_41 = lean_apply_4(x_11, lean_box(0), lean_box(0), x_38, x_40); +return x_41; } } +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; 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_42 = lean_float_to_string(x_14); +x_43 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_43, 0, x_42); +x_44 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_44, 0, x_43); +x_45 = l_Lean_withTraceNode___rarg___lambda__5___closed__2; +x_46 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_44); +x_47 = l_Lean_withTraceNode___rarg___lambda__5___closed__4; +x_48 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_48, 0, x_46); +lean_ctor_set(x_48, 1, x_47); +x_49 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_49, 0, x_48); +lean_ctor_set(x_49, 1, x_15); +x_50 = l_Lean_withTraceNode___rarg___lambda__5___closed__6; +x_51 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_51, 0, x_49); +lean_ctor_set(x_51, 1, x_50); +x_52 = lean_ctor_get(x_1, 0); +lean_inc(x_52); +x_53 = lean_ctor_get(x_52, 1); +lean_inc(x_53); +lean_dec(x_52); +x_54 = lean_box(0); +x_55 = lean_apply_2(x_53, lean_box(0), x_54); +x_56 = lean_box(x_8); +lean_inc(x_11); +x_57 = lean_alloc_closure((void*)(l_Lean_withTraceNode___rarg___lambda__4___boxed), 13, 12); +lean_closure_set(x_57, 0, x_1); +lean_closure_set(x_57, 1, x_2); +lean_closure_set(x_57, 2, x_3); +lean_closure_set(x_57, 3, x_4); +lean_closure_set(x_57, 4, x_5); +lean_closure_set(x_57, 5, x_6); +lean_closure_set(x_57, 6, x_7); +lean_closure_set(x_57, 7, x_56); +lean_closure_set(x_57, 8, x_9); +lean_closure_set(x_57, 9, x_10); +lean_closure_set(x_57, 10, x_11); +lean_closure_set(x_57, 11, x_51); +x_58 = lean_apply_4(x_11, lean_box(0), lean_box(0), x_55, x_57); +return x_58; +} } -LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___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, uint8_t x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___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, uint8_t x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, uint8_t x_13, double x_14, lean_object* x_15) { _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; -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); -lean_inc(x_13); -x_15 = lean_apply_1(x_1, x_13); -x_16 = lean_box(x_9); +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_inc(x_2); +x_16 = lean_apply_1(x_1, x_2); +x_17 = lean_box(x_9); +x_18 = lean_box(x_13); +x_19 = lean_box_float(x_14); lean_inc(x_11); -x_17 = lean_alloc_closure((void*)(l_Lean_withTraceNode___rarg___lambda__5___boxed), 13, 12); -lean_closure_set(x_17, 0, x_2); -lean_closure_set(x_17, 1, x_3); -lean_closure_set(x_17, 2, x_4); -lean_closure_set(x_17, 3, x_5); -lean_closure_set(x_17, 4, x_6); -lean_closure_set(x_17, 5, x_7); -lean_closure_set(x_17, 6, x_8); -lean_closure_set(x_17, 7, x_16); -lean_closure_set(x_17, 8, x_10); -lean_closure_set(x_17, 9, x_13); -lean_closure_set(x_17, 10, x_11); -lean_closure_set(x_17, 11, x_14); -x_18 = lean_apply_4(x_11, lean_box(0), lean_box(0), x_15, x_17); -return x_18; +x_20 = lean_alloc_closure((void*)(l_Lean_withTraceNode___rarg___lambda__5___boxed), 15, 14); +lean_closure_set(x_20, 0, x_3); +lean_closure_set(x_20, 1, x_4); +lean_closure_set(x_20, 2, x_5); +lean_closure_set(x_20, 3, x_6); +lean_closure_set(x_20, 4, x_7); +lean_closure_set(x_20, 5, x_8); +lean_closure_set(x_20, 6, x_15); +lean_closure_set(x_20, 7, x_17); +lean_closure_set(x_20, 8, x_10); +lean_closure_set(x_20, 9, x_2); +lean_closure_set(x_20, 10, x_11); +lean_closure_set(x_20, 11, x_12); +lean_closure_set(x_20, 12, x_18); +lean_closure_set(x_20, 13, x_19); +x_21 = lean_apply_4(x_11, lean_box(0), lean_box(0), x_16, x_20); +return x_21; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__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, uint8_t x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, uint8_t x_13, double 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; +x_16 = lean_ctor_get(x_1, 0); +lean_inc(x_16); +x_17 = lean_box(x_9); +x_18 = lean_box(x_13); +x_19 = lean_box_float(x_14); +lean_inc(x_11); +x_20 = lean_alloc_closure((void*)(l_Lean_withTraceNode___rarg___lambda__6___boxed), 15, 14); +lean_closure_set(x_20, 0, x_2); +lean_closure_set(x_20, 1, x_3); +lean_closure_set(x_20, 2, x_4); +lean_closure_set(x_20, 3, x_5); +lean_closure_set(x_20, 4, x_1); +lean_closure_set(x_20, 5, x_6); +lean_closure_set(x_20, 6, x_7); +lean_closure_set(x_20, 7, x_8); +lean_closure_set(x_20, 8, x_17); +lean_closure_set(x_20, 9, x_10); +lean_closure_set(x_20, 10, x_11); +lean_closure_set(x_20, 11, x_12); +lean_closure_set(x_20, 12, x_18); +lean_closure_set(x_20, 13, x_19); +x_21 = lean_apply_4(x_11, lean_box(0), lean_box(0), x_16, x_20); +return x_21; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__8(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_inc(x_1); +x_6 = l_MonadExcept_ofExcept___rarg(x_1, x_2, x_3); +x_7 = lean_alloc_closure((void*)(l_Lean_getTraces___rarg___lambda__1), 2, 1); +lean_closure_set(x_7, 0, x_1); +x_8 = lean_apply_4(x_4, lean_box(0), lean_box(0), x_6, x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___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, uint8_t x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, uint8_t x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; lean_object* x_15; uint8_t x_16; lean_object* x_36; uint8_t x_37; +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_36 = l_Lean_shouldProfile___rarg___lambda__1___closed__2; +x_37 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_11, x_36); +if (x_37 == 0) +{ +uint8_t x_38; +x_38 = 0; +x_16 = x_38; +goto block_35; +} +else +{ +double x_39; double x_40; uint8_t x_41; +x_39 = l_Lean_trace_profiler_threshold_getSecs(x_11); +x_40 = lean_unbox_float(x_15); +x_41 = lean_float_decLt(x_39, x_40); +x_16 = x_41; +goto block_35; +} +block_35: +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_box(x_8); +x_18 = lean_box(x_16); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_6); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_14); +x_19 = lean_alloc_closure((void*)(l_Lean_withTraceNode___rarg___lambda__7___boxed), 15, 14); +lean_closure_set(x_19, 0, x_1); +lean_closure_set(x_19, 1, x_2); +lean_closure_set(x_19, 2, x_14); +lean_closure_set(x_19, 3, x_3); +lean_closure_set(x_19, 4, x_4); +lean_closure_set(x_19, 5, x_5); +lean_closure_set(x_19, 6, x_6); +lean_closure_set(x_19, 7, x_7); +lean_closure_set(x_19, 8, x_17); +lean_closure_set(x_19, 9, x_9); +lean_closure_set(x_19, 10, x_10); +lean_closure_set(x_19, 11, x_11); +lean_closure_set(x_19, 12, x_18); +lean_closure_set(x_19, 13, x_15); +if (x_12 == 0) +{ +if (x_16 == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +lean_dec(x_19); +x_20 = lean_ctor_get(x_4, 0); +lean_inc(x_20); +lean_dec(x_4); +x_21 = lean_alloc_closure((void*)(l_Lean_PersistentArray_append___rarg), 2, 1); +lean_closure_set(x_21, 0, x_6); +x_22 = lean_apply_1(x_20, x_21); +lean_inc(x_10); +x_23 = lean_alloc_closure((void*)(l_Lean_withTraceNode___rarg___lambda__8___boxed), 5, 4); +lean_closure_set(x_23, 0, x_3); +lean_closure_set(x_23, 1, x_9); +lean_closure_set(x_23, 2, x_14); +lean_closure_set(x_23, 3, x_10); +x_24 = lean_apply_4(x_10, lean_box(0), lean_box(0), x_22, x_23); +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_dec(x_14); +lean_dec(x_9); +lean_dec(x_6); +lean_dec(x_4); +x_25 = lean_ctor_get(x_3, 0); +lean_inc(x_25); +lean_dec(x_3); +x_26 = lean_ctor_get(x_25, 1); +lean_inc(x_26); +lean_dec(x_25); +x_27 = lean_box(0); +x_28 = lean_apply_2(x_26, lean_box(0), x_27); +x_29 = lean_apply_4(x_10, lean_box(0), lean_box(0), x_28, x_19); +return x_29; +} +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +lean_dec(x_14); +lean_dec(x_9); +lean_dec(x_6); +lean_dec(x_4); +x_30 = lean_ctor_get(x_3, 0); +lean_inc(x_30); +lean_dec(x_3); +x_31 = lean_ctor_get(x_30, 1); +lean_inc(x_31); +lean_dec(x_30); +x_32 = lean_box(0); +x_33 = lean_apply_2(x_31, lean_box(0), x_32); +x_34 = lean_apply_4(x_10, lean_box(0), lean_box(0), x_33, x_19); +return x_34; +} } } -LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__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, uint8_t x_13, lean_object* x_14) { +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__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, uint8_t x_11, lean_object* x_12, uint8_t x_13, lean_object* x_14) { _start: { -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_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_15 = lean_ctor_get(x_1, 1); lean_inc(x_15); lean_inc(x_2); @@ -2962,123 +3423,181 @@ x_18 = lean_alloc_closure((void*)(l_Lean_withTraceNode___rarg___lambda__2), 2, 1 lean_closure_set(x_18, 0, x_2); x_19 = lean_apply_3(x_15, lean_box(0), x_17, x_18); lean_inc(x_2); -x_20 = l_Lean_withOptProfile___rarg(x_2, x_5, x_6, x_19); -x_21 = lean_box(x_13); +x_20 = l_Lean_withSeconds___rarg(x_2, x_5, x_19); +x_21 = lean_box(x_11); +x_22 = lean_box(x_13); lean_inc(x_3); -x_22 = lean_alloc_closure((void*)(l_Lean_withTraceNode___rarg___lambda__6___boxed), 12, 11); -lean_closure_set(x_22, 0, x_7); -lean_closure_set(x_22, 1, x_2); -lean_closure_set(x_22, 2, x_8); -lean_closure_set(x_22, 3, x_9); -lean_closure_set(x_22, 4, x_10); -lean_closure_set(x_22, 5, x_14); -lean_closure_set(x_22, 6, x_11); -lean_closure_set(x_22, 7, x_12); -lean_closure_set(x_22, 8, x_21); -lean_closure_set(x_22, 9, x_1); -lean_closure_set(x_22, 10, x_3); -x_23 = lean_apply_4(x_3, lean_box(0), lean_box(0), x_20, x_22); -return x_23; +x_23 = lean_alloc_closure((void*)(l_Lean_withTraceNode___rarg___lambda__9___boxed), 13, 12); +lean_closure_set(x_23, 0, x_6); +lean_closure_set(x_23, 1, x_7); +lean_closure_set(x_23, 2, x_2); +lean_closure_set(x_23, 3, x_8); +lean_closure_set(x_23, 4, x_9); +lean_closure_set(x_23, 5, x_14); +lean_closure_set(x_23, 6, x_10); +lean_closure_set(x_23, 7, x_21); +lean_closure_set(x_23, 8, x_1); +lean_closure_set(x_23, 9, x_3); +lean_closure_set(x_23, 10, x_12); +lean_closure_set(x_23, 11, x_22); +x_24 = lean_apply_4(x_3, lean_box(0), lean_box(0), x_20, x_23); +return x_24; } } -LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___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, uint8_t x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__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, uint8_t x_11, lean_object* x_12, uint8_t 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_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_inc(x_2); lean_inc(x_1); -x_14 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___rarg(x_1, x_2); -x_15 = lean_box(x_12); +x_15 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___rarg(x_1, x_2); +x_16 = lean_box(x_11); +x_17 = lean_box(x_13); lean_inc(x_4); -x_16 = lean_alloc_closure((void*)(l_Lean_withTraceNode___rarg___lambda__7___boxed), 14, 13); -lean_closure_set(x_16, 0, x_3); -lean_closure_set(x_16, 1, x_1); -lean_closure_set(x_16, 2, x_4); -lean_closure_set(x_16, 3, x_5); -lean_closure_set(x_16, 4, x_6); -lean_closure_set(x_16, 5, x_7); -lean_closure_set(x_16, 6, x_8); -lean_closure_set(x_16, 7, x_2); -lean_closure_set(x_16, 8, x_9); -lean_closure_set(x_16, 9, x_10); -lean_closure_set(x_16, 10, x_11); -lean_closure_set(x_16, 11, x_13); -lean_closure_set(x_16, 12, x_15); -x_17 = lean_apply_4(x_4, lean_box(0), lean_box(0), x_14, x_16); -return x_17; +x_18 = lean_alloc_closure((void*)(l_Lean_withTraceNode___rarg___lambda__10___boxed), 14, 13); +lean_closure_set(x_18, 0, x_3); +lean_closure_set(x_18, 1, x_1); +lean_closure_set(x_18, 2, x_4); +lean_closure_set(x_18, 3, x_5); +lean_closure_set(x_18, 4, x_6); +lean_closure_set(x_18, 5, x_7); +lean_closure_set(x_18, 6, x_8); +lean_closure_set(x_18, 7, x_2); +lean_closure_set(x_18, 8, x_9); +lean_closure_set(x_18, 9, x_10); +lean_closure_set(x_18, 10, x_16); +lean_closure_set(x_18, 11, x_12); +lean_closure_set(x_18, 12, x_17); +x_19 = lean_apply_4(x_4, lean_box(0), lean_box(0), x_15, x_18); +return x_19; } } -LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___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, uint8_t x_12, uint8_t x_13) { +LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__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, uint8_t x_11, lean_object* x_12, uint8_t x_13) { _start: { +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_box(x_11); +x_15 = lean_box(x_13); +lean_inc(x_12); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_1); +x_16 = lean_alloc_closure((void*)(l_Lean_withTraceNode___rarg___lambda__11___boxed), 14, 13); +lean_closure_set(x_16, 0, x_1); +lean_closure_set(x_16, 1, x_2); +lean_closure_set(x_16, 2, x_3); +lean_closure_set(x_16, 3, x_4); +lean_closure_set(x_16, 4, x_5); +lean_closure_set(x_16, 5, x_6); +lean_closure_set(x_16, 6, x_7); +lean_closure_set(x_16, 7, x_8); +lean_closure_set(x_16, 8, x_9); +lean_closure_set(x_16, 9, x_10); +lean_closure_set(x_16, 10, x_14); +lean_closure_set(x_16, 11, x_12); +lean_closure_set(x_16, 12, x_15); if (x_13 == 0) { -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_object* x_17; uint8_t x_18; +x_17 = l_Lean_shouldProfile___rarg___lambda__1___closed__2; +x_18 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_12, x_17); +lean_dec(x_12); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +lean_dec(x_16); +x_19 = lean_alloc_closure((void*)(l_Lean_getTraces___rarg___lambda__1), 2, 1); +lean_closure_set(x_19, 0, x_1); +x_20 = lean_apply_4(x_4, lean_box(0), lean_box(0), x_5, x_19); +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_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_1; +x_21 = lean_ctor_get(x_1, 0); +lean_inc(x_21); +lean_dec(x_1); +x_22 = lean_ctor_get(x_21, 1); +lean_inc(x_22); +lean_dec(x_21); +x_23 = lean_box(0); +x_24 = lean_apply_2(x_22, lean_box(0), x_23); +x_25 = lean_apply_4(x_4, lean_box(0), lean_box(0), x_24, x_16); +return x_25; +} } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_14 = lean_ctor_get(x_2, 0); -lean_inc(x_14); -x_15 = lean_box(x_12); -lean_inc(x_6); -x_16 = lean_alloc_closure((void*)(l_Lean_withTraceNode___rarg___lambda__8___boxed), 13, 12); -lean_closure_set(x_16, 0, x_3); -lean_closure_set(x_16, 1, x_4); -lean_closure_set(x_16, 2, x_5); -lean_closure_set(x_16, 3, x_6); -lean_closure_set(x_16, 4, x_1); -lean_closure_set(x_16, 5, x_7); -lean_closure_set(x_16, 6, x_8); -lean_closure_set(x_16, 7, x_9); -lean_closure_set(x_16, 8, x_2); -lean_closure_set(x_16, 9, x_10); -lean_closure_set(x_16, 10, x_11); -lean_closure_set(x_16, 11, x_15); -x_17 = lean_apply_4(x_6, lean_box(0), lean_box(0), x_14, x_16); -return x_17; +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +lean_dec(x_12); +lean_dec(x_5); +x_26 = lean_ctor_get(x_1, 0); +lean_inc(x_26); +lean_dec(x_1); +x_27 = lean_ctor_get(x_26, 1); +lean_inc(x_27); +lean_dec(x_26); +x_28 = lean_box(0); +x_29 = lean_apply_2(x_27, lean_box(0), x_28); +x_30 = lean_apply_4(x_4, lean_box(0), lean_box(0), x_29, x_16); +return x_30; } } } -LEAN_EXPORT lean_object* l_Lean_withTraceNode___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_object* x_12, uint8_t x_13) { +LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___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, 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) { _start: { -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_1, 1); -lean_inc(x_14); -lean_inc(x_10); -lean_inc(x_6); +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +lean_inc(x_4); lean_inc(x_1); -x_15 = l_Lean_isTracingEnabledFor___rarg(x_1, x_6, x_3, x_10); +x_15 = l_Lean_isTracingEnabledFor___rarg(x_1, x_2, x_3, x_4); x_16 = lean_box(x_13); -lean_inc(x_14); -x_17 = lean_alloc_closure((void*)(l_Lean_withTraceNode___rarg___lambda__9___boxed), 13, 12); -lean_closure_set(x_17, 0, x_12); -lean_closure_set(x_17, 1, x_4); -lean_closure_set(x_17, 2, x_1); -lean_closure_set(x_17, 3, x_2); +lean_inc(x_7); +x_17 = lean_alloc_closure((void*)(l_Lean_withTraceNode___rarg___lambda__12___boxed), 13, 12); +lean_closure_set(x_17, 0, x_1); +lean_closure_set(x_17, 1, x_5); +lean_closure_set(x_17, 2, x_6); +lean_closure_set(x_17, 3, x_7); lean_closure_set(x_17, 4, x_8); -lean_closure_set(x_17, 5, x_14); -lean_closure_set(x_17, 6, x_9); -lean_closure_set(x_17, 7, x_6); -lean_closure_set(x_17, 8, x_11); -lean_closure_set(x_17, 9, x_5); -lean_closure_set(x_17, 10, x_10); -lean_closure_set(x_17, 11, x_16); -x_18 = lean_apply_4(x_14, lean_box(0), lean_box(0), x_15, x_17); +lean_closure_set(x_17, 5, x_9); +lean_closure_set(x_17, 6, x_10); +lean_closure_set(x_17, 7, x_11); +lean_closure_set(x_17, 8, x_12); +lean_closure_set(x_17, 9, x_4); +lean_closure_set(x_17, 10, x_16); +lean_closure_set(x_17, 11, x_14); +x_18 = lean_apply_4(x_7, lean_box(0), lean_box(0), x_15, x_17); return x_18; } } +LEAN_EXPORT lean_object* l_Lean_withTraceNode___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_object* x_12, uint8_t x_13) { +_start: +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_14 = lean_ctor_get(x_1, 1); +lean_inc(x_14); +x_15 = lean_box(x_13); +lean_inc(x_14); +lean_inc(x_6); +x_16 = lean_alloc_closure((void*)(l_Lean_withTraceNode___rarg___lambda__13___boxed), 14, 13); +lean_closure_set(x_16, 0, x_1); +lean_closure_set(x_16, 1, x_6); +lean_closure_set(x_16, 2, x_3); +lean_closure_set(x_16, 3, x_10); +lean_closure_set(x_16, 4, x_2); +lean_closure_set(x_16, 5, x_8); +lean_closure_set(x_16, 6, x_14); +lean_closure_set(x_16, 7, x_12); +lean_closure_set(x_16, 8, x_9); +lean_closure_set(x_16, 9, x_4); +lean_closure_set(x_16, 10, x_11); +lean_closure_set(x_16, 11, x_5); +lean_closure_set(x_16, 12, x_15); +x_17 = lean_apply_4(x_14, lean_box(0), lean_box(0), x_6, x_16); +return x_17; +} +} LEAN_EXPORT lean_object* l_Lean_withTraceNode(lean_object* x_1, lean_object* x_2) { _start: { @@ -3107,55 +3626,115 @@ lean_dec(x_13); return x_15; } } -LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___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_EXPORT lean_object* l_Lean_withTraceNode___rarg___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: { -uint8_t x_14; lean_object* x_15; -x_14 = lean_unbox(x_8); +uint8_t x_16; uint8_t x_17; double x_18; lean_object* x_19; +x_16 = lean_unbox(x_8); lean_dec(x_8); -x_15 = l_Lean_withTraceNode___rarg___lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_14, x_9, x_10, x_11, x_12, x_13); -return x_15; +x_17 = lean_unbox(x_13); +lean_dec(x_13); +x_18 = lean_unbox_float(x_14); +lean_dec(x_14); +x_19 = l_Lean_withTraceNode___rarg___lambda__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_16, x_9, x_10, x_11, x_12, x_17, x_18, x_15); +lean_dec(x_12); +return x_19; } } -LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__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_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__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) { _start: { -uint8_t x_13; lean_object* x_14; -x_13 = lean_unbox(x_9); +uint8_t x_16; uint8_t x_17; double x_18; lean_object* x_19; +x_16 = lean_unbox(x_9); lean_dec(x_9); -x_14 = l_Lean_withTraceNode___rarg___lambda__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_13, x_10, x_11, x_12); -return x_14; +x_17 = lean_unbox(x_13); +lean_dec(x_13); +x_18 = lean_unbox_float(x_14); +lean_dec(x_14); +x_19 = l_Lean_withTraceNode___rarg___lambda__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_16, x_10, x_11, x_12, x_17, x_18, x_15); +return x_19; } } -LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__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_Lean_withTraceNode___rarg___lambda__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: { -uint8_t x_15; lean_object* x_16; -x_15 = lean_unbox(x_13); +uint8_t x_16; uint8_t x_17; double x_18; lean_object* x_19; +x_16 = lean_unbox(x_9); +lean_dec(x_9); +x_17 = lean_unbox(x_13); lean_dec(x_13); -x_16 = l_Lean_withTraceNode___rarg___lambda__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_15, x_14); -return x_16; +x_18 = lean_unbox_float(x_14); +lean_dec(x_14); +x_19 = l_Lean_withTraceNode___rarg___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_16, x_10, x_11, x_12, x_17, x_18, x_15); +lean_dec(x_15); +return x_19; } } -LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__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_EXPORT lean_object* l_Lean_withTraceNode___rarg___lambda__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -uint8_t x_14; lean_object* x_15; -x_14 = lean_unbox(x_12); -lean_dec(x_12); -x_15 = l_Lean_withTraceNode___rarg___lambda__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_14, x_13); -return x_15; +lean_object* x_6; +x_6 = l_Lean_withTraceNode___rarg___lambda__8(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_5); +return x_6; } } LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___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, lean_object* x_12, lean_object* x_13) { _start: { uint8_t x_14; uint8_t x_15; lean_object* x_16; -x_14 = lean_unbox(x_12); +x_14 = lean_unbox(x_8); +lean_dec(x_8); +x_15 = lean_unbox(x_12); lean_dec(x_12); +x_16 = l_Lean_withTraceNode___rarg___lambda__9(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_14, x_9, x_10, x_11, x_15, x_13); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___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) { +_start: +{ +uint8_t x_15; uint8_t x_16; lean_object* x_17; +x_15 = lean_unbox(x_11); +lean_dec(x_11); +x_16 = lean_unbox(x_13); +lean_dec(x_13); +x_17 = l_Lean_withTraceNode___rarg___lambda__10(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_15, x_12, x_16, x_14); +return x_17; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___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, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +uint8_t x_15; uint8_t x_16; lean_object* x_17; +x_15 = lean_unbox(x_11); +lean_dec(x_11); +x_16 = lean_unbox(x_13); +lean_dec(x_13); +x_17 = l_Lean_withTraceNode___rarg___lambda__11(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_15, x_12, x_16, x_14); +lean_dec(x_14); +return x_17; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___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, lean_object* x_7, 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; uint8_t x_15; lean_object* x_16; +x_14 = lean_unbox(x_11); +lean_dec(x_11); x_15 = lean_unbox(x_13); lean_dec(x_13); -x_16 = l_Lean_withTraceNode___rarg___lambda__9(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_14, x_15); +x_16 = l_Lean_withTraceNode___rarg___lambda__12(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_14, x_12, x_15); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___rarg___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, 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_15; lean_object* x_16; +x_15 = lean_unbox(x_13); +lean_dec(x_13); +x_16 = l_Lean_withTraceNode___rarg___lambda__13(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_15, x_14); return x_16; } } @@ -3172,30 +3751,28 @@ return x_15; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1___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, uint8_t 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_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; x_13 = lean_ctor_get(x_1, 1); lean_inc(x_13); -lean_inc(x_9); -lean_inc(x_6); -lean_inc(x_1); -x_14 = l_Lean_isTracingEnabledFor___rarg(x_1, x_6, x_3, x_9); -x_15 = lean_box(x_12); +x_14 = lean_box(x_12); lean_inc(x_13); -x_16 = lean_alloc_closure((void*)(l_Lean_withTraceNode___rarg___lambda__9___boxed), 13, 12); -lean_closure_set(x_16, 0, x_11); -lean_closure_set(x_16, 1, x_4); -lean_closure_set(x_16, 2, x_1); -lean_closure_set(x_16, 3, x_2); -lean_closure_set(x_16, 4, x_7); -lean_closure_set(x_16, 5, x_13); -lean_closure_set(x_16, 6, x_8); -lean_closure_set(x_16, 7, x_6); -lean_closure_set(x_16, 8, x_10); -lean_closure_set(x_16, 9, x_5); -lean_closure_set(x_16, 10, x_9); -lean_closure_set(x_16, 11, x_15); -x_17 = lean_apply_4(x_13, lean_box(0), lean_box(0), x_14, x_16); -return x_17; +lean_inc(x_6); +x_15 = lean_alloc_closure((void*)(l_Lean_withTraceNode___rarg___lambda__13___boxed), 14, 13); +lean_closure_set(x_15, 0, x_1); +lean_closure_set(x_15, 1, x_6); +lean_closure_set(x_15, 2, x_3); +lean_closure_set(x_15, 3, x_9); +lean_closure_set(x_15, 4, x_2); +lean_closure_set(x_15, 5, x_7); +lean_closure_set(x_15, 6, x_13); +lean_closure_set(x_15, 7, x_11); +lean_closure_set(x_15, 8, x_8); +lean_closure_set(x_15, 9, x_4); +lean_closure_set(x_15, 10, x_10); +lean_closure_set(x_15, 11, x_5); +lean_closure_set(x_15, 12, x_14); +x_16 = lean_apply_4(x_13, lean_box(0), lean_box(0), x_6, x_15); +return x_16; } } LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_withTraceNode_x27___spec__1(lean_object* x_1, lean_object* x_2) { @@ -3316,19 +3893,11 @@ LEAN_EXPORT lean_object* l_Lean_withTraceNode_x27___rarg___boxed(lean_object* x_ uint8_t x_12; lean_object* x_13; x_12 = lean_unbox(x_11); lean_dec(x_11); -x_13 = l_Lean_withTraceNode_x27___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_12); -return x_13; -} -} -static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Lean", 4); -return x_1; +x_13 = l_Lean_withTraceNode_x27___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_12); +return x_13; } } -static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__2() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2356____closed__1() { _start: { lean_object* x_1; @@ -3336,7 +3905,7 @@ x_1 = lean_mk_string_from_bytes("Parser", 6); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__3() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2356____closed__2() { _start: { lean_object* x_1; @@ -3344,7 +3913,7 @@ x_1 = lean_mk_string_from_bytes("Tactic", 6); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__4() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2356____closed__3() { _start: { lean_object* x_1; @@ -3352,19 +3921,19 @@ x_1 = lean_mk_string_from_bytes("tacticSeq", 9); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__5() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2356____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___auto____x40_Lean_Util_Trace___hyg_1952____closed__1; -x_2 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__2; -x_3 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__3; -x_4 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__4; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__5; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__1; +x_3 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__2; +x_4 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__3; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__6() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2356____closed__5() { _start: { lean_object* x_1; @@ -3372,19 +3941,19 @@ x_1 = lean_mk_string_from_bytes("tacticSeq1Indented", 18); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__7() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2356____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__1; -x_2 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__2; -x_3 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__3; -x_4 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__6; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__5; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__1; +x_3 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__2; +x_4 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__5; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__8() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2356____closed__7() { _start: { lean_object* x_1; @@ -3392,17 +3961,17 @@ x_1 = lean_mk_string_from_bytes("null", 4); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__9() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2356____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__8; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__7; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__10() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2356____closed__9() { _start: { lean_object* x_1; @@ -3410,41 +3979,41 @@ x_1 = lean_mk_string_from_bytes("exact", 5); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__11() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2356____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__1; -x_2 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__2; -x_3 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__3; -x_4 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__10; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__5; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__1; +x_3 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__2; +x_4 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__9; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__12() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2356____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(2); -x_2 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__10; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__9; 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___auto____x40_Lean_Util_Trace___hyg_1952____closed__13() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2356____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_instInhabitedTraceState___closed__1; -x_2 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__12; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__11; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__14() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2356____closed__13() { _start: { lean_object* x_1; @@ -3452,7 +4021,7 @@ x_1 = lean_mk_string_from_bytes("Term", 4); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__15() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2356____closed__14() { _start: { lean_object* x_1; @@ -3460,19 +4029,19 @@ x_1 = lean_mk_string_from_bytes("declName", 8); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__16() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2356____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__1; -x_2 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__2; -x_3 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__14; -x_4 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__15; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__5; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__1; +x_3 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__13; +x_4 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__14; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__17() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2356____closed__16() { _start: { lean_object* x_1; @@ -3480,35 +4049,35 @@ x_1 = lean_mk_string_from_bytes("decl_name%", 10); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__18() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2356____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(2); -x_2 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__17; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2356____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___auto____x40_Lean_Util_Trace___hyg_1952____closed__19() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2356____closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_instInhabitedTraceState___closed__1; -x_2 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__18; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__17; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__20() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2356____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___auto____x40_Lean_Util_Trace___hyg_1952____closed__16; -x_3 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__19; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__15; +x_3 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__18; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -3516,23 +4085,23 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__21() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2356____closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__13; -x_2 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__20; +x_1 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__12; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__19; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__22() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2356____closed__21() { _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___auto____x40_Lean_Util_Trace___hyg_1952____closed__11; -x_3 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__21; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__10; +x_3 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__20; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -3540,23 +4109,23 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__23() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2356____closed__22() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_instInhabitedTraceState___closed__1; -x_2 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__22; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__21; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__24() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2356____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___auto____x40_Lean_Util_Trace___hyg_1952____closed__9; -x_3 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__23; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__8; +x_3 = l___auto____x40_Lean_Util_Trace___hyg_2356____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); @@ -3564,23 +4133,23 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__25() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2356____closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_instInhabitedTraceState___closed__1; -x_2 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__24; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__23; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__26() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2356____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___auto____x40_Lean_Util_Trace___hyg_1952____closed__7; -x_3 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__25; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__6; +x_3 = l___auto____x40_Lean_Util_Trace___hyg_2356____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); @@ -3588,23 +4157,23 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__27() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2356____closed__26() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_instInhabitedTraceState___closed__1; -x_2 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__26; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__25; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__28() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2356____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___auto____x40_Lean_Util_Trace___hyg_1952____closed__5; -x_3 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__27; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__4; +x_3 = l___auto____x40_Lean_Util_Trace___hyg_2356____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); @@ -3612,11 +4181,11 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_1952_() { +static lean_object* _init_l___auto____x40_Lean_Util_Trace___hyg_2356_() { _start: { lean_object* x_1; -x_1 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__28; +x_1 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__27; return x_1; } } @@ -3765,7 +4334,7 @@ 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 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__1; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__5; x_2 = l_Lean_doElemTrace_x5b___x5d_______closed__1; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; @@ -4021,9 +4590,9 @@ static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__1; -x_2 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__2; -x_3 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__14; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__5; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__1; +x_3 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__13; x_4 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; @@ -4049,9 +4618,9 @@ static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__1; -x_2 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__2; -x_3 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__14; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__5; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__1; +x_3 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__13; x_4 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__4; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; @@ -4069,9 +4638,9 @@ static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__1; -x_2 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__2; -x_3 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__14; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__5; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__1; +x_3 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__13; x_4 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__6; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; @@ -4089,9 +4658,9 @@ static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__1; -x_2 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__2; -x_3 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__14; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__5; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__1; +x_3 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__13; x_4 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__8; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; @@ -4117,9 +4686,9 @@ static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__1; -x_2 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__2; -x_3 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__14; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__5; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__1; +x_3 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__13; x_4 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__11; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; @@ -4137,9 +4706,9 @@ static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__1; -x_2 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__2; -x_3 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__14; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__5; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__1; +x_3 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__13; x_4 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__13; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; @@ -4192,9 +4761,9 @@ static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__1; -x_2 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__2; -x_3 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__14; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__5; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__1; +x_3 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__13; x_4 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__19; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; @@ -4220,9 +4789,9 @@ static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__1; -x_2 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__2; -x_3 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__14; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__5; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__1; +x_3 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__13; x_4 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__22; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; @@ -4240,9 +4809,9 @@ static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__1; -x_2 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__2; -x_3 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__14; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__5; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__1; +x_3 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__13; x_4 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__24; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; @@ -4268,9 +4837,9 @@ static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__1; -x_2 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__2; -x_3 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__14; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__5; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__1; +x_3 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__13; x_4 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__27; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; @@ -4296,9 +4865,9 @@ static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__1; -x_2 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__2; -x_3 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__14; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__5; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__1; +x_3 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__13; x_4 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__30; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; @@ -4333,7 +4902,7 @@ static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__1; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__5; x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__34; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; @@ -4413,9 +4982,9 @@ static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__1; -x_2 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__2; -x_3 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__14; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__5; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__1; +x_3 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__13; x_4 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__42; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; @@ -4450,7 +5019,7 @@ static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__1; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__5; x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__46; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; @@ -4514,9 +5083,9 @@ static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__1; -x_2 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__2; -x_3 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__14; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__5; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__1; +x_3 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__13; x_4 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__52; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; @@ -4570,7 +5139,7 @@ lean_inc(x_7); x_13 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_13, 0, x_7); lean_ctor_set(x_13, 1, x_12); -x_14 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__9; +x_14 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__8; x_15 = l_Lean_instInhabitedTraceState___closed__1; lean_inc(x_7); x_16 = lean_alloc_ctor(1, 3, 0); @@ -4790,9 +5359,9 @@ static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__1; -x_2 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__2; -x_3 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__14; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__5; +x_2 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__1; +x_3 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__13; x_4 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__3; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; @@ -4837,7 +5406,7 @@ static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__1; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__5; x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__6; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; @@ -4901,7 +5470,7 @@ static lean_object* _init_l_Lean___aux__Lean__Util__Trace______macroRules__Lean_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Lean_Util_Trace___hyg_1952____closed__1; +x_1 = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__5; x_2 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___closed__14; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; @@ -4993,7 +5562,7 @@ lean_ctor_set(x_29, 0, x_18); 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___auto____x40_Lean_Util_Trace___hyg_1952____closed__9; +x_30 = l___auto____x40_Lean_Util_Trace___hyg_2356____closed__8; lean_inc(x_18); x_31 = l_Lean_Syntax_node1(x_18, x_30, x_29); x_32 = l_Lean___aux__Lean__Util__Trace______macroRules__Lean__doElemTrace_x5b___x5d______1___lambda__1___closed__40; @@ -5273,282 +5842,525 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___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, uint8_t x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___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, uint8_t x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, uint8_t x_13, double x_14, lean_object* x_15) { _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; lean_object* x_21; -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_12); -x_14 = lean_apply_1(x_1, x_12); -x_15 = l_Lean_stringToMessageData(x_14); -lean_dec(x_14); -x_16 = l_Lean_withTraceNode___rarg___lambda__5___closed__6; -x_17 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_17, 0, x_16); -lean_ctor_set(x_17, 1, x_15); -x_18 = l_Lean_withTraceNodeBefore___rarg___lambda__2___closed__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; uint8_t x_25; +lean_inc(x_2); +x_16 = lean_apply_1(x_1, x_2); +x_17 = l_Lean_stringToMessageData(x_16); +lean_dec(x_16); +x_18 = l_Lean_withTraceNode___rarg___lambda__5___closed__6; x_19 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_19, 0, x_17); -lean_ctor_set(x_19, 1, x_18); -x_20 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_2); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_17); +x_20 = l_Lean_withTraceNodeBefore___rarg___lambda__2___closed__2; x_21 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_21, 0, x_20); -lean_ctor_set(x_21, 1, x_16); -if (lean_obj_tag(x_13) == 0) +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); +x_22 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_3); +x_23 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_18); +x_24 = l_Lean_shouldProfile___rarg___lambda__1___closed__1; +x_25 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_12, x_24); +if (x_25 == 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_8, 0); -lean_inc(x_22); -x_23 = lean_ctor_get(x_22, 1); -lean_inc(x_23); -lean_dec(x_22); -x_24 = lean_box(0); -x_25 = lean_apply_2(x_23, lean_box(0), x_24); -x_26 = lean_box(x_7); -lean_inc(x_10); -x_27 = lean_alloc_closure((void*)(l_Lean_withTraceNodeBefore___rarg___lambda__1___boxed), 11, 10); -lean_closure_set(x_27, 0, x_3); -lean_closure_set(x_27, 1, x_4); -lean_closure_set(x_27, 2, x_5); -lean_closure_set(x_27, 3, x_6); -lean_closure_set(x_27, 4, x_26); -lean_closure_set(x_27, 5, x_8); -lean_closure_set(x_27, 6, x_9); -lean_closure_set(x_27, 7, x_12); -lean_closure_set(x_27, 8, x_10); -lean_closure_set(x_27, 9, x_21); -x_28 = lean_apply_4(x_10, lean_box(0), lean_box(0), x_25, x_27); -return x_28; +if (x_13 == 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; +x_26 = lean_ctor_get(x_9, 0); +lean_inc(x_26); +x_27 = lean_ctor_get(x_26, 1); +lean_inc(x_27); +lean_dec(x_26); +x_28 = lean_box(0); +x_29 = lean_apply_2(x_27, lean_box(0), x_28); +x_30 = lean_box(x_8); +lean_inc(x_11); +x_31 = lean_alloc_closure((void*)(l_Lean_withTraceNodeBefore___rarg___lambda__1___boxed), 11, 10); +lean_closure_set(x_31, 0, x_4); +lean_closure_set(x_31, 1, x_5); +lean_closure_set(x_31, 2, x_6); +lean_closure_set(x_31, 3, x_7); +lean_closure_set(x_31, 4, x_30); +lean_closure_set(x_31, 5, x_9); +lean_closure_set(x_31, 6, x_10); +lean_closure_set(x_31, 7, x_2); +lean_closure_set(x_31, 8, x_11); +lean_closure_set(x_31, 9, x_23); +x_32 = lean_apply_4(x_11, lean_box(0), lean_box(0), x_29, x_31); +return x_32; } else { -lean_object* x_29; double x_30; lean_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_29 = lean_ctor_get(x_13, 0); -lean_inc(x_29); -lean_dec(x_13); -x_30 = lean_unbox_float(x_29); -lean_dec(x_29); -x_31 = lean_float_to_string(x_30); -x_32 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_32, 0, x_31); -x_33 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_33, 0, x_32); -x_34 = l_Lean_withTraceNode___rarg___lambda__5___closed__2; -x_35 = lean_alloc_ctor(7, 2, 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; +x_33 = lean_float_to_string(x_14); +x_34 = lean_alloc_ctor(3, 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); -lean_ctor_set(x_35, 1, x_33); -x_36 = l_Lean_withTraceNode___rarg___lambda__5___closed__4; +x_36 = l_Lean_withTraceNode___rarg___lambda__5___closed__2; x_37 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_37, 0, x_35); -lean_ctor_set(x_37, 1, x_36); -x_38 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_38, 0, x_37); -lean_ctor_set(x_38, 1, x_21); +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set(x_37, 1, x_35); +x_38 = l_Lean_withTraceNode___rarg___lambda__5___closed__4; x_39 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_39, 1, x_16); -x_40 = lean_ctor_get(x_8, 0); -lean_inc(x_40); -x_41 = lean_ctor_get(x_40, 1); -lean_inc(x_41); -lean_dec(x_40); -x_42 = lean_box(0); -x_43 = lean_apply_2(x_41, lean_box(0), x_42); -x_44 = lean_box(x_7); -lean_inc(x_10); -x_45 = lean_alloc_closure((void*)(l_Lean_withTraceNodeBefore___rarg___lambda__1___boxed), 11, 10); -lean_closure_set(x_45, 0, x_3); -lean_closure_set(x_45, 1, x_4); -lean_closure_set(x_45, 2, x_5); -lean_closure_set(x_45, 3, x_6); -lean_closure_set(x_45, 4, x_44); -lean_closure_set(x_45, 5, x_8); -lean_closure_set(x_45, 6, x_9); -lean_closure_set(x_45, 7, x_12); -lean_closure_set(x_45, 8, x_10); -lean_closure_set(x_45, 9, x_39); -x_46 = lean_apply_4(x_10, lean_box(0), lean_box(0), x_43, x_45); -return x_46; +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +x_40 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_23); +x_41 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_18); +x_42 = lean_ctor_get(x_9, 0); +lean_inc(x_42); +x_43 = lean_ctor_get(x_42, 1); +lean_inc(x_43); +lean_dec(x_42); +x_44 = lean_box(0); +x_45 = lean_apply_2(x_43, lean_box(0), x_44); +x_46 = lean_box(x_8); +lean_inc(x_11); +x_47 = lean_alloc_closure((void*)(l_Lean_withTraceNodeBefore___rarg___lambda__1___boxed), 11, 10); +lean_closure_set(x_47, 0, x_4); +lean_closure_set(x_47, 1, x_5); +lean_closure_set(x_47, 2, x_6); +lean_closure_set(x_47, 3, x_7); +lean_closure_set(x_47, 4, x_46); +lean_closure_set(x_47, 5, x_9); +lean_closure_set(x_47, 6, x_10); +lean_closure_set(x_47, 7, x_2); +lean_closure_set(x_47, 8, x_11); +lean_closure_set(x_47, 9, x_41); +x_48 = lean_apply_4(x_11, lean_box(0), lean_box(0), x_45, x_47); +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; +x_49 = lean_float_to_string(x_14); +x_50 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_50, 0, x_49); +x_51 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_51, 0, x_50); +x_52 = l_Lean_withTraceNode___rarg___lambda__5___closed__2; +x_53 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_53, 0, x_52); +lean_ctor_set(x_53, 1, x_51); +x_54 = l_Lean_withTraceNode___rarg___lambda__5___closed__4; +x_55 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +x_56 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_56, 0, x_55); +lean_ctor_set(x_56, 1, x_23); +x_57 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_57, 0, x_56); +lean_ctor_set(x_57, 1, x_18); +x_58 = lean_ctor_get(x_9, 0); +lean_inc(x_58); +x_59 = lean_ctor_get(x_58, 1); +lean_inc(x_59); +lean_dec(x_58); +x_60 = lean_box(0); +x_61 = lean_apply_2(x_59, lean_box(0), x_60); +x_62 = lean_box(x_8); +lean_inc(x_11); +x_63 = lean_alloc_closure((void*)(l_Lean_withTraceNodeBefore___rarg___lambda__1___boxed), 11, 10); +lean_closure_set(x_63, 0, x_4); +lean_closure_set(x_63, 1, x_5); +lean_closure_set(x_63, 2, x_6); +lean_closure_set(x_63, 3, x_7); +lean_closure_set(x_63, 4, x_62); +lean_closure_set(x_63, 5, x_9); +lean_closure_set(x_63, 6, x_10); +lean_closure_set(x_63, 7, x_2); +lean_closure_set(x_63, 8, x_11); +lean_closure_set(x_63, 9, x_57); +x_64 = lean_apply_4(x_11, lean_box(0), lean_box(0), x_61, x_63); +return x_64; } } } -LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___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, uint8_t x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___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, uint8_t 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) { _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; -x_14 = lean_ctor_get(x_1, 1); +lean_object* x_14; lean_object* x_15; uint8_t x_16; lean_object* x_36; uint8_t x_37; +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_36 = l_Lean_shouldProfile___rarg___lambda__1___closed__2; +x_37 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_11, x_36); +if (x_37 == 0) +{ +uint8_t x_38; +x_38 = 0; +x_16 = x_38; +goto block_35; +} +else +{ +double x_39; double x_40; uint8_t x_41; +x_39 = l_Lean_trace_profiler_threshold_getSecs(x_11); +x_40 = lean_unbox_float(x_15); +x_41 = lean_float_decLt(x_39, x_40); +x_16 = x_41; +goto block_35; +} +block_35: +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_box(x_7); +x_18 = lean_box(x_16); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_4); +lean_inc(x_3); lean_inc(x_14); +x_19 = lean_alloc_closure((void*)(l_Lean_withTraceNodeBefore___rarg___lambda__2___boxed), 15, 14); +lean_closure_set(x_19, 0, x_1); +lean_closure_set(x_19, 1, x_14); +lean_closure_set(x_19, 2, x_2); +lean_closure_set(x_19, 3, x_3); +lean_closure_set(x_19, 4, x_4); +lean_closure_set(x_19, 5, x_5); +lean_closure_set(x_19, 6, x_6); +lean_closure_set(x_19, 7, x_17); +lean_closure_set(x_19, 8, x_8); +lean_closure_set(x_19, 9, x_9); +lean_closure_set(x_19, 10, x_10); +lean_closure_set(x_19, 11, x_11); +lean_closure_set(x_19, 12, x_18); +lean_closure_set(x_19, 13, x_15); +if (x_12 == 0) +{ +if (x_16 == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +lean_dec(x_19); +x_20 = lean_ctor_get(x_3, 0); +lean_inc(x_20); +lean_dec(x_3); +x_21 = lean_alloc_closure((void*)(l_Lean_PersistentArray_append___rarg), 2, 1); +lean_closure_set(x_21, 0, x_4); +x_22 = lean_apply_1(x_20, x_21); +lean_inc(x_10); +x_23 = lean_alloc_closure((void*)(l_Lean_withTraceNode___rarg___lambda__8___boxed), 5, 4); +lean_closure_set(x_23, 0, x_8); +lean_closure_set(x_23, 1, x_9); +lean_closure_set(x_23, 2, x_14); +lean_closure_set(x_23, 3, x_10); +x_24 = lean_apply_4(x_10, lean_box(0), lean_box(0), x_22, x_23); +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_dec(x_14); +lean_dec(x_9); +lean_dec(x_4); +lean_dec(x_3); +x_25 = lean_ctor_get(x_8, 0); +lean_inc(x_25); +lean_dec(x_8); +x_26 = lean_ctor_get(x_25, 1); +lean_inc(x_26); +lean_dec(x_25); +x_27 = lean_box(0); +x_28 = lean_apply_2(x_26, lean_box(0), x_27); +x_29 = lean_apply_4(x_10, lean_box(0), lean_box(0), x_28, x_19); +return x_29; +} +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +lean_dec(x_14); +lean_dec(x_9); +lean_dec(x_4); +lean_dec(x_3); +x_30 = lean_ctor_get(x_8, 0); +lean_inc(x_30); +lean_dec(x_8); +x_31 = lean_ctor_get(x_30, 1); +lean_inc(x_31); +lean_dec(x_30); +x_32 = lean_box(0); +x_33 = lean_apply_2(x_31, lean_box(0), x_32); +x_34 = lean_apply_4(x_10, lean_box(0), lean_box(0), x_33, x_19); +return x_34; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___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, uint8_t x_11, lean_object* x_12, uint8_t x_13, lean_object* x_14) { +_start: +{ +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_15 = lean_ctor_get(x_1, 1); +lean_inc(x_15); lean_inc(x_2); -x_15 = lean_alloc_closure((void*)(l_Lean_withTraceNode___rarg___lambda__1), 2, 1); -lean_closure_set(x_15, 0, x_2); +x_16 = lean_alloc_closure((void*)(l_Lean_withTraceNode___rarg___lambda__1), 2, 1); +lean_closure_set(x_16, 0, x_2); lean_inc(x_3); -x_16 = lean_apply_4(x_3, lean_box(0), lean_box(0), x_4, x_15); +x_17 = lean_apply_4(x_3, lean_box(0), lean_box(0), x_4, x_16); lean_inc(x_2); -x_17 = lean_alloc_closure((void*)(l_Lean_withTraceNode___rarg___lambda__2), 2, 1); -lean_closure_set(x_17, 0, x_2); -x_18 = lean_apply_3(x_14, lean_box(0), x_16, x_17); +x_18 = lean_alloc_closure((void*)(l_Lean_withTraceNode___rarg___lambda__2), 2, 1); +lean_closure_set(x_18, 0, x_2); +x_19 = lean_apply_3(x_15, lean_box(0), x_17, x_18); lean_inc(x_2); -x_19 = l_Lean_withOptProfile___rarg(x_2, x_5, x_6, x_18); -x_20 = lean_box(x_12); +x_20 = l_Lean_withSeconds___rarg(x_2, x_5, x_19); +x_21 = lean_box(x_11); +x_22 = lean_box(x_13); lean_inc(x_3); -x_21 = lean_alloc_closure((void*)(l_Lean_withTraceNodeBefore___rarg___lambda__2___boxed), 11, 10); -lean_closure_set(x_21, 0, x_7); -lean_closure_set(x_21, 1, x_13); -lean_closure_set(x_21, 2, x_8); -lean_closure_set(x_21, 3, x_9); -lean_closure_set(x_21, 4, x_10); -lean_closure_set(x_21, 5, x_11); -lean_closure_set(x_21, 6, x_20); -lean_closure_set(x_21, 7, x_2); -lean_closure_set(x_21, 8, x_1); -lean_closure_set(x_21, 9, x_3); -x_22 = lean_apply_4(x_3, lean_box(0), lean_box(0), x_19, x_21); -return x_22; +x_23 = lean_alloc_closure((void*)(l_Lean_withTraceNodeBefore___rarg___lambda__3___boxed), 13, 12); +lean_closure_set(x_23, 0, x_6); +lean_closure_set(x_23, 1, x_14); +lean_closure_set(x_23, 2, x_7); +lean_closure_set(x_23, 3, x_8); +lean_closure_set(x_23, 4, x_9); +lean_closure_set(x_23, 5, x_10); +lean_closure_set(x_23, 6, x_21); +lean_closure_set(x_23, 7, x_2); +lean_closure_set(x_23, 8, x_1); +lean_closure_set(x_23, 9, x_3); +lean_closure_set(x_23, 10, x_12); +lean_closure_set(x_23, 11, x_22); +x_24 = lean_apply_4(x_3, lean_box(0), lean_box(0), x_20, x_23); +return x_24; } } -LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___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, lean_object* x_14, uint8_t x_15, lean_object* x_16) { +LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___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, uint8_t x_14, lean_object* x_15, uint8_t x_16, lean_object* x_17) { _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_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_inc(x_2); -x_17 = lean_apply_4(x_2, lean_box(0), lean_box(0), x_3, x_1); -lean_inc(x_4); -x_18 = lean_alloc_closure((void*)(l___private_Lean_Util_Trace_0__Lean_addTraceNode___rarg___lambda__2___boxed), 4, 3); -lean_closure_set(x_18, 0, x_4); -lean_closure_set(x_18, 1, x_5); -lean_closure_set(x_18, 2, x_17); +x_18 = lean_apply_4(x_2, lean_box(0), lean_box(0), x_3, x_1); +lean_inc(x_17); +x_19 = lean_alloc_closure((void*)(l___private_Lean_Util_Trace_0__Lean_addTraceNode___rarg___lambda__2___boxed), 4, 3); +lean_closure_set(x_19, 0, x_17); +lean_closure_set(x_19, 1, x_4); +lean_closure_set(x_19, 2, x_18); lean_inc(x_2); -x_19 = lean_apply_4(x_2, lean_box(0), lean_box(0), x_6, x_18); -x_20 = lean_box(x_15); +x_20 = lean_apply_4(x_2, lean_box(0), lean_box(0), x_5, x_19); +x_21 = lean_box(x_14); +x_22 = lean_box(x_16); lean_inc(x_2); -x_21 = lean_alloc_closure((void*)(l_Lean_withTraceNodeBefore___rarg___lambda__3___boxed), 13, 12); -lean_closure_set(x_21, 0, x_7); -lean_closure_set(x_21, 1, x_8); -lean_closure_set(x_21, 2, x_2); -lean_closure_set(x_21, 3, x_9); -lean_closure_set(x_21, 4, x_10); -lean_closure_set(x_21, 5, x_11); -lean_closure_set(x_21, 6, x_12); -lean_closure_set(x_21, 7, x_13); -lean_closure_set(x_21, 8, x_16); -lean_closure_set(x_21, 9, x_14); -lean_closure_set(x_21, 10, x_4); -lean_closure_set(x_21, 11, x_20); -x_22 = lean_apply_4(x_2, lean_box(0), lean_box(0), x_19, x_21); -return x_22; -} -} -LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___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, uint8_t x_14, lean_object* x_15) { +x_23 = lean_alloc_closure((void*)(l_Lean_withTraceNodeBefore___rarg___lambda__4___boxed), 14, 13); +lean_closure_set(x_23, 0, x_6); +lean_closure_set(x_23, 1, x_7); +lean_closure_set(x_23, 2, x_2); +lean_closure_set(x_23, 3, x_8); +lean_closure_set(x_23, 4, x_9); +lean_closure_set(x_23, 5, x_10); +lean_closure_set(x_23, 6, x_11); +lean_closure_set(x_23, 7, x_12); +lean_closure_set(x_23, 8, x_13); +lean_closure_set(x_23, 9, x_17); +lean_closure_set(x_23, 10, x_21); +lean_closure_set(x_23, 11, x_15); +lean_closure_set(x_23, 12, x_22); +x_24 = lean_apply_4(x_2, lean_box(0), lean_box(0), x_20, x_23); +return x_24; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___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, uint8_t x_12, lean_object* x_13, uint8_t 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_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_16 = lean_ctor_get(x_1, 0); +lean_inc(x_16); +x_17 = lean_box(x_12); +x_18 = lean_box(x_14); +lean_inc(x_16); +lean_inc(x_3); +x_19 = lean_alloc_closure((void*)(l_Lean_withTraceNodeBefore___rarg___lambda__5___boxed), 17, 16); +lean_closure_set(x_19, 0, x_2); +lean_closure_set(x_19, 1, x_3); +lean_closure_set(x_19, 2, x_4); +lean_closure_set(x_19, 3, x_1); +lean_closure_set(x_19, 4, x_16); +lean_closure_set(x_19, 5, x_5); +lean_closure_set(x_19, 6, x_6); +lean_closure_set(x_19, 7, x_7); +lean_closure_set(x_19, 8, x_8); +lean_closure_set(x_19, 9, x_9); +lean_closure_set(x_19, 10, x_10); +lean_closure_set(x_19, 11, x_15); +lean_closure_set(x_19, 12, x_11); +lean_closure_set(x_19, 13, x_17); +lean_closure_set(x_19, 14, x_13); +lean_closure_set(x_19, 15, x_18); +x_20 = lean_apply_4(x_3, lean_box(0), lean_box(0), x_16, x_19); +return x_20; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__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, uint8_t x_12, lean_object* x_13, uint8_t 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_inc(x_2); lean_inc(x_1); x_16 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___rarg(x_1, x_2); -x_17 = lean_box(x_14); -lean_inc(x_4); -x_18 = lean_alloc_closure((void*)(l_Lean_withTraceNodeBefore___rarg___lambda__4___boxed), 16, 15); -lean_closure_set(x_18, 0, x_3); -lean_closure_set(x_18, 1, x_4); -lean_closure_set(x_18, 2, x_5); -lean_closure_set(x_18, 3, x_15); -lean_closure_set(x_18, 4, x_6); -lean_closure_set(x_18, 5, x_7); -lean_closure_set(x_18, 6, x_8); -lean_closure_set(x_18, 7, x_1); -lean_closure_set(x_18, 8, x_9); -lean_closure_set(x_18, 9, x_10); -lean_closure_set(x_18, 10, x_11); -lean_closure_set(x_18, 11, x_12); -lean_closure_set(x_18, 12, x_2); -lean_closure_set(x_18, 13, x_13); -lean_closure_set(x_18, 14, x_17); -x_19 = lean_apply_4(x_4, lean_box(0), lean_box(0), x_16, x_18); -return x_19; +x_17 = lean_box(x_12); +x_18 = lean_box(x_14); +lean_inc(x_5); +x_19 = lean_alloc_closure((void*)(l_Lean_withTraceNodeBefore___rarg___lambda__6___boxed), 15, 14); +lean_closure_set(x_19, 0, x_3); +lean_closure_set(x_19, 1, x_4); +lean_closure_set(x_19, 2, x_5); +lean_closure_set(x_19, 3, x_6); +lean_closure_set(x_19, 4, x_7); +lean_closure_set(x_19, 5, x_1); +lean_closure_set(x_19, 6, x_8); +lean_closure_set(x_19, 7, x_9); +lean_closure_set(x_19, 8, x_10); +lean_closure_set(x_19, 9, x_2); +lean_closure_set(x_19, 10, x_11); +lean_closure_set(x_19, 11, x_17); +lean_closure_set(x_19, 12, x_13); +lean_closure_set(x_19, 13, x_18); +x_20 = lean_apply_4(x_5, lean_box(0), lean_box(0), x_16, x_19); +return x_20; } } -LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___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, uint8_t x_13, uint8_t x_14) { +LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___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, uint8_t x_12, lean_object* x_13, uint8_t x_14) { _start: { +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_box(x_12); +x_16 = lean_box(x_14); +lean_inc(x_13); +lean_inc(x_8); +lean_inc(x_5); +lean_inc(x_1); +x_17 = lean_alloc_closure((void*)(l_Lean_withTraceNodeBefore___rarg___lambda__7___boxed), 15, 14); +lean_closure_set(x_17, 0, x_1); +lean_closure_set(x_17, 1, x_2); +lean_closure_set(x_17, 2, x_3); +lean_closure_set(x_17, 3, x_4); +lean_closure_set(x_17, 4, x_5); +lean_closure_set(x_17, 5, x_6); +lean_closure_set(x_17, 6, x_7); +lean_closure_set(x_17, 7, x_8); +lean_closure_set(x_17, 8, x_9); +lean_closure_set(x_17, 9, x_10); +lean_closure_set(x_17, 10, x_11); +lean_closure_set(x_17, 11, x_15); +lean_closure_set(x_17, 12, x_13); +lean_closure_set(x_17, 13, x_16); if (x_14 == 0) { -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); +lean_object* x_18; uint8_t x_19; +x_18 = l_Lean_shouldProfile___rarg___lambda__1___closed__2; +x_19 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_13, x_18); +lean_dec(x_13); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; +lean_dec(x_17); +x_20 = lean_alloc_closure((void*)(l_Lean_getTraces___rarg___lambda__1), 2, 1); +lean_closure_set(x_20, 0, x_1); +x_21 = lean_apply_4(x_5, lean_box(0), lean_box(0), x_8, x_20); +return x_21; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; 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_1; +x_22 = lean_ctor_get(x_1, 0); +lean_inc(x_22); +lean_dec(x_1); +x_23 = lean_ctor_get(x_22, 1); +lean_inc(x_23); +lean_dec(x_22); +x_24 = lean_box(0); +x_25 = lean_apply_2(x_23, lean_box(0), x_24); +x_26 = lean_apply_4(x_5, lean_box(0), lean_box(0), x_25, x_17); +return x_26; +} } else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_15 = lean_ctor_get(x_2, 0); -lean_inc(x_15); -x_16 = lean_box(x_13); -lean_inc(x_15); -lean_inc(x_6); -x_17 = lean_alloc_closure((void*)(l_Lean_withTraceNodeBefore___rarg___lambda__5___boxed), 15, 14); -lean_closure_set(x_17, 0, x_3); -lean_closure_set(x_17, 1, x_4); -lean_closure_set(x_17, 2, x_5); -lean_closure_set(x_17, 3, x_6); -lean_closure_set(x_17, 4, x_7); -lean_closure_set(x_17, 5, x_2); -lean_closure_set(x_17, 6, x_15); -lean_closure_set(x_17, 7, x_8); -lean_closure_set(x_17, 8, x_1); -lean_closure_set(x_17, 9, x_9); -lean_closure_set(x_17, 10, x_10); -lean_closure_set(x_17, 11, x_11); -lean_closure_set(x_17, 12, x_12); -lean_closure_set(x_17, 13, x_16); -x_18 = lean_apply_4(x_6, lean_box(0), lean_box(0), x_15, x_17); -return x_18; +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +lean_dec(x_13); +lean_dec(x_8); +x_27 = lean_ctor_get(x_1, 0); +lean_inc(x_27); +lean_dec(x_1); +x_28 = lean_ctor_get(x_27, 1); +lean_inc(x_28); +lean_dec(x_27); +x_29 = lean_box(0); +x_30 = lean_apply_2(x_28, lean_box(0), x_29); +x_31 = lean_apply_4(x_5, lean_box(0), lean_box(0), x_30, x_17); +return x_31; } } } -LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___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_object* x_12, lean_object* x_13, uint8_t x_14) { +LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___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, lean_object* x_12, lean_object* x_13, uint8_t x_14, lean_object* x_15) { _start: { -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_1, 1); -lean_inc(x_15); -lean_inc(x_11); -lean_inc(x_7); +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +lean_inc(x_4); lean_inc(x_1); -x_16 = l_Lean_isTracingEnabledFor___rarg(x_1, x_7, x_3, x_11); +x_16 = l_Lean_isTracingEnabledFor___rarg(x_1, x_2, x_3, x_4); x_17 = lean_box(x_14); -lean_inc(x_15); -x_18 = lean_alloc_closure((void*)(l_Lean_withTraceNodeBefore___rarg___lambda__6___boxed), 14, 13); -lean_closure_set(x_18, 0, x_13); +lean_inc(x_8); +x_18 = lean_alloc_closure((void*)(l_Lean_withTraceNodeBefore___rarg___lambda__8___boxed), 14, 13); +lean_closure_set(x_18, 0, x_1); lean_closure_set(x_18, 1, x_5); -lean_closure_set(x_18, 2, x_1); -lean_closure_set(x_18, 3, x_2); -lean_closure_set(x_18, 4, x_6); -lean_closure_set(x_18, 5, x_15); -lean_closure_set(x_18, 6, x_12); -lean_closure_set(x_18, 7, x_8); -lean_closure_set(x_18, 8, x_9); -lean_closure_set(x_18, 9, x_7); -lean_closure_set(x_18, 10, x_10); -lean_closure_set(x_18, 11, x_11); -lean_closure_set(x_18, 12, x_17); -x_19 = lean_apply_4(x_15, lean_box(0), lean_box(0), x_16, x_18); +lean_closure_set(x_18, 2, x_6); +lean_closure_set(x_18, 3, x_7); +lean_closure_set(x_18, 4, x_8); +lean_closure_set(x_18, 5, x_9); +lean_closure_set(x_18, 6, x_10); +lean_closure_set(x_18, 7, x_11); +lean_closure_set(x_18, 8, x_12); +lean_closure_set(x_18, 9, x_13); +lean_closure_set(x_18, 10, x_4); +lean_closure_set(x_18, 11, x_17); +lean_closure_set(x_18, 12, x_15); +x_19 = lean_apply_4(x_8, lean_box(0), lean_box(0), x_16, x_18); return x_19; } } +LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___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_object* x_12, lean_object* x_13, uint8_t x_14) { +_start: +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_15 = lean_ctor_get(x_1, 1); +lean_inc(x_15); +x_16 = lean_box(x_14); +lean_inc(x_15); +lean_inc(x_7); +x_17 = lean_alloc_closure((void*)(l_Lean_withTraceNodeBefore___rarg___lambda__9___boxed), 15, 14); +lean_closure_set(x_17, 0, x_1); +lean_closure_set(x_17, 1, x_7); +lean_closure_set(x_17, 2, x_3); +lean_closure_set(x_17, 3, x_11); +lean_closure_set(x_17, 4, x_2); +lean_closure_set(x_17, 5, x_5); +lean_closure_set(x_17, 6, x_6); +lean_closure_set(x_17, 7, x_15); +lean_closure_set(x_17, 8, x_12); +lean_closure_set(x_17, 9, x_8); +lean_closure_set(x_17, 10, x_13); +lean_closure_set(x_17, 11, x_9); +lean_closure_set(x_17, 12, x_10); +lean_closure_set(x_17, 13, x_16); +x_18 = lean_apply_4(x_15, lean_box(0), lean_box(0), x_7, x_17); +return x_18; +} +} LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore(lean_object* x_1, lean_object* x_2) { _start: { @@ -5568,55 +6380,119 @@ lean_dec(x_11); return x_13; } } -LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___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_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___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: { -uint8_t x_12; lean_object* x_13; -x_12 = lean_unbox(x_7); -lean_dec(x_7); -x_13 = l_Lean_withTraceNodeBefore___rarg___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_12, x_8, x_9, x_10, x_11); -return x_13; +uint8_t x_16; uint8_t x_17; double x_18; lean_object* x_19; +x_16 = lean_unbox(x_8); +lean_dec(x_8); +x_17 = lean_unbox(x_13); +lean_dec(x_13); +x_18 = lean_unbox_float(x_14); +lean_dec(x_14); +x_19 = l_Lean_withTraceNodeBefore___rarg___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_16, x_9, x_10, x_11, x_12, x_17, x_18, x_15); +lean_dec(x_15); +lean_dec(x_12); +return x_19; } } LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___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) { _start: { -uint8_t x_14; lean_object* x_15; -x_14 = lean_unbox(x_12); +uint8_t x_14; uint8_t x_15; lean_object* x_16; +x_14 = lean_unbox(x_7); +lean_dec(x_7); +x_15 = lean_unbox(x_12); lean_dec(x_12); -x_15 = l_Lean_withTraceNodeBefore___rarg___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_14, x_13); -return x_15; +x_16 = l_Lean_withTraceNodeBefore___rarg___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_14, x_8, x_9, x_10, x_11, x_15, x_13); +return x_16; } } -LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___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, lean_object* x_15, lean_object* x_16) { +LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___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: { -uint8_t x_17; lean_object* x_18; -x_17 = lean_unbox(x_15); +uint8_t x_15; uint8_t x_16; lean_object* x_17; +x_15 = lean_unbox(x_11); +lean_dec(x_11); +x_16 = lean_unbox(x_13); +lean_dec(x_13); +x_17 = l_Lean_withTraceNodeBefore___rarg___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_15, x_12, x_16, x_14); +return x_17; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__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]; +_start: +{ +uint8_t x_18; uint8_t x_19; lean_object* x_20; +x_18 = lean_unbox(x_14); +lean_dec(x_14); +x_19 = lean_unbox(x_16); +lean_dec(x_16); +x_20 = l_Lean_withTraceNodeBefore___rarg___lambda__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_18, x_15, x_19, x_17); +return x_20; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__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) { +_start: +{ +uint8_t x_16; uint8_t x_17; lean_object* x_18; +x_16 = lean_unbox(x_12); +lean_dec(x_12); +x_17 = lean_unbox(x_14); +lean_dec(x_14); +x_18 = l_Lean_withTraceNodeBefore___rarg___lambda__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_16, x_13, x_17, x_15); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__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: +{ +uint8_t x_16; uint8_t x_17; lean_object* x_18; +x_16 = lean_unbox(x_12); +lean_dec(x_12); +x_17 = lean_unbox(x_14); +lean_dec(x_14); +x_18 = l_Lean_withTraceNodeBefore___rarg___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_16, x_13, x_17, x_15); lean_dec(x_15); -x_18 = l_Lean_withTraceNodeBefore___rarg___lambda__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_17, x_16); return x_18; } } -LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___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) { +LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__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) { _start: { -uint8_t x_16; lean_object* x_17; +uint8_t x_15; uint8_t x_16; lean_object* x_17; +x_15 = lean_unbox(x_12); +lean_dec(x_12); x_16 = lean_unbox(x_14); lean_dec(x_14); -x_17 = l_Lean_withTraceNodeBefore___rarg___lambda__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_16, x_15); +x_17 = l_Lean_withTraceNodeBefore___rarg___lambda__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_15, x_13, x_16); return x_17; } } -LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___rarg___lambda__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_Lean_withTraceNodeBefore___rarg___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, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { -uint8_t x_15; uint8_t x_16; lean_object* x_17; -x_15 = lean_unbox(x_13); -lean_dec(x_13); +uint8_t x_16; lean_object* x_17; x_16 = lean_unbox(x_14); lean_dec(x_14); -x_17 = l_Lean_withTraceNodeBefore___rarg___lambda__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_15, x_16); +x_17 = l_Lean_withTraceNodeBefore___rarg___lambda__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_13, x_16, x_15); return x_17; } } @@ -5688,14 +6564,45 @@ lean_mark_persistent(l_Lean_isTracingEnabledFor___rarg___closed__1); l_Lean_withSeconds___rarg___lambda__1___closed__1 = _init_l_Lean_withSeconds___rarg___lambda__1___closed__1(); l_Lean_withSeconds___rarg___closed__1 = _init_l_Lean_withSeconds___rarg___closed__1(); lean_mark_persistent(l_Lean_withSeconds___rarg___closed__1); -l_Lean_withOptProfile___rarg___lambda__3___closed__1 = _init_l_Lean_withOptProfile___rarg___lambda__3___closed__1(); -lean_mark_persistent(l_Lean_withOptProfile___rarg___lambda__3___closed__1); -l_Lean_withOptProfile___rarg___lambda__3___closed__2 = _init_l_Lean_withOptProfile___rarg___lambda__3___closed__2(); -lean_mark_persistent(l_Lean_withOptProfile___rarg___lambda__3___closed__2); -l_Lean_withOptProfile___rarg___closed__1 = _init_l_Lean_withOptProfile___rarg___closed__1(); -lean_mark_persistent(l_Lean_withOptProfile___rarg___closed__1); -l_Lean_withOptProfile___rarg___closed__2 = _init_l_Lean_withOptProfile___rarg___closed__2(); -lean_mark_persistent(l_Lean_withOptProfile___rarg___closed__2); +l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__1 = _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__1); +l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__2 = _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__2); +l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__3 = _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__3); +l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__4 = _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__4(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__4); +l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__5 = _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__5(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__5); +l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__6 = _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__6(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382____closed__6); +if (builtin) {res = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1382_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +l_Lean_trace_profiler = lean_io_result_get_value(res); +lean_mark_persistent(l_Lean_trace_profiler); +lean_dec_ref(res); +}l_Lean_initFn____x40_Lean_Util_Trace___hyg_1478____closed__1 = _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1478____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Util_Trace___hyg_1478____closed__1); +l_Lean_initFn____x40_Lean_Util_Trace___hyg_1478____closed__2 = _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1478____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Util_Trace___hyg_1478____closed__2); +l_Lean_initFn____x40_Lean_Util_Trace___hyg_1478____closed__3 = _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1478____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Util_Trace___hyg_1478____closed__3); +l_Lean_initFn____x40_Lean_Util_Trace___hyg_1478____closed__4 = _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1478____closed__4(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Util_Trace___hyg_1478____closed__4); +l_Lean_initFn____x40_Lean_Util_Trace___hyg_1478____closed__5 = _init_l_Lean_initFn____x40_Lean_Util_Trace___hyg_1478____closed__5(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Util_Trace___hyg_1478____closed__5); +if (builtin) {res = l_Lean_initFn____x40_Lean_Util_Trace___hyg_1478_(lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +l_Lean_trace_profiler_threshold = lean_io_result_get_value(res); +lean_mark_persistent(l_Lean_trace_profiler_threshold); +lean_dec_ref(res); +}l_Lean_trace_profiler_threshold_getSecs___closed__1 = _init_l_Lean_trace_profiler_threshold_getSecs___closed__1(); +lean_mark_persistent(l_Lean_trace_profiler_threshold_getSecs___closed__1); +l_Lean_trace_profiler_threshold_getSecs___closed__2 = _init_l_Lean_trace_profiler_threshold_getSecs___closed__2(); +l_Lean_shouldProfile___rarg___lambda__1___closed__1 = _init_l_Lean_shouldProfile___rarg___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_shouldProfile___rarg___lambda__1___closed__1); +l_Lean_shouldProfile___rarg___lambda__1___closed__2 = _init_l_Lean_shouldProfile___rarg___lambda__1___closed__2(); +lean_mark_persistent(l_Lean_shouldProfile___rarg___lambda__1___closed__2); l_Lean_withTraceNode___rarg___lambda__5___closed__1 = _init_l_Lean_withTraceNode___rarg___lambda__5___closed__1(); lean_mark_persistent(l_Lean_withTraceNode___rarg___lambda__5___closed__1); l_Lean_withTraceNode___rarg___lambda__5___closed__2 = _init_l_Lean_withTraceNode___rarg___lambda__5___closed__2(); @@ -5710,64 +6617,62 @@ l_Lean_withTraceNode___rarg___lambda__5___closed__6 = _init_l_Lean_withTraceNode lean_mark_persistent(l_Lean_withTraceNode___rarg___lambda__5___closed__6); l_Lean_withTraceNode_x27___rarg___closed__1 = _init_l_Lean_withTraceNode_x27___rarg___closed__1(); lean_mark_persistent(l_Lean_withTraceNode_x27___rarg___closed__1); -l___auto____x40_Lean_Util_Trace___hyg_1952____closed__1 = _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__1(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_1952____closed__1); -l___auto____x40_Lean_Util_Trace___hyg_1952____closed__2 = _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__2(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_1952____closed__2); -l___auto____x40_Lean_Util_Trace___hyg_1952____closed__3 = _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__3(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_1952____closed__3); -l___auto____x40_Lean_Util_Trace___hyg_1952____closed__4 = _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__4(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_1952____closed__4); -l___auto____x40_Lean_Util_Trace___hyg_1952____closed__5 = _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__5(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_1952____closed__5); -l___auto____x40_Lean_Util_Trace___hyg_1952____closed__6 = _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__6(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_1952____closed__6); -l___auto____x40_Lean_Util_Trace___hyg_1952____closed__7 = _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__7(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_1952____closed__7); -l___auto____x40_Lean_Util_Trace___hyg_1952____closed__8 = _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__8(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_1952____closed__8); -l___auto____x40_Lean_Util_Trace___hyg_1952____closed__9 = _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__9(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_1952____closed__9); -l___auto____x40_Lean_Util_Trace___hyg_1952____closed__10 = _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__10(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_1952____closed__10); -l___auto____x40_Lean_Util_Trace___hyg_1952____closed__11 = _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__11(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_1952____closed__11); -l___auto____x40_Lean_Util_Trace___hyg_1952____closed__12 = _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__12(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_1952____closed__12); -l___auto____x40_Lean_Util_Trace___hyg_1952____closed__13 = _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__13(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_1952____closed__13); -l___auto____x40_Lean_Util_Trace___hyg_1952____closed__14 = _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__14(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_1952____closed__14); -l___auto____x40_Lean_Util_Trace___hyg_1952____closed__15 = _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__15(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_1952____closed__15); -l___auto____x40_Lean_Util_Trace___hyg_1952____closed__16 = _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__16(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_1952____closed__16); -l___auto____x40_Lean_Util_Trace___hyg_1952____closed__17 = _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__17(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_1952____closed__17); -l___auto____x40_Lean_Util_Trace___hyg_1952____closed__18 = _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__18(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_1952____closed__18); -l___auto____x40_Lean_Util_Trace___hyg_1952____closed__19 = _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__19(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_1952____closed__19); -l___auto____x40_Lean_Util_Trace___hyg_1952____closed__20 = _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__20(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_1952____closed__20); -l___auto____x40_Lean_Util_Trace___hyg_1952____closed__21 = _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__21(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_1952____closed__21); -l___auto____x40_Lean_Util_Trace___hyg_1952____closed__22 = _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__22(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_1952____closed__22); -l___auto____x40_Lean_Util_Trace___hyg_1952____closed__23 = _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__23(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_1952____closed__23); -l___auto____x40_Lean_Util_Trace___hyg_1952____closed__24 = _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__24(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_1952____closed__24); -l___auto____x40_Lean_Util_Trace___hyg_1952____closed__25 = _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__25(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_1952____closed__25); -l___auto____x40_Lean_Util_Trace___hyg_1952____closed__26 = _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__26(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_1952____closed__26); -l___auto____x40_Lean_Util_Trace___hyg_1952____closed__27 = _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__27(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_1952____closed__27); -l___auto____x40_Lean_Util_Trace___hyg_1952____closed__28 = _init_l___auto____x40_Lean_Util_Trace___hyg_1952____closed__28(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_1952____closed__28); -l___auto____x40_Lean_Util_Trace___hyg_1952_ = _init_l___auto____x40_Lean_Util_Trace___hyg_1952_(); -lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_1952_); +l___auto____x40_Lean_Util_Trace___hyg_2356____closed__1 = _init_l___auto____x40_Lean_Util_Trace___hyg_2356____closed__1(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2356____closed__1); +l___auto____x40_Lean_Util_Trace___hyg_2356____closed__2 = _init_l___auto____x40_Lean_Util_Trace___hyg_2356____closed__2(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2356____closed__2); +l___auto____x40_Lean_Util_Trace___hyg_2356____closed__3 = _init_l___auto____x40_Lean_Util_Trace___hyg_2356____closed__3(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2356____closed__3); +l___auto____x40_Lean_Util_Trace___hyg_2356____closed__4 = _init_l___auto____x40_Lean_Util_Trace___hyg_2356____closed__4(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2356____closed__4); +l___auto____x40_Lean_Util_Trace___hyg_2356____closed__5 = _init_l___auto____x40_Lean_Util_Trace___hyg_2356____closed__5(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2356____closed__5); +l___auto____x40_Lean_Util_Trace___hyg_2356____closed__6 = _init_l___auto____x40_Lean_Util_Trace___hyg_2356____closed__6(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2356____closed__6); +l___auto____x40_Lean_Util_Trace___hyg_2356____closed__7 = _init_l___auto____x40_Lean_Util_Trace___hyg_2356____closed__7(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2356____closed__7); +l___auto____x40_Lean_Util_Trace___hyg_2356____closed__8 = _init_l___auto____x40_Lean_Util_Trace___hyg_2356____closed__8(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2356____closed__8); +l___auto____x40_Lean_Util_Trace___hyg_2356____closed__9 = _init_l___auto____x40_Lean_Util_Trace___hyg_2356____closed__9(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2356____closed__9); +l___auto____x40_Lean_Util_Trace___hyg_2356____closed__10 = _init_l___auto____x40_Lean_Util_Trace___hyg_2356____closed__10(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2356____closed__10); +l___auto____x40_Lean_Util_Trace___hyg_2356____closed__11 = _init_l___auto____x40_Lean_Util_Trace___hyg_2356____closed__11(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2356____closed__11); +l___auto____x40_Lean_Util_Trace___hyg_2356____closed__12 = _init_l___auto____x40_Lean_Util_Trace___hyg_2356____closed__12(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2356____closed__12); +l___auto____x40_Lean_Util_Trace___hyg_2356____closed__13 = _init_l___auto____x40_Lean_Util_Trace___hyg_2356____closed__13(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2356____closed__13); +l___auto____x40_Lean_Util_Trace___hyg_2356____closed__14 = _init_l___auto____x40_Lean_Util_Trace___hyg_2356____closed__14(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2356____closed__14); +l___auto____x40_Lean_Util_Trace___hyg_2356____closed__15 = _init_l___auto____x40_Lean_Util_Trace___hyg_2356____closed__15(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2356____closed__15); +l___auto____x40_Lean_Util_Trace___hyg_2356____closed__16 = _init_l___auto____x40_Lean_Util_Trace___hyg_2356____closed__16(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2356____closed__16); +l___auto____x40_Lean_Util_Trace___hyg_2356____closed__17 = _init_l___auto____x40_Lean_Util_Trace___hyg_2356____closed__17(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2356____closed__17); +l___auto____x40_Lean_Util_Trace___hyg_2356____closed__18 = _init_l___auto____x40_Lean_Util_Trace___hyg_2356____closed__18(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2356____closed__18); +l___auto____x40_Lean_Util_Trace___hyg_2356____closed__19 = _init_l___auto____x40_Lean_Util_Trace___hyg_2356____closed__19(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2356____closed__19); +l___auto____x40_Lean_Util_Trace___hyg_2356____closed__20 = _init_l___auto____x40_Lean_Util_Trace___hyg_2356____closed__20(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2356____closed__20); +l___auto____x40_Lean_Util_Trace___hyg_2356____closed__21 = _init_l___auto____x40_Lean_Util_Trace___hyg_2356____closed__21(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2356____closed__21); +l___auto____x40_Lean_Util_Trace___hyg_2356____closed__22 = _init_l___auto____x40_Lean_Util_Trace___hyg_2356____closed__22(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2356____closed__22); +l___auto____x40_Lean_Util_Trace___hyg_2356____closed__23 = _init_l___auto____x40_Lean_Util_Trace___hyg_2356____closed__23(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2356____closed__23); +l___auto____x40_Lean_Util_Trace___hyg_2356____closed__24 = _init_l___auto____x40_Lean_Util_Trace___hyg_2356____closed__24(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2356____closed__24); +l___auto____x40_Lean_Util_Trace___hyg_2356____closed__25 = _init_l___auto____x40_Lean_Util_Trace___hyg_2356____closed__25(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2356____closed__25); +l___auto____x40_Lean_Util_Trace___hyg_2356____closed__26 = _init_l___auto____x40_Lean_Util_Trace___hyg_2356____closed__26(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2356____closed__26); +l___auto____x40_Lean_Util_Trace___hyg_2356____closed__27 = _init_l___auto____x40_Lean_Util_Trace___hyg_2356____closed__27(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2356____closed__27); +l___auto____x40_Lean_Util_Trace___hyg_2356_ = _init_l___auto____x40_Lean_Util_Trace___hyg_2356_(); +lean_mark_persistent(l___auto____x40_Lean_Util_Trace___hyg_2356_); l_Lean_registerTraceClass___closed__1 = _init_l_Lean_registerTraceClass___closed__1(); lean_mark_persistent(l_Lean_registerTraceClass___closed__1); l_Lean_registerTraceClass___closed__2 = _init_l_Lean_registerTraceClass___closed__2(); diff --git a/stage0/stdlib/Lean/Widget/Diff.c b/stage0/stdlib/Lean/Widget/Diff.c index 30fb10de5ca3..a9b206048abe 100644 --- a/stage0/stdlib/Lean/Widget/Diff.c +++ b/stage0/stdlib/Lean/Widget/Diff.c @@ -13,7 +13,6 @@ #ifdef __cplusplus extern "C" { #endif -uint8_t l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Widget_ExprDiffTag_toString___closed__3; LEAN_EXPORT lean_object* l_Lean_Widget_exprDiff___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_ExprDiff_withChange(lean_object*, lean_object*, uint8_t); @@ -108,7 +107,6 @@ LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Widget_exprDiffCore_piDiff__ LEAN_EXPORT lean_object* l___private_Lean_Widget_Diff_0__Lean_Widget_ExprDiffTag_toCtorIdx(uint8_t); LEAN_EXPORT lean_object* l_List_isSuffixOf_x3f___at_Lean_Widget_exprDiffCore_piDiff___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_diffHypotheses___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_diffInteractiveGoal___closed__6; LEAN_EXPORT lean_object* l_Lean_Widget_exprDiffCore___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_Lean_Widget_diffInteractiveGoal(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -132,6 +130,7 @@ static lean_object* l_Array_forInUnsafe_loop___at_Lean_Widget_diffHypothesesBund LEAN_EXPORT lean_object* l_Lean_Widget_exprDiffCore___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 uint8_t l_Lean_Widget_ExprDiffTag_toDiffTag(uint8_t, uint8_t); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_diffInteractiveGoals___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*); +uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); uint8_t l_Lean_LocalContext_contains(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_CodeWithInfos_mergePosMap___at_Lean_Widget_addDiffTags___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_ExprDiff_withChangePos___boxed(lean_object*, lean_object*, lean_object*); @@ -168,6 +167,7 @@ static lean_object* l_Lean_Widget_exprDiffCore_piDiff___lambda__3___closed__1; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_Widget_diffInteractiveGoals___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l___private_Lean_Expr_0__Lean_beqBinderInfo____x40_Lean_Expr___hyg_473_(uint8_t, uint8_t); LEAN_EXPORT lean_object* l_Lean_Widget_ExprDiff_withChange___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_getForallBinderNames(lean_object*); @@ -7414,7 +7414,7 @@ lean_inc(x_44); x_45 = lean_ctor_get(x_43, 1); lean_inc(x_45); lean_dec(x_43); -x_46 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_44, x_4, x_5, x_6, x_7, x_45); +x_46 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_44, x_4, x_5, x_6, x_7, x_45); x_47 = lean_ctor_get(x_46, 0); lean_inc(x_47); x_48 = lean_ctor_get(x_46, 1); @@ -7469,7 +7469,7 @@ lean_dec(x_53); x_61 = lean_ctor_get(x_60, 2); lean_inc(x_61); lean_dec(x_60); -x_62 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_61, x_4, x_5, x_6, x_7, x_51); +x_62 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_61, x_4, x_5, x_6, x_7, x_51); x_63 = lean_ctor_get(x_62, 0); lean_inc(x_63); x_64 = lean_ctor_get(x_62, 1); @@ -7699,7 +7699,7 @@ lean_inc(x_100); x_101 = lean_ctor_get(x_99, 1); lean_inc(x_101); lean_dec(x_99); -x_102 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_100, x_4, x_5, x_6, x_7, x_101); +x_102 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_100, x_4, x_5, x_6, x_7, x_101); x_103 = lean_ctor_get(x_102, 0); lean_inc(x_103); x_104 = lean_ctor_get(x_102, 1); @@ -7753,7 +7753,7 @@ lean_dec(x_109); x_117 = lean_ctor_get(x_116, 2); lean_inc(x_117); lean_dec(x_116); -x_118 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_117, x_4, x_5, x_6, x_7, x_107); +x_118 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_117, x_4, x_5, x_6, x_7, x_107); x_119 = lean_ctor_get(x_118, 0); lean_inc(x_119); x_120 = lean_ctor_get(x_118, 1); @@ -8003,7 +8003,7 @@ lean_inc(x_160); x_161 = lean_ctor_get(x_159, 1); lean_inc(x_161); lean_dec(x_159); -x_162 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_160, x_4, x_5, x_6, x_7, x_161); +x_162 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_160, x_4, x_5, x_6, x_7, x_161); x_163 = lean_ctor_get(x_162, 0); lean_inc(x_163); x_164 = lean_ctor_get(x_162, 1); @@ -8057,7 +8057,7 @@ lean_dec(x_169); x_177 = lean_ctor_get(x_176, 2); lean_inc(x_177); lean_dec(x_176); -x_178 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_177, x_4, x_5, x_6, x_7, x_167); +x_178 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_177, x_4, x_5, x_6, x_7, x_167); x_179 = lean_ctor_get(x_178, 0); lean_inc(x_179); x_180 = lean_ctor_get(x_178, 1); @@ -8839,7 +8839,7 @@ lean_object* x_9; lean_object* x_10; uint8_t x_11; x_9 = lean_ctor_get(x_6, 2); lean_inc(x_9); x_10 = l_Lean_Widget_diffInteractiveGoals___closed__1; -x_11 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_9, x_10); +x_11 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_9, x_10); lean_dec(x_9); if (x_11 == 0) { diff --git a/stage0/stdlib/Lean/Widget/InteractiveCode.c b/stage0/stdlib/Lean/Widget/InteractiveCode.c index e31907cdafed..3da0cba92413 100644 --- a/stage0/stdlib/Lean/Widget/InteractiveCode.c +++ b/stage0/stdlib/Lean/Widget/InteractiveCode.c @@ -13,19 +13,20 @@ #ifdef __cplusplus extern "C" { #endif +static lean_object* l_Lean_Widget_ppExprTagged___lambda__2___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonDiffTag____x40_Lean_Widget_InteractiveCode___hyg_79____lambda__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonDiffTag____x40_Lean_Widget_InteractiveCode___hyg_17____closed__1; -static lean_object* l_Lean_Widget_ppExprTagged___closed__5; +static lean_object* l_Lean_Widget_ppExprTagged___lambda__2___closed__3; lean_object* l_List_join___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_TaggedText_mapM___at_Lean_Widget_CodeWithInfos_mergePosMap___spec__2(lean_object*); static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonDiffTag____x40_Lean_Widget_InteractiveCode___hyg_79____closed__3; lean_object* lean_mk_empty_array_with_capacity(lean_object*); static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonDiffTag____x40_Lean_Widget_InteractiveCode___hyg_17____closed__2; LEAN_EXPORT lean_object* l_Lean_Widget_ppExprTagged(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Widget_ppExprTagged___closed__6; static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonDiffTag____x40_Lean_Widget_InteractiveCode___hyg_79____lambda__2___closed__3; LEAN_EXPORT lean_object* l_Lean_Widget_tagCodeInfos(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Json_mkObj(lean_object*); +static lean_object* l_Lean_Widget_ppExprTagged___lambda__2___closed__7; LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Widget_CodeWithInfos_mergePosMap___spec__1___rarg___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Widget_instFromJsonDiffTag___closed__1; static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_377____closed__4; @@ -46,6 +47,7 @@ LEAN_EXPORT lean_object* l_Lean_Widget_tagCodeInfos_go___lambda__1(lean_object*, LEAN_EXPORT lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonDiffTag____x40_Lean_Widget_InteractiveCode___hyg_79____lambda__1(lean_object*); static lean_object* l_Lean_Widget_instRpcEncodableSubexprInfo___closed__2; static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_377____closed__23; +LEAN_EXPORT lean_object* l_Lean_Widget_ppExprTagged___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonDiffTag____x40_Lean_Widget_InteractiveCode___hyg_17____closed__3; static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_377____closed__1; static lean_object* l_Lean_Widget_DiffTag_noConfusion___rarg___closed__1; @@ -54,7 +56,6 @@ extern lean_object* l_Lean_pp_explicit; LEAN_EXPORT lean_object* l_Lean_Widget_CodeWithInfos_mergePosMap___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableSubexprInfo_dec____x40_Lean_Widget_InteractiveCode___hyg_317_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_377____spec__1___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Widget_ppExprTagged___closed__8; LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Widget_CodeWithInfos_mergePosMap___spec__1(lean_object*); static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_377____closed__21; static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_377____closed__29; @@ -69,7 +70,7 @@ LEAN_EXPORT lean_object* l_Lean_Widget_ppExprTagged___lambda__1___boxed(lean_obj LEAN_EXPORT lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonDiffTag____x40_Lean_Widget_InteractiveCode___hyg_79____lambda__4___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_instToJsonRpcEncodablePacket; LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableSubexprInfo_dec____x40_Lean_Widget_InteractiveCode___hyg_317____spec__1(lean_object*, lean_object*); -static lean_object* l_Lean_Widget_ppExprTagged___closed__9; +static lean_object* l_Lean_Widget_ppExprTagged___lambda__2___closed__6; lean_object* l_Lean_Json_getStr_x3f(lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableSubexprInfo_dec____x40_Lean_Widget_InteractiveCode___hyg_317____boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonDiffTag____x40_Lean_Widget_InteractiveCode___hyg_17____closed__4; @@ -85,6 +86,7 @@ static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJ LEAN_EXPORT lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonDiffTag____x40_Lean_Widget_InteractiveCode___hyg_79____lambda__6(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_ppExprTagged___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_Widget_CodeWithInfos_mergePosMap___spec__3(lean_object*); +extern lean_object* l_Lean_pp_raw; static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonDiffTag____x40_Lean_Widget_InteractiveCode___hyg_79____closed__2; LEAN_EXPORT lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonDiffTag____x40_Lean_Widget_InteractiveCode___hyg_79____lambda__3___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_377____closed__19; @@ -93,6 +95,7 @@ static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJ LEAN_EXPORT lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonDiffTag____x40_Lean_Widget_InteractiveCode___hyg_79____lambda__5___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonDiffTag____x40_Lean_Widget_InteractiveCode___hyg_79____lambda__2(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Json_getObjValD(lean_object*, lean_object*); +static lean_object* l_Lean_Widget_ppExprTagged___lambda__2___closed__4; LEAN_EXPORT lean_object* l_Lean_Widget_TaggedText_mapM___at_Lean_Widget_CodeWithInfos_mergePosMap___spec__2___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_instToJsonDiffTag___closed__1; static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonDiffTag____x40_Lean_Widget_InteractiveCode___hyg_17____closed__9; @@ -114,11 +117,12 @@ LEAN_EXPORT lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_ static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_377____spec__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Widget_instFromJsonDiffTag; LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableSubexprInfo_dec____x40_Lean_Widget_InteractiveCode___hyg_317____spec__2(lean_object*, lean_object*); -static lean_object* l_Lean_Widget_ppExprTagged___closed__3; lean_object* l_Lean_PrettyPrinter_ppExprWithInfos(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_588_(lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +static lean_object* l_Lean_Widget_ppExprTagged___lambda__2___closed__9; LEAN_EXPORT lean_object* l_Lean_Widget_DiffTag_noConfusion(lean_object*); +uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_377_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_DiffTag_noConfusion___rarg___lambda__1(lean_object*); static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_377____closed__18; @@ -135,7 +139,6 @@ lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppExplicit(lean_object*, lea LEAN_EXPORT lean_object* l_Lean_Widget_SubexprInfo_withDiffTag(uint8_t, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_377____closed__24; static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_377____closed__35; -static lean_object* l_Lean_Widget_ppExprTagged___closed__2; static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_377____closed__12; static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonDiffTag____x40_Lean_Widget_InteractiveCode___hyg_79____lambda__5___closed__1; static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_377____closed__17; @@ -149,9 +152,7 @@ uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Widget_CodeWithInfos_mergePosMap___spec__1___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_588____spec__1(lean_object*, lean_object*); lean_object* l_Lean_Widget_TaggedText_rewrite___rarg(lean_object*, lean_object*); -static lean_object* l_Lean_Widget_ppExprTagged___closed__4; LEAN_EXPORT lean_object* l_Lean_Widget_DiffTag_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Widget_ppExprTagged___closed__7; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_377____closed__7; static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonDiffTag____x40_Lean_Widget_InteractiveCode___hyg_79____lambda__6___closed__1; @@ -164,6 +165,8 @@ LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Widget_ static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_377____closed__13; lean_object* l_Except_orElseLazy___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableSubexprInfo; +LEAN_EXPORT lean_object* l_Lean_Widget_ppExprTagged___lambda__2(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Widget_ppExprTagged___lambda__2___closed__2; static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonDiffTag____x40_Lean_Widget_InteractiveCode___hyg_79____lambda__3___closed__1; lean_object* l_Lean_PrettyPrinter_Delaborator_orElse___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_Widget_InteractiveCode_0__Lean_Widget_fromJsonDiffTag____x40_Lean_Widget_InteractiveCode___hyg_79____lambda__6___boxed(lean_object*, lean_object*, lean_object*); @@ -179,6 +182,7 @@ static lean_object* l_Lean_Widget_instRpcEncodableSubexprInfo___closed__1; static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonDiffTag____x40_Lean_Widget_InteractiveCode___hyg_79____lambda__1___closed__2; static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonDiffTag____x40_Lean_Widget_InteractiveCode___hyg_79____lambda__4___closed__1; static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_toJsonDiffTag____x40_Lean_Widget_InteractiveCode___hyg_17____closed__10; +static lean_object* l_Lean_Widget_ppExprTagged___lambda__2___closed__8; static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonDiffTag____x40_Lean_Widget_InteractiveCode___hyg_79____lambda__1___closed__1; lean_object* lean_string_append(lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); @@ -189,6 +193,7 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_CodeWithInfos_m extern lean_object* l_Lean_Widget_instTypeNameInfoWithCtx; uint8_t lean_usize_dec_lt(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Widget_CodeWithInfos_mergePosMap(lean_object*, lean_object*); +lean_object* lean_expr_dbg_to_string(lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_CodeWithInfos_mergePosMap___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_377____closed__3; LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableSubexprInfo_dec____x40_Lean_Widget_InteractiveCode___hyg_317____spec__1___boxed(lean_object*, lean_object*); @@ -198,6 +203,7 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_CodeWithInfos_m static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_377____closed__22; LEAN_EXPORT lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonDiffTag____x40_Lean_Widget_InteractiveCode___hyg_79____lambda__5(lean_object*, lean_object*, lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); +static lean_object* l_Lean_Widget_ppExprTagged___lambda__2___closed__5; lean_object* l_Lean_Widget_TaggedText_stripTags___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_DiffTag_noConfusion___rarg___lambda__1___boxed(lean_object*); static lean_object* l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_377____closed__14; @@ -2920,7 +2926,7 @@ x_9 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicit(x_2, x_3, x_4, x_5, x_6, return x_9; } } -static lean_object* _init_l_Lean_Widget_ppExprTagged___closed__1() { +static lean_object* _init_l_Lean_Widget_ppExprTagged___lambda__2___closed__1() { _start: { lean_object* x_1; @@ -2928,11 +2934,11 @@ x_1 = lean_mk_string_from_bytes("", 0); return x_1; } } -static lean_object* _init_l_Lean_Widget_ppExprTagged___closed__2() { +static lean_object* _init_l_Lean_Widget_ppExprTagged___lambda__2___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Widget_ppExprTagged___closed__1; +x_1 = l_Lean_Widget_ppExprTagged___lambda__2___closed__1; x_2 = l___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonDiffTag____x40_Lean_Widget_InteractiveCode___hyg_79____closed__1; x_3 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_3, 0, x_1); @@ -2941,7 +2947,7 @@ lean_ctor_set(x_3, 2, x_2); return x_3; } } -static lean_object* _init_l_Lean_Widget_ppExprTagged___closed__3() { +static lean_object* _init_l_Lean_Widget_ppExprTagged___lambda__2___closed__3() { _start: { lean_object* x_1; @@ -2949,7 +2955,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delab), 7, 0); return x_1; } } -static lean_object* _init_l_Lean_Widget_ppExprTagged___closed__4() { +static lean_object* _init_l_Lean_Widget_ppExprTagged___lambda__2___closed__4() { _start: { uint8_t x_1; lean_object* x_2; @@ -2959,7 +2965,7 @@ lean_ctor_set_uint8(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Widget_ppExprTagged___closed__5() { +static lean_object* _init_l_Lean_Widget_ppExprTagged___lambda__2___closed__5() { _start: { lean_object* x_1; @@ -2967,7 +2973,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabAppImplic return x_1; } } -static lean_object* _init_l_Lean_Widget_ppExprTagged___closed__6() { +static lean_object* _init_l_Lean_Widget_ppExprTagged___lambda__2___closed__6() { _start: { lean_object* x_1; @@ -2975,27 +2981,27 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Widget_ppExprTagged___lambda__1___boxed) return x_1; } } -static lean_object* _init_l_Lean_Widget_ppExprTagged___closed__7() { +static lean_object* _init_l_Lean_Widget_ppExprTagged___lambda__2___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Widget_ppExprTagged___closed__5; -x_2 = l_Lean_Widget_ppExprTagged___closed__6; +x_1 = l_Lean_Widget_ppExprTagged___lambda__2___closed__5; +x_2 = l_Lean_Widget_ppExprTagged___lambda__2___closed__6; x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_orElse___rarg), 9, 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_Widget_ppExprTagged___closed__8() { +static lean_object* _init_l_Lean_Widget_ppExprTagged___lambda__2___closed__8() { _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_pp_explicit; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l_Lean_Widget_ppExprTagged___closed__4; -x_4 = l_Lean_Widget_ppExprTagged___closed__7; +x_3 = l_Lean_Widget_ppExprTagged___lambda__2___closed__4; +x_4 = l_Lean_Widget_ppExprTagged___lambda__2___closed__7; x_5 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_withOptionAtCurrPos___rarg), 10, 3); lean_closure_set(x_5, 0, x_2); lean_closure_set(x_5, 1, x_3); @@ -3003,15 +3009,15 @@ lean_closure_set(x_5, 2, x_4); return x_5; } } -static lean_object* _init_l_Lean_Widget_ppExprTagged___closed__9() { +static lean_object* _init_l_Lean_Widget_ppExprTagged___lambda__2___closed__9() { _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_pp_tagAppFns; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l_Lean_Widget_ppExprTagged___closed__4; -x_4 = l_Lean_Widget_ppExprTagged___closed__8; +x_3 = l_Lean_Widget_ppExprTagged___lambda__2___closed__4; +x_4 = l_Lean_Widget_ppExprTagged___lambda__2___closed__8; x_5 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_withOptionAtCurrPos___rarg), 10, 3); lean_closure_set(x_5, 0, x_2); lean_closure_set(x_5, 1, x_3); @@ -3019,150 +3025,192 @@ lean_closure_set(x_5, 2, x_4); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Widget_ppExprTagged(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_EXPORT lean_object* l_Lean_Widget_ppExprTagged___lambda__2(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) { _start: { -lean_object* x_8; lean_object* x_9; -x_8 = lean_box(0); +lean_object* x_9; lean_object* x_10; +x_9 = lean_box(0); if (x_2 == 0) { -lean_object* x_48; -x_48 = l_Lean_Widget_ppExprTagged___closed__3; -x_9 = x_48; -goto block_47; +lean_object* x_49; +x_49 = l_Lean_Widget_ppExprTagged___lambda__2___closed__3; +x_10 = x_49; +goto block_48; } else { -lean_object* x_49; -x_49 = l_Lean_Widget_ppExprTagged___closed__9; -x_9 = x_49; -goto block_47; +lean_object* x_50; +x_50 = l_Lean_Widget_ppExprTagged___lambda__2___closed__9; +x_10 = x_50; +goto block_48; } -block_47: +block_48: { -lean_object* x_10; +lean_object* x_11; +lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_4); -x_10 = l_Lean_PrettyPrinter_ppExprWithInfos(x_1, x_8, x_9, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_10) == 0) +x_11 = l_Lean_PrettyPrinter_ppExprWithInfos(x_1, x_9, x_10, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_11) == 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; 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_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_10, 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; 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; uint8_t x_31; +x_12 = lean_ctor_get(x_11, 0); lean_inc(x_12); -lean_dec(x_10); -x_13 = lean_ctor_get(x_11, 0); +x_13 = lean_ctor_get(x_11, 1); lean_inc(x_13); -x_14 = lean_ctor_get(x_11, 1); -lean_inc(x_14); lean_dec(x_11); -x_15 = lean_unsigned_to_nat(0u); -x_16 = l_Std_Format_defWidth; -x_17 = l_Lean_Widget_TaggedText_prettyTagged(x_13, x_15, x_16); -x_18 = lean_st_ref_get(x_6, x_12); -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -x_20 = lean_ctor_get(x_18, 1); +x_14 = lean_ctor_get(x_12, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_12, 1); +lean_inc(x_15); +lean_dec(x_12); +x_16 = lean_unsigned_to_nat(0u); +x_17 = l_Std_Format_defWidth; +x_18 = l_Lean_Widget_TaggedText_prettyTagged(x_14, x_16, x_17); +x_19 = lean_st_ref_get(x_7, x_13); +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_st_ref_get(x_4, x_20); -lean_dec(x_4); -x_23 = lean_ctor_get(x_22, 0); -lean_inc(x_23); -x_24 = lean_ctor_get(x_22, 1); +x_22 = lean_ctor_get(x_20, 0); +lean_inc(x_22); +lean_dec(x_20); +x_23 = lean_st_ref_get(x_5, x_21); +lean_dec(x_5); +x_24 = lean_ctor_get(x_23, 0); lean_inc(x_24); -lean_dec(x_22); -x_25 = lean_ctor_get(x_23, 0); +x_25 = lean_ctor_get(x_23, 1); lean_inc(x_25); lean_dec(x_23); -x_26 = lean_ctor_get(x_5, 2); +x_26 = lean_ctor_get(x_24, 0); lean_inc(x_26); -x_27 = lean_ctor_get(x_5, 6); +lean_dec(x_24); +x_27 = lean_ctor_get(x_6, 2); lean_inc(x_27); -x_28 = lean_ctor_get(x_5, 7); +x_28 = lean_ctor_get(x_6, 6); lean_inc(x_28); -lean_dec(x_5); -x_29 = lean_st_ref_get(x_6, x_24); +x_29 = lean_ctor_get(x_6, 7); +lean_inc(x_29); lean_dec(x_6); -x_30 = !lean_is_exclusive(x_29); -if (x_30 == 0) +x_30 = lean_st_ref_get(x_7, x_25); +lean_dec(x_7); +x_31 = !lean_is_exclusive(x_30); +if (x_31 == 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_29, 0); -x_32 = lean_ctor_get(x_31, 2); -lean_inc(x_32); -lean_dec(x_31); -x_33 = l_Lean_Widget_ppExprTagged___closed__2; -x_34 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_34, 0, x_21); -lean_ctor_set(x_34, 1, x_33); -lean_ctor_set(x_34, 2, x_25); -lean_ctor_set(x_34, 3, x_26); -lean_ctor_set(x_34, 4, x_27); -lean_ctor_set(x_34, 5, x_28); -lean_ctor_set(x_34, 6, x_32); -x_35 = l_Lean_Widget_tagCodeInfos_go(x_34, x_14, x_17); -lean_ctor_set(x_29, 0, x_35); -return x_29; -} -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; -x_36 = lean_ctor_get(x_29, 0); -x_37 = lean_ctor_get(x_29, 1); -lean_inc(x_37); -lean_inc(x_36); -lean_dec(x_29); -x_38 = lean_ctor_get(x_36, 2); +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_30, 0); +x_33 = lean_ctor_get(x_32, 2); +lean_inc(x_33); +lean_dec(x_32); +x_34 = l_Lean_Widget_ppExprTagged___lambda__2___closed__2; +x_35 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_35, 0, x_22); +lean_ctor_set(x_35, 1, x_34); +lean_ctor_set(x_35, 2, x_26); +lean_ctor_set(x_35, 3, x_27); +lean_ctor_set(x_35, 4, x_28); +lean_ctor_set(x_35, 5, x_29); +lean_ctor_set(x_35, 6, x_33); +x_36 = l_Lean_Widget_tagCodeInfos_go(x_35, x_15, x_18); +lean_ctor_set(x_30, 0, x_36); +return x_30; +} +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; +x_37 = lean_ctor_get(x_30, 0); +x_38 = lean_ctor_get(x_30, 1); lean_inc(x_38); -lean_dec(x_36); -x_39 = l_Lean_Widget_ppExprTagged___closed__2; -x_40 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_40, 0, x_21); -lean_ctor_set(x_40, 1, x_39); -lean_ctor_set(x_40, 2, x_25); -lean_ctor_set(x_40, 3, x_26); -lean_ctor_set(x_40, 4, x_27); -lean_ctor_set(x_40, 5, x_28); -lean_ctor_set(x_40, 6, x_38); -x_41 = l_Lean_Widget_tagCodeInfos_go(x_40, x_14, x_17); -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 -{ -uint8_t x_43; +lean_inc(x_37); +lean_dec(x_30); +x_39 = lean_ctor_get(x_37, 2); +lean_inc(x_39); +lean_dec(x_37); +x_40 = l_Lean_Widget_ppExprTagged___lambda__2___closed__2; +x_41 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_41, 0, x_22); +lean_ctor_set(x_41, 1, x_40); +lean_ctor_set(x_41, 2, x_26); +lean_ctor_set(x_41, 3, x_27); +lean_ctor_set(x_41, 4, x_28); +lean_ctor_set(x_41, 5, x_29); +lean_ctor_set(x_41, 6, x_39); +x_42 = l_Lean_Widget_tagCodeInfos_go(x_41, x_15, x_18); +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; +} +} +else +{ +uint8_t x_44; +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); -x_43 = !lean_is_exclusive(x_10); -if (x_43 == 0) +x_44 = !lean_is_exclusive(x_11); +if (x_44 == 0) { -return x_10; +return x_11; } else { -lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_44 = lean_ctor_get(x_10, 0); -x_45 = lean_ctor_get(x_10, 1); +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_11, 0); +x_46 = lean_ctor_get(x_11, 1); +lean_inc(x_46); lean_inc(x_45); -lean_inc(x_44); -lean_dec(x_10); -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_dec(x_11); +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; +} +} } } } +static lean_object* _init_l_Lean_Widget_ppExprTagged___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_pp_raw; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Widget_ppExprTagged(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) { +_start: +{ +lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_8 = lean_ctor_get(x_5, 2); +lean_inc(x_8); +x_9 = l_Lean_Widget_ppExprTagged___closed__1; +x_10 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_8, x_9); +lean_dec(x_8); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_box(0); +x_12 = l_Lean_Widget_ppExprTagged___lambda__2(x_1, x_2, x_11, x_3, x_4, x_5, x_6, x_7); +return x_12; +} +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_13 = lean_expr_dbg_to_string(x_1); +lean_dec(x_1); +x_14 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_14, 0, 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_7); +return x_15; +} } } LEAN_EXPORT lean_object* l_Lean_Widget_ppExprTagged___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) { @@ -3174,6 +3222,17 @@ lean_dec(x_1); return x_9; } } +LEAN_EXPORT lean_object* l_Lean_Widget_ppExprTagged___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) { +_start: +{ +uint8_t x_9; lean_object* x_10; +x_9 = lean_unbox(x_2); +lean_dec(x_2); +x_10 = l_Lean_Widget_ppExprTagged___lambda__2(x_1, x_9, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_3); +return x_10; +} +} LEAN_EXPORT lean_object* l_Lean_Widget_ppExprTagged___boxed(lean_object* x_1, 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: { @@ -3363,24 +3422,26 @@ l_Lean_Widget_instRpcEncodableSubexprInfo = _init_l_Lean_Widget_instRpcEncodable lean_mark_persistent(l_Lean_Widget_instRpcEncodableSubexprInfo); l_Lean_Widget_tagCodeInfos_go___lambda__1___closed__1 = _init_l_Lean_Widget_tagCodeInfos_go___lambda__1___closed__1(); lean_mark_persistent(l_Lean_Widget_tagCodeInfos_go___lambda__1___closed__1); +l_Lean_Widget_ppExprTagged___lambda__2___closed__1 = _init_l_Lean_Widget_ppExprTagged___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Widget_ppExprTagged___lambda__2___closed__1); +l_Lean_Widget_ppExprTagged___lambda__2___closed__2 = _init_l_Lean_Widget_ppExprTagged___lambda__2___closed__2(); +lean_mark_persistent(l_Lean_Widget_ppExprTagged___lambda__2___closed__2); +l_Lean_Widget_ppExprTagged___lambda__2___closed__3 = _init_l_Lean_Widget_ppExprTagged___lambda__2___closed__3(); +lean_mark_persistent(l_Lean_Widget_ppExprTagged___lambda__2___closed__3); +l_Lean_Widget_ppExprTagged___lambda__2___closed__4 = _init_l_Lean_Widget_ppExprTagged___lambda__2___closed__4(); +lean_mark_persistent(l_Lean_Widget_ppExprTagged___lambda__2___closed__4); +l_Lean_Widget_ppExprTagged___lambda__2___closed__5 = _init_l_Lean_Widget_ppExprTagged___lambda__2___closed__5(); +lean_mark_persistent(l_Lean_Widget_ppExprTagged___lambda__2___closed__5); +l_Lean_Widget_ppExprTagged___lambda__2___closed__6 = _init_l_Lean_Widget_ppExprTagged___lambda__2___closed__6(); +lean_mark_persistent(l_Lean_Widget_ppExprTagged___lambda__2___closed__6); +l_Lean_Widget_ppExprTagged___lambda__2___closed__7 = _init_l_Lean_Widget_ppExprTagged___lambda__2___closed__7(); +lean_mark_persistent(l_Lean_Widget_ppExprTagged___lambda__2___closed__7); +l_Lean_Widget_ppExprTagged___lambda__2___closed__8 = _init_l_Lean_Widget_ppExprTagged___lambda__2___closed__8(); +lean_mark_persistent(l_Lean_Widget_ppExprTagged___lambda__2___closed__8); +l_Lean_Widget_ppExprTagged___lambda__2___closed__9 = _init_l_Lean_Widget_ppExprTagged___lambda__2___closed__9(); +lean_mark_persistent(l_Lean_Widget_ppExprTagged___lambda__2___closed__9); l_Lean_Widget_ppExprTagged___closed__1 = _init_l_Lean_Widget_ppExprTagged___closed__1(); lean_mark_persistent(l_Lean_Widget_ppExprTagged___closed__1); -l_Lean_Widget_ppExprTagged___closed__2 = _init_l_Lean_Widget_ppExprTagged___closed__2(); -lean_mark_persistent(l_Lean_Widget_ppExprTagged___closed__2); -l_Lean_Widget_ppExprTagged___closed__3 = _init_l_Lean_Widget_ppExprTagged___closed__3(); -lean_mark_persistent(l_Lean_Widget_ppExprTagged___closed__3); -l_Lean_Widget_ppExprTagged___closed__4 = _init_l_Lean_Widget_ppExprTagged___closed__4(); -lean_mark_persistent(l_Lean_Widget_ppExprTagged___closed__4); -l_Lean_Widget_ppExprTagged___closed__5 = _init_l_Lean_Widget_ppExprTagged___closed__5(); -lean_mark_persistent(l_Lean_Widget_ppExprTagged___closed__5); -l_Lean_Widget_ppExprTagged___closed__6 = _init_l_Lean_Widget_ppExprTagged___closed__6(); -lean_mark_persistent(l_Lean_Widget_ppExprTagged___closed__6); -l_Lean_Widget_ppExprTagged___closed__7 = _init_l_Lean_Widget_ppExprTagged___closed__7(); -lean_mark_persistent(l_Lean_Widget_ppExprTagged___closed__7); -l_Lean_Widget_ppExprTagged___closed__8 = _init_l_Lean_Widget_ppExprTagged___closed__8(); -lean_mark_persistent(l_Lean_Widget_ppExprTagged___closed__8); -l_Lean_Widget_ppExprTagged___closed__9 = _init_l_Lean_Widget_ppExprTagged___closed__9(); -lean_mark_persistent(l_Lean_Widget_ppExprTagged___closed__9); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lean/Widget/InteractiveDiagnostic.c b/stage0/stdlib/Lean/Widget/InteractiveDiagnostic.c index f7c1fd347433..fbf388ffdb24 100644 --- a/stage0/stdlib/Lean/Widget/InteractiveDiagnostic.c +++ b/stage0/stdlib/Lean/Widget/InteractiveDiagnostic.c @@ -19,7 +19,6 @@ static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget lean_object* l___private_Init_Util_0__outOfBounds___rarg(lean_object*); lean_object* l_List_join___rarg(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableMsgEmbed_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_593____spec__3(lean_object*, size_t, size_t, lean_object*, lean_object*); -lean_object* l_Lean_Syntax_decodeNameLit(lean_object*); lean_object* lean_format_pretty(lean_object*, lean_object*); static lean_object* l_Lean_Widget_instRpcEncodableStrictOrLazy_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_4____at_Lean_Widget_instRpcEncodableMsgEmbed_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_593____spec__4___closed__2; static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1810____closed__58; @@ -66,6 +65,7 @@ static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget lean_object* l_Lean_Widget_TaggedText_prettyTagged(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1810____closed__23; LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableDiagnosticWith_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_1726____spec__2___boxed(lean_object*, lean_object*); +uint8_t l_Lean_Name_isAnonymous(lean_object*); static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1810____closed__21; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableDiagnosticWith_enc____x40_Lean_Widget_InteractiveDiagnostic___hyg_1726____spec__2(size_t, size_t, lean_object*, lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); @@ -127,6 +127,7 @@ lean_object* l_Lean_Widget_goalToInteractive(lean_object*, lean_object*, lean_ob LEAN_EXPORT lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1041____boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_msgToInteractive_fmtToTT___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_msgToInteractive___closed__1; +LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableMsgEmbed_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_593____lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_msgToInteractive___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_88____lambda__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1810____closed__50; @@ -176,6 +177,7 @@ lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableDiagnosticWith_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_1726_(lean_object*); lean_object* l_Lean_Server_instRpcEncodableWithRpcRef_rpcDecode___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1810____closed__46; +LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableMsgEmbed_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_593____lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_msgToInteractiveDiagnostic___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_msgToInteractiveAux_go___lambda__1___closed__2; static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1810____closed__59; @@ -290,6 +292,7 @@ static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget static lean_object* l_Lean_Widget_instImpl____x40_Lean_Widget_InteractiveDiagnostic___hyg_574____closed__3; static lean_object* l_Lean_Widget_instRpcEncodableDiagnosticWith_enc____x40_Lean_Widget_InteractiveDiagnostic___hyg_1726____rarg___closed__6; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_msgToInteractive_fmtToTT___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); +lean_object* l_String_toName(lean_object*); lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_377____spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableMsgEmbed; LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableStrictOrLazy___rarg(lean_object*, lean_object*); @@ -310,7 +313,6 @@ static lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableDiag static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_88____lambda__2___closed__1; lean_object* l___private_Lean_Widget_TaggedText_0__Lean_Widget_fromJsonTaggedText____x40_Lean_Widget_TaggedText___hyg_432____at_Lean_Widget_TaggedText_instRpcEncodableTaggedText___spec__5(lean_object*); static lean_object* l_Lean_Widget_msgToInteractiveDiagnostic___closed__8; -static lean_object* l_Lean_Widget_instRpcEncodableMsgEmbed_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_593____closed__5; static lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_1810____closed__26; LEAN_EXPORT lean_object* l_Lean_Widget_msgToInteractive(lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableStrictOrLazy_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_4____rarg(lean_object*, lean_object*, lean_object*, lean_object*); @@ -3376,6 +3378,15 @@ return x_80; } } } +LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableMsgEmbed_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_593____lambda__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_3, 0, x_1); +return x_3; +} +} static lean_object* _init_l_Lean_Widget_instRpcEncodableMsgEmbed_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_593____closed__1() { _start: { @@ -3396,19 +3407,11 @@ static lean_object* _init_l_Lean_Widget_instRpcEncodableMsgEmbed_dec____x40_Lean _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("`", 1); -return x_1; -} -} -static lean_object* _init_l_Lean_Widget_instRpcEncodableMsgEmbed_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_593____closed__4() { -_start: -{ -lean_object* x_1; x_1 = lean_mk_string_from_bytes("expected a `Name`, got '", 24); return x_1; } } -static lean_object* _init_l_Lean_Widget_instRpcEncodableMsgEmbed_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_593____closed__5() { +static lean_object* _init_l_Lean_Widget_instRpcEncodableMsgEmbed_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_593____closed__4() { _start: { lean_object* x_1; lean_object* x_2; @@ -3668,99 +3671,90 @@ x_89 = l_Lean_Widget_instRpcEncodableMsgEmbed_dec____x40_Lean_Widget_Interactive x_90 = lean_string_dec_eq(x_88, x_89); if (x_90 == 0) { -lean_object* x_91; lean_object* x_92; lean_object* x_93; -x_91 = l_Lean_Widget_instRpcEncodableMsgEmbed_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_593____closed__3; -x_92 = lean_string_append(x_91, x_88); -lean_dec(x_88); -x_93 = l_Lean_Syntax_decodeNameLit(x_92); -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; lean_object* x_99; -x_94 = lean_unsigned_to_nat(80u); -x_95 = l_Lean_Json_pretty(x_37, x_94); -x_96 = l_Lean_Widget_instRpcEncodableMsgEmbed_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_593____closed__4; -x_97 = lean_string_append(x_96, x_95); -lean_dec(x_95); -x_98 = l_Lean_Widget_instRpcEncodableStrictOrLazy_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_4____at_Lean_Widget_instRpcEncodableMsgEmbed_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_593____spec__4___closed__2; -x_99 = lean_string_append(x_97, x_98); -lean_ctor_set_tag(x_83, 0); -lean_ctor_set(x_83, 0, x_99); +lean_object* x_91; uint8_t x_92; +x_91 = l_String_toName(x_88); +x_92 = l_Lean_Name_isAnonymous(x_91); +if (x_92 == 0) +{ +lean_dec(x_37); +lean_ctor_set(x_83, 0, x_91); x_47 = x_83; goto block_82; } else { -lean_object* x_100; -lean_dec(x_37); -x_100 = lean_ctor_get(x_93, 0); -lean_inc(x_100); -lean_dec(x_93); -lean_ctor_set(x_83, 0, x_100); +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_91); +x_93 = lean_unsigned_to_nat(80u); +x_94 = l_Lean_Json_pretty(x_37, x_93); +x_95 = l_Lean_Widget_instRpcEncodableMsgEmbed_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_593____closed__3; +x_96 = lean_string_append(x_95, x_94); +lean_dec(x_94); +x_97 = l_Lean_Widget_instRpcEncodableStrictOrLazy_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_4____at_Lean_Widget_instRpcEncodableMsgEmbed_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_593____spec__4___closed__2; +x_98 = lean_string_append(x_96, x_97); +lean_ctor_set_tag(x_83, 0); +lean_ctor_set(x_83, 0, x_98); x_47 = x_83; goto block_82; } } else { -lean_object* x_101; +lean_object* x_99; lean_free_object(x_83); lean_dec(x_88); lean_dec(x_37); -x_101 = l_Lean_Widget_instRpcEncodableMsgEmbed_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_593____closed__5; -x_47 = x_101; +x_99 = l_Lean_Widget_instRpcEncodableMsgEmbed_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_593____closed__4; +x_47 = x_99; goto block_82; } } else { -lean_object* x_102; lean_object* x_103; uint8_t x_104; -x_102 = lean_ctor_get(x_83, 0); -lean_inc(x_102); +lean_object* x_100; lean_object* x_101; uint8_t x_102; +x_100 = lean_ctor_get(x_83, 0); +lean_inc(x_100); lean_dec(x_83); -x_103 = l_Lean_Widget_instRpcEncodableMsgEmbed_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_593____closed__2; -x_104 = lean_string_dec_eq(x_102, x_103); +x_101 = l_Lean_Widget_instRpcEncodableMsgEmbed_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_593____closed__2; +x_102 = lean_string_dec_eq(x_100, x_101); +if (x_102 == 0) +{ +lean_object* x_103; uint8_t x_104; +x_103 = l_String_toName(x_100); +x_104 = l_Lean_Name_isAnonymous(x_103); if (x_104 == 0) { -lean_object* x_105; lean_object* x_106; lean_object* x_107; -x_105 = l_Lean_Widget_instRpcEncodableMsgEmbed_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_593____closed__3; -x_106 = lean_string_append(x_105, x_102); -lean_dec(x_102); -x_107 = l_Lean_Syntax_decodeNameLit(x_106); -if (lean_obj_tag(x_107) == 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 = lean_unsigned_to_nat(80u); -x_109 = l_Lean_Json_pretty(x_37, x_108); -x_110 = l_Lean_Widget_instRpcEncodableMsgEmbed_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_593____closed__4; -x_111 = lean_string_append(x_110, x_109); -lean_dec(x_109); -x_112 = l_Lean_Widget_instRpcEncodableStrictOrLazy_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_4____at_Lean_Widget_instRpcEncodableMsgEmbed_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_593____spec__4___closed__2; -x_113 = lean_string_append(x_111, x_112); -x_114 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_114, 0, x_113); -x_47 = x_114; +lean_object* x_105; +lean_dec(x_37); +x_105 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_105, 0, x_103); +x_47 = x_105; goto block_82; } else { -lean_object* x_115; lean_object* x_116; -lean_dec(x_37); -x_115 = lean_ctor_get(x_107, 0); -lean_inc(x_115); +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_103); +x_106 = lean_unsigned_to_nat(80u); +x_107 = l_Lean_Json_pretty(x_37, x_106); +x_108 = l_Lean_Widget_instRpcEncodableMsgEmbed_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_593____closed__3; +x_109 = lean_string_append(x_108, x_107); lean_dec(x_107); -x_116 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_116, 0, x_115); -x_47 = x_116; +x_110 = l_Lean_Widget_instRpcEncodableStrictOrLazy_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_4____at_Lean_Widget_instRpcEncodableMsgEmbed_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_593____spec__4___closed__2; +x_111 = lean_string_append(x_109, x_110); +x_112 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_112, 0, x_111); +x_47 = x_112; goto block_82; } } else { -lean_object* x_117; -lean_dec(x_102); +lean_object* x_113; +lean_dec(x_100); lean_dec(x_37); -x_117 = l_Lean_Widget_instRpcEncodableMsgEmbed_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_593____closed__5; -x_47 = x_117; +x_113 = l_Lean_Widget_instRpcEncodableMsgEmbed_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_593____closed__4; +x_47 = x_113; goto block_82; } } @@ -4026,6 +4020,15 @@ x_7 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableMsgEmbed_dec____x4 return x_7; } } +LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableMsgEmbed_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_593____lambda__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_Widget_instRpcEncodableMsgEmbed_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_593____lambda__1(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} static lean_object* _init_l_Lean_Widget_instRpcEncodableMsgEmbed___closed__1() { _start: { @@ -21061,8 +21064,6 @@ l_Lean_Widget_instRpcEncodableMsgEmbed_dec____x40_Lean_Widget_InteractiveDiagnos lean_mark_persistent(l_Lean_Widget_instRpcEncodableMsgEmbed_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_593____closed__3); l_Lean_Widget_instRpcEncodableMsgEmbed_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_593____closed__4 = _init_l_Lean_Widget_instRpcEncodableMsgEmbed_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_593____closed__4(); lean_mark_persistent(l_Lean_Widget_instRpcEncodableMsgEmbed_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_593____closed__4); -l_Lean_Widget_instRpcEncodableMsgEmbed_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_593____closed__5 = _init_l_Lean_Widget_instRpcEncodableMsgEmbed_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_593____closed__5(); -lean_mark_persistent(l_Lean_Widget_instRpcEncodableMsgEmbed_dec____x40_Lean_Widget_InteractiveDiagnostic___hyg_593____closed__5); l_Lean_Widget_instRpcEncodableMsgEmbed___closed__1 = _init_l_Lean_Widget_instRpcEncodableMsgEmbed___closed__1(); lean_mark_persistent(l_Lean_Widget_instRpcEncodableMsgEmbed___closed__1); l_Lean_Widget_instRpcEncodableMsgEmbed = _init_l_Lean_Widget_instRpcEncodableMsgEmbed(); diff --git a/stage0/stdlib/Lean/Widget/InteractiveGoal.c b/stage0/stdlib/Lean/Widget/InteractiveGoal.c index ea049a2cd3d7..b30fb645c651 100644 --- a/stage0/stdlib/Lean/Widget/InteractiveGoal.c +++ b/stage0/stdlib/Lean/Widget/InteractiveGoal.c @@ -14,13 +14,11 @@ extern "C" { #endif static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_2609____closed__9; -uint8_t l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_2609_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableInteractiveGoal_dec____x40_Lean_Widget_InteractiveGoal___hyg_1387____boxed(lean_object*, lean_object*); lean_object* l_Lean_Meta_getGoalPrefix(lean_object*); static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_2609____closed__18; lean_object* l_List_join___rarg(lean_object*); -lean_object* l_Lean_Syntax_decodeNameLit(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_Widget_goalToInteractive___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_Array_forInUnsafe_loop___at_Lean_Widget_InteractiveGoalCore_pretty___spec__3___closed__2; static lean_object* l_Lean_Widget_withGoalCtx___rarg___lambda__2___closed__3; @@ -47,6 +45,7 @@ LEAN_EXPORT lean_object* l_Lean_Widget_withGoalCtx___at_Lean_Widget_goalToIntera LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_enc____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__1(size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_InteractiveGoalCore_pretty___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_instRpcEncodableInteractiveGoals___closed__2; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableInteractiveGoal_dec____x40_Lean_Widget_InteractiveGoal___hyg_1387_(lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_184____closed__12; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Widget_goalToInteractive___spec__4___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*); @@ -127,7 +126,6 @@ static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJ static lean_object* l_Lean_Widget_instInhabitedInteractiveHypothesisBundle___closed__4; static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_1514____closed__2; static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_184____closed__49; -lean_object* l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_1514____closed__35; lean_object* l_Lean_Json_getStr_x3f(lean_object*); static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_184____closed__36; @@ -198,7 +196,6 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodabl LEAN_EXPORT lean_object* l_Lean_Widget_InteractiveGoals_append(lean_object*, lean_object*); lean_object* lean_array_to_list(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_TaggedText_mapM___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__5(lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveGoal_dec____x40_Lean_Widget_InteractiveGoal___hyg_1387____spec__2___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_1514____closed__12; @@ -228,6 +225,7 @@ LEAN_EXPORT lean_object* l_List_filterTR_loop___at_Lean_Widget_InteractiveGoalCo static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_184____closed__25; static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_184____closed__58; static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_184____closed__50; +uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_184____closed__30; static lean_object* l_Lean_Widget_instInhabitedInteractiveHypothesisBundle___closed__1; static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_1514____closed__18; @@ -271,6 +269,7 @@ LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodabl LEAN_EXPORT lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_184____boxed(lean_object*); static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_3717____closed__7; static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_184____closed__44; +lean_object* l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_enc____x40_Lean_Widget_InteractiveGoal___hyg_5____closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveGoal_dec____x40_Lean_Widget_InteractiveGoal___hyg_1387____spec__1(size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableInteractiveGoals_dec____x40_Lean_Widget_InteractiveGoal___hyg_3693____boxed(lean_object*, lean_object*); @@ -297,9 +296,11 @@ static lean_object* l_Array_forInUnsafe_loop___at_Lean_Widget_InteractiveGoalCor LEAN_EXPORT lean_object* l_Lean_Widget_instFromJsonRpcEncodablePacket__2; static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_3717____closed__6; static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_184____closed__26; +lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveGoal_enc____x40_Lean_Widget_InteractiveGoal___hyg_1387____spec__1(size_t, size_t, lean_object*, lean_object*); static lean_object* l_Lean_Widget_InteractiveGoalCore_pretty___closed__1; static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_184____closed__35; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___lambda__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_184____closed__7; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Widget_goalToInteractive___spec__4___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_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__7___boxed(lean_object*, lean_object*); @@ -312,11 +313,11 @@ lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_184____closed__20; 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_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_String_toName(lean_object*); static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_1514____closed__28; lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_377____spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_184____closed__28; static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_1514____closed__21; -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__5; LEAN_EXPORT lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_2925____boxed(lean_object*); static lean_object* l___private_Lean_Widget_InteractiveGoal_0__Lean_Widget_fromJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveGoal___hyg_184____closed__32; LEAN_EXPORT lean_object* l_Lean_Widget_instRpcEncodableInteractiveGoals_enc____x40_Lean_Widget_InteractiveGoal___hyg_3693_(lean_object*, lean_object*); @@ -2177,23 +2178,24 @@ return x_8; } } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__1() { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___lambda__1(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("[anonymous]", 11); -return x_1; +lean_object* x_3; +x_3 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_3, 0, x_1); +return x_3; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__2() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("`", 1); +x_1 = lean_mk_string_from_bytes("[anonymous]", 11); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__3() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__2() { _start: { lean_object* x_1; @@ -2201,7 +2203,7 @@ x_1 = lean_mk_string_from_bytes("expected a `Name`, got '", 24); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__4() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__3() { _start: { lean_object* x_1; @@ -2209,7 +2211,7 @@ x_1 = lean_mk_string_from_bytes("'", 1); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__5() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__4() { _start: { lean_object* x_1; lean_object* x_2; @@ -2338,293 +2340,285 @@ x_34 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypoth x_35 = lean_string_dec_eq(x_33, x_34); if (x_35 == 0) { -lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_36 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__2; -x_37 = lean_string_append(x_36, x_33); -lean_dec(x_33); -x_38 = l_Lean_Syntax_decodeNameLit(x_37); -if (lean_obj_tag(x_38) == 0) +lean_object* x_36; uint8_t x_37; +x_36 = l_String_toName(x_33); +x_37 = l_Lean_Name_isAnonymous(x_36); +if (x_37 == 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; -x_39 = lean_unsigned_to_nat(80u); -x_40 = l_Lean_Json_pretty(x_7, x_39); -x_41 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__3; -x_42 = lean_string_append(x_41, x_40); -lean_dec(x_40); -x_43 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__4; -x_44 = lean_string_append(x_42, x_43); -lean_ctor_set_tag(x_10, 0); -lean_ctor_set(x_10, 0, x_44); -x_45 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__1(x_10, x_4); -if (lean_obj_tag(x_45) == 0) +lean_object* x_38; +lean_dec(x_7); +lean_ctor_set(x_10, 0, x_36); +x_38 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__1(x_10, x_4); +if (lean_obj_tag(x_38) == 0) { -uint8_t x_46; +uint8_t x_39; lean_dec(x_9); -x_46 = !lean_is_exclusive(x_45); -if (x_46 == 0) +x_39 = !lean_is_exclusive(x_38); +if (x_39 == 0) { -return x_45; +return x_38; } else { -lean_object* x_47; lean_object* x_48; -x_47 = lean_ctor_get(x_45, 0); -lean_inc(x_47); -lean_dec(x_45); -x_48 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_48, 0, x_47); -return x_48; +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_alloc_ctor(0, 1, 0); +lean_ctor_set(x_41, 0, x_40); +return x_41; } } else { -lean_object* x_49; size_t x_50; size_t x_51; lean_object* x_52; -x_49 = lean_ctor_get(x_45, 0); -lean_inc(x_49); -lean_dec(x_45); -x_50 = 1; -x_51 = lean_usize_add(x_2, x_50); -x_52 = lean_array_uset(x_9, x_2, x_49); -x_2 = x_51; -x_3 = x_52; +lean_object* x_42; size_t x_43; size_t x_44; lean_object* x_45; +x_42 = lean_ctor_get(x_38, 0); +lean_inc(x_42); +lean_dec(x_38); +x_43 = 1; +x_44 = lean_usize_add(x_2, x_43); +x_45 = lean_array_uset(x_9, x_2, x_42); +x_2 = x_44; +x_3 = x_45; goto _start; } } else { -lean_object* x_54; lean_object* x_55; -lean_dec(x_7); -x_54 = lean_ctor_get(x_38, 0); -lean_inc(x_54); -lean_dec(x_38); -lean_ctor_set(x_10, 0, x_54); -x_55 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__1(x_10, x_4); -if (lean_obj_tag(x_55) == 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_dec(x_36); +x_47 = lean_unsigned_to_nat(80u); +x_48 = l_Lean_Json_pretty(x_7, x_47); +x_49 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__2; +x_50 = lean_string_append(x_49, x_48); +lean_dec(x_48); +x_51 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__3; +x_52 = lean_string_append(x_50, x_51); +lean_ctor_set_tag(x_10, 0); +lean_ctor_set(x_10, 0, x_52); +x_53 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__1(x_10, x_4); +if (lean_obj_tag(x_53) == 0) { -uint8_t x_56; +uint8_t x_54; lean_dec(x_9); -x_56 = !lean_is_exclusive(x_55); -if (x_56 == 0) +x_54 = !lean_is_exclusive(x_53); +if (x_54 == 0) { -return x_55; +return x_53; } else { -lean_object* x_57; lean_object* x_58; -x_57 = lean_ctor_get(x_55, 0); -lean_inc(x_57); -lean_dec(x_55); -x_58 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_58, 0, x_57); -return x_58; +lean_object* x_55; lean_object* x_56; +x_55 = lean_ctor_get(x_53, 0); +lean_inc(x_55); +lean_dec(x_53); +x_56 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_56, 0, x_55); +return x_56; } } else { -lean_object* x_59; size_t x_60; size_t x_61; lean_object* x_62; -x_59 = lean_ctor_get(x_55, 0); -lean_inc(x_59); -lean_dec(x_55); -x_60 = 1; -x_61 = lean_usize_add(x_2, x_60); -x_62 = lean_array_uset(x_9, x_2, x_59); -x_2 = x_61; -x_3 = x_62; +lean_object* x_57; size_t x_58; size_t x_59; lean_object* x_60; +x_57 = lean_ctor_get(x_53, 0); +lean_inc(x_57); +lean_dec(x_53); +x_58 = 1; +x_59 = lean_usize_add(x_2, x_58); +x_60 = lean_array_uset(x_9, x_2, x_57); +x_2 = x_59; +x_3 = x_60; goto _start; } } } else { -lean_object* x_64; lean_object* x_65; +lean_object* x_62; lean_object* x_63; lean_free_object(x_10); lean_dec(x_33); lean_dec(x_7); -x_64 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__5; -x_65 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__1(x_64, x_4); -if (lean_obj_tag(x_65) == 0) +x_62 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__4; +x_63 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__1(x_62, x_4); +if (lean_obj_tag(x_63) == 0) { -uint8_t x_66; +uint8_t x_64; lean_dec(x_9); -x_66 = !lean_is_exclusive(x_65); -if (x_66 == 0) +x_64 = !lean_is_exclusive(x_63); +if (x_64 == 0) { -return x_65; +return x_63; } else { -lean_object* x_67; lean_object* x_68; -x_67 = lean_ctor_get(x_65, 0); -lean_inc(x_67); -lean_dec(x_65); -x_68 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_68, 0, x_67); -return x_68; +lean_object* x_65; lean_object* x_66; +x_65 = lean_ctor_get(x_63, 0); +lean_inc(x_65); +lean_dec(x_63); +x_66 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_66, 0, x_65); +return x_66; } } else { -lean_object* x_69; size_t x_70; size_t x_71; lean_object* x_72; -x_69 = lean_ctor_get(x_65, 0); -lean_inc(x_69); -lean_dec(x_65); -x_70 = 1; -x_71 = lean_usize_add(x_2, x_70); -x_72 = lean_array_uset(x_9, x_2, x_69); -x_2 = x_71; -x_3 = x_72; +lean_object* x_67; size_t x_68; size_t x_69; lean_object* x_70; +x_67 = lean_ctor_get(x_63, 0); +lean_inc(x_67); +lean_dec(x_63); +x_68 = 1; +x_69 = lean_usize_add(x_2, x_68); +x_70 = lean_array_uset(x_9, x_2, x_67); +x_2 = x_69; +x_3 = x_70; goto _start; } } } else { -lean_object* x_74; lean_object* x_75; uint8_t x_76; -x_74 = lean_ctor_get(x_10, 0); -lean_inc(x_74); +lean_object* x_72; lean_object* x_73; uint8_t x_74; +x_72 = lean_ctor_get(x_10, 0); +lean_inc(x_72); lean_dec(x_10); -x_75 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__1; -x_76 = lean_string_dec_eq(x_74, x_75); +x_73 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__1; +x_74 = lean_string_dec_eq(x_72, x_73); +if (x_74 == 0) +{ +lean_object* x_75; uint8_t x_76; +x_75 = l_String_toName(x_72); +x_76 = l_Lean_Name_isAnonymous(x_75); if (x_76 == 0) { -lean_object* x_77; lean_object* x_78; lean_object* x_79; -x_77 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__2; -x_78 = lean_string_append(x_77, x_74); -lean_dec(x_74); -x_79 = l_Lean_Syntax_decodeNameLit(x_78); -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_85; lean_object* x_86; lean_object* x_87; -x_80 = lean_unsigned_to_nat(80u); -x_81 = l_Lean_Json_pretty(x_7, x_80); -x_82 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__3; -x_83 = lean_string_append(x_82, x_81); -lean_dec(x_81); -x_84 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__4; -x_85 = lean_string_append(x_83, x_84); -x_86 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_86, 0, x_85); -x_87 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__1(x_86, x_4); -if (lean_obj_tag(x_87) == 0) +lean_object* x_77; lean_object* x_78; +lean_dec(x_7); +x_77 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_77, 0, x_75); +x_78 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__1(x_77, x_4); +if (lean_obj_tag(x_78) == 0) { -lean_object* x_88; lean_object* x_89; lean_object* x_90; +lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_dec(x_9); -x_88 = lean_ctor_get(x_87, 0); -lean_inc(x_88); -if (lean_is_exclusive(x_87)) { - lean_ctor_release(x_87, 0); - x_89 = x_87; +x_79 = lean_ctor_get(x_78, 0); +lean_inc(x_79); +if (lean_is_exclusive(x_78)) { + lean_ctor_release(x_78, 0); + x_80 = x_78; } else { - lean_dec_ref(x_87); - x_89 = lean_box(0); + lean_dec_ref(x_78); + x_80 = lean_box(0); } -if (lean_is_scalar(x_89)) { - x_90 = lean_alloc_ctor(0, 1, 0); +if (lean_is_scalar(x_80)) { + x_81 = lean_alloc_ctor(0, 1, 0); } else { - x_90 = x_89; + x_81 = x_80; } -lean_ctor_set(x_90, 0, x_88); -return x_90; +lean_ctor_set(x_81, 0, x_79); +return x_81; } else { -lean_object* x_91; size_t x_92; size_t x_93; lean_object* x_94; -x_91 = lean_ctor_get(x_87, 0); -lean_inc(x_91); -lean_dec(x_87); -x_92 = 1; -x_93 = lean_usize_add(x_2, x_92); -x_94 = lean_array_uset(x_9, x_2, x_91); -x_2 = x_93; -x_3 = x_94; +lean_object* x_82; size_t x_83; size_t x_84; lean_object* x_85; +x_82 = lean_ctor_get(x_78, 0); +lean_inc(x_82); +lean_dec(x_78); +x_83 = 1; +x_84 = lean_usize_add(x_2, x_83); +x_85 = lean_array_uset(x_9, x_2, x_82); +x_2 = x_84; +x_3 = x_85; goto _start; } } else { -lean_object* x_96; lean_object* x_97; lean_object* x_98; -lean_dec(x_7); -x_96 = lean_ctor_get(x_79, 0); -lean_inc(x_96); -lean_dec(x_79); -x_97 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_97, 0, x_96); -x_98 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__1(x_97, x_4); -if (lean_obj_tag(x_98) == 0) -{ -lean_object* x_99; lean_object* x_100; lean_object* x_101; +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_dec(x_75); +x_87 = lean_unsigned_to_nat(80u); +x_88 = l_Lean_Json_pretty(x_7, x_87); +x_89 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__2; +x_90 = lean_string_append(x_89, x_88); +lean_dec(x_88); +x_91 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__3; +x_92 = lean_string_append(x_90, x_91); +x_93 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_93, 0, x_92); +x_94 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__1(x_93, x_4); +if (lean_obj_tag(x_94) == 0) +{ +lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_dec(x_9); -x_99 = lean_ctor_get(x_98, 0); -lean_inc(x_99); -if (lean_is_exclusive(x_98)) { - lean_ctor_release(x_98, 0); - x_100 = x_98; +x_95 = lean_ctor_get(x_94, 0); +lean_inc(x_95); +if (lean_is_exclusive(x_94)) { + lean_ctor_release(x_94, 0); + x_96 = x_94; } else { - lean_dec_ref(x_98); - x_100 = lean_box(0); + lean_dec_ref(x_94); + x_96 = lean_box(0); } -if (lean_is_scalar(x_100)) { - x_101 = lean_alloc_ctor(0, 1, 0); +if (lean_is_scalar(x_96)) { + x_97 = lean_alloc_ctor(0, 1, 0); } else { - x_101 = x_100; + x_97 = x_96; } -lean_ctor_set(x_101, 0, x_99); -return x_101; +lean_ctor_set(x_97, 0, x_95); +return x_97; } else { -lean_object* x_102; size_t x_103; size_t x_104; lean_object* x_105; -x_102 = lean_ctor_get(x_98, 0); -lean_inc(x_102); -lean_dec(x_98); -x_103 = 1; -x_104 = lean_usize_add(x_2, x_103); -x_105 = lean_array_uset(x_9, x_2, x_102); -x_2 = x_104; -x_3 = x_105; +lean_object* x_98; size_t x_99; size_t x_100; lean_object* x_101; +x_98 = lean_ctor_get(x_94, 0); +lean_inc(x_98); +lean_dec(x_94); +x_99 = 1; +x_100 = lean_usize_add(x_2, x_99); +x_101 = lean_array_uset(x_9, x_2, x_98); +x_2 = x_100; +x_3 = x_101; goto _start; } } } else { -lean_object* x_107; lean_object* x_108; -lean_dec(x_74); +lean_object* x_103; lean_object* x_104; +lean_dec(x_72); lean_dec(x_7); -x_107 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__5; -x_108 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__1(x_107, x_4); -if (lean_obj_tag(x_108) == 0) +x_103 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__4; +x_104 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__1(x_103, x_4); +if (lean_obj_tag(x_104) == 0) { -lean_object* x_109; lean_object* x_110; lean_object* x_111; +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_dec(x_9); -x_109 = lean_ctor_get(x_108, 0); -lean_inc(x_109); -if (lean_is_exclusive(x_108)) { - lean_ctor_release(x_108, 0); - x_110 = x_108; +x_105 = lean_ctor_get(x_104, 0); +lean_inc(x_105); +if (lean_is_exclusive(x_104)) { + lean_ctor_release(x_104, 0); + x_106 = x_104; } else { - lean_dec_ref(x_108); - x_110 = lean_box(0); + lean_dec_ref(x_104); + x_106 = lean_box(0); } -if (lean_is_scalar(x_110)) { - x_111 = lean_alloc_ctor(0, 1, 0); +if (lean_is_scalar(x_106)) { + x_107 = lean_alloc_ctor(0, 1, 0); } else { - x_111 = x_110; + x_107 = x_106; } -lean_ctor_set(x_111, 0, x_109); -return x_111; +lean_ctor_set(x_107, 0, x_105); +return x_107; } else { -lean_object* x_112; size_t x_113; size_t x_114; lean_object* x_115; -x_112 = lean_ctor_get(x_108, 0); -lean_inc(x_112); -lean_dec(x_108); -x_113 = 1; -x_114 = lean_usize_add(x_2, x_113); -x_115 = lean_array_uset(x_9, x_2, x_112); -x_2 = x_114; -x_3 = x_115; +lean_object* x_108; size_t x_109; size_t x_110; lean_object* x_111; +x_108 = lean_ctor_get(x_104, 0); +lean_inc(x_108); +lean_dec(x_104); +x_109 = 1; +x_110 = lean_usize_add(x_2, x_109); +x_111 = lean_array_uset(x_9, x_2, x_108); +x_2 = x_110; +x_3 = x_111; goto _start; } } @@ -2795,293 +2789,285 @@ x_34 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypoth x_35 = lean_string_dec_eq(x_33, x_34); if (x_35 == 0) { -lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_36 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__2; -x_37 = lean_string_append(x_36, x_33); -lean_dec(x_33); -x_38 = l_Lean_Syntax_decodeNameLit(x_37); -if (lean_obj_tag(x_38) == 0) +lean_object* x_36; uint8_t x_37; +x_36 = l_String_toName(x_33); +x_37 = l_Lean_Name_isAnonymous(x_36); +if (x_37 == 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; -x_39 = lean_unsigned_to_nat(80u); -x_40 = l_Lean_Json_pretty(x_7, x_39); -x_41 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__3; -x_42 = lean_string_append(x_41, x_40); -lean_dec(x_40); -x_43 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__4; -x_44 = lean_string_append(x_42, x_43); -lean_ctor_set_tag(x_10, 0); -lean_ctor_set(x_10, 0, x_44); -x_45 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__3(x_10, x_4); -if (lean_obj_tag(x_45) == 0) +lean_object* x_38; +lean_dec(x_7); +lean_ctor_set(x_10, 0, x_36); +x_38 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__3(x_10, x_4); +if (lean_obj_tag(x_38) == 0) { -uint8_t x_46; +uint8_t x_39; lean_dec(x_9); -x_46 = !lean_is_exclusive(x_45); -if (x_46 == 0) +x_39 = !lean_is_exclusive(x_38); +if (x_39 == 0) { -return x_45; +return x_38; } else { -lean_object* x_47; lean_object* x_48; -x_47 = lean_ctor_get(x_45, 0); -lean_inc(x_47); -lean_dec(x_45); -x_48 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_48, 0, x_47); -return x_48; +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_alloc_ctor(0, 1, 0); +lean_ctor_set(x_41, 0, x_40); +return x_41; } } else { -lean_object* x_49; size_t x_50; size_t x_51; lean_object* x_52; -x_49 = lean_ctor_get(x_45, 0); -lean_inc(x_49); -lean_dec(x_45); -x_50 = 1; -x_51 = lean_usize_add(x_2, x_50); -x_52 = lean_array_uset(x_9, x_2, x_49); -x_2 = x_51; -x_3 = x_52; +lean_object* x_42; size_t x_43; size_t x_44; lean_object* x_45; +x_42 = lean_ctor_get(x_38, 0); +lean_inc(x_42); +lean_dec(x_38); +x_43 = 1; +x_44 = lean_usize_add(x_2, x_43); +x_45 = lean_array_uset(x_9, x_2, x_42); +x_2 = x_44; +x_3 = x_45; goto _start; } } else { -lean_object* x_54; lean_object* x_55; -lean_dec(x_7); -x_54 = lean_ctor_get(x_38, 0); -lean_inc(x_54); -lean_dec(x_38); -lean_ctor_set(x_10, 0, x_54); -x_55 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__3(x_10, x_4); -if (lean_obj_tag(x_55) == 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_dec(x_36); +x_47 = lean_unsigned_to_nat(80u); +x_48 = l_Lean_Json_pretty(x_7, x_47); +x_49 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__2; +x_50 = lean_string_append(x_49, x_48); +lean_dec(x_48); +x_51 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__3; +x_52 = lean_string_append(x_50, x_51); +lean_ctor_set_tag(x_10, 0); +lean_ctor_set(x_10, 0, x_52); +x_53 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__3(x_10, x_4); +if (lean_obj_tag(x_53) == 0) { -uint8_t x_56; +uint8_t x_54; lean_dec(x_9); -x_56 = !lean_is_exclusive(x_55); -if (x_56 == 0) +x_54 = !lean_is_exclusive(x_53); +if (x_54 == 0) { -return x_55; +return x_53; } else { -lean_object* x_57; lean_object* x_58; -x_57 = lean_ctor_get(x_55, 0); -lean_inc(x_57); -lean_dec(x_55); -x_58 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_58, 0, x_57); -return x_58; +lean_object* x_55; lean_object* x_56; +x_55 = lean_ctor_get(x_53, 0); +lean_inc(x_55); +lean_dec(x_53); +x_56 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_56, 0, x_55); +return x_56; } } else { -lean_object* x_59; size_t x_60; size_t x_61; lean_object* x_62; -x_59 = lean_ctor_get(x_55, 0); -lean_inc(x_59); -lean_dec(x_55); -x_60 = 1; -x_61 = lean_usize_add(x_2, x_60); -x_62 = lean_array_uset(x_9, x_2, x_59); -x_2 = x_61; -x_3 = x_62; +lean_object* x_57; size_t x_58; size_t x_59; lean_object* x_60; +x_57 = lean_ctor_get(x_53, 0); +lean_inc(x_57); +lean_dec(x_53); +x_58 = 1; +x_59 = lean_usize_add(x_2, x_58); +x_60 = lean_array_uset(x_9, x_2, x_57); +x_2 = x_59; +x_3 = x_60; goto _start; } } } else { -lean_object* x_64; lean_object* x_65; +lean_object* x_62; lean_object* x_63; lean_free_object(x_10); lean_dec(x_33); lean_dec(x_7); -x_64 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__5; -x_65 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__3(x_64, x_4); -if (lean_obj_tag(x_65) == 0) +x_62 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__4; +x_63 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__3(x_62, x_4); +if (lean_obj_tag(x_63) == 0) { -uint8_t x_66; +uint8_t x_64; lean_dec(x_9); -x_66 = !lean_is_exclusive(x_65); -if (x_66 == 0) +x_64 = !lean_is_exclusive(x_63); +if (x_64 == 0) { -return x_65; +return x_63; } else { -lean_object* x_67; lean_object* x_68; -x_67 = lean_ctor_get(x_65, 0); -lean_inc(x_67); -lean_dec(x_65); -x_68 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_68, 0, x_67); -return x_68; +lean_object* x_65; lean_object* x_66; +x_65 = lean_ctor_get(x_63, 0); +lean_inc(x_65); +lean_dec(x_63); +x_66 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_66, 0, x_65); +return x_66; } } else { -lean_object* x_69; size_t x_70; size_t x_71; lean_object* x_72; -x_69 = lean_ctor_get(x_65, 0); -lean_inc(x_69); -lean_dec(x_65); -x_70 = 1; -x_71 = lean_usize_add(x_2, x_70); -x_72 = lean_array_uset(x_9, x_2, x_69); -x_2 = x_71; -x_3 = x_72; +lean_object* x_67; size_t x_68; size_t x_69; lean_object* x_70; +x_67 = lean_ctor_get(x_63, 0); +lean_inc(x_67); +lean_dec(x_63); +x_68 = 1; +x_69 = lean_usize_add(x_2, x_68); +x_70 = lean_array_uset(x_9, x_2, x_67); +x_2 = x_69; +x_3 = x_70; goto _start; } } } else { -lean_object* x_74; lean_object* x_75; uint8_t x_76; -x_74 = lean_ctor_get(x_10, 0); -lean_inc(x_74); +lean_object* x_72; lean_object* x_73; uint8_t x_74; +x_72 = lean_ctor_get(x_10, 0); +lean_inc(x_72); lean_dec(x_10); -x_75 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__1; -x_76 = lean_string_dec_eq(x_74, x_75); +x_73 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__1; +x_74 = lean_string_dec_eq(x_72, x_73); +if (x_74 == 0) +{ +lean_object* x_75; uint8_t x_76; +x_75 = l_String_toName(x_72); +x_76 = l_Lean_Name_isAnonymous(x_75); if (x_76 == 0) { -lean_object* x_77; lean_object* x_78; lean_object* x_79; -x_77 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__2; -x_78 = lean_string_append(x_77, x_74); -lean_dec(x_74); -x_79 = l_Lean_Syntax_decodeNameLit(x_78); -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_85; lean_object* x_86; lean_object* x_87; -x_80 = lean_unsigned_to_nat(80u); -x_81 = l_Lean_Json_pretty(x_7, x_80); -x_82 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__3; -x_83 = lean_string_append(x_82, x_81); -lean_dec(x_81); -x_84 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__4; -x_85 = lean_string_append(x_83, x_84); -x_86 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_86, 0, x_85); -x_87 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__3(x_86, x_4); -if (lean_obj_tag(x_87) == 0) +lean_object* x_77; lean_object* x_78; +lean_dec(x_7); +x_77 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_77, 0, x_75); +x_78 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__3(x_77, x_4); +if (lean_obj_tag(x_78) == 0) { -lean_object* x_88; lean_object* x_89; lean_object* x_90; +lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_dec(x_9); -x_88 = lean_ctor_get(x_87, 0); -lean_inc(x_88); -if (lean_is_exclusive(x_87)) { - lean_ctor_release(x_87, 0); - x_89 = x_87; +x_79 = lean_ctor_get(x_78, 0); +lean_inc(x_79); +if (lean_is_exclusive(x_78)) { + lean_ctor_release(x_78, 0); + x_80 = x_78; } else { - lean_dec_ref(x_87); - x_89 = lean_box(0); + lean_dec_ref(x_78); + x_80 = lean_box(0); } -if (lean_is_scalar(x_89)) { - x_90 = lean_alloc_ctor(0, 1, 0); +if (lean_is_scalar(x_80)) { + x_81 = lean_alloc_ctor(0, 1, 0); } else { - x_90 = x_89; + x_81 = x_80; } -lean_ctor_set(x_90, 0, x_88); -return x_90; +lean_ctor_set(x_81, 0, x_79); +return x_81; } else { -lean_object* x_91; size_t x_92; size_t x_93; lean_object* x_94; -x_91 = lean_ctor_get(x_87, 0); -lean_inc(x_91); -lean_dec(x_87); -x_92 = 1; -x_93 = lean_usize_add(x_2, x_92); -x_94 = lean_array_uset(x_9, x_2, x_91); -x_2 = x_93; -x_3 = x_94; +lean_object* x_82; size_t x_83; size_t x_84; lean_object* x_85; +x_82 = lean_ctor_get(x_78, 0); +lean_inc(x_82); +lean_dec(x_78); +x_83 = 1; +x_84 = lean_usize_add(x_2, x_83); +x_85 = lean_array_uset(x_9, x_2, x_82); +x_2 = x_84; +x_3 = x_85; goto _start; } } else { -lean_object* x_96; lean_object* x_97; lean_object* x_98; -lean_dec(x_7); -x_96 = lean_ctor_get(x_79, 0); -lean_inc(x_96); -lean_dec(x_79); -x_97 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_97, 0, x_96); -x_98 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__3(x_97, x_4); -if (lean_obj_tag(x_98) == 0) -{ -lean_object* x_99; lean_object* x_100; lean_object* x_101; +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_dec(x_75); +x_87 = lean_unsigned_to_nat(80u); +x_88 = l_Lean_Json_pretty(x_7, x_87); +x_89 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__2; +x_90 = lean_string_append(x_89, x_88); +lean_dec(x_88); +x_91 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__3; +x_92 = lean_string_append(x_90, x_91); +x_93 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_93, 0, x_92); +x_94 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__3(x_93, x_4); +if (lean_obj_tag(x_94) == 0) +{ +lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_dec(x_9); -x_99 = lean_ctor_get(x_98, 0); -lean_inc(x_99); -if (lean_is_exclusive(x_98)) { - lean_ctor_release(x_98, 0); - x_100 = x_98; +x_95 = lean_ctor_get(x_94, 0); +lean_inc(x_95); +if (lean_is_exclusive(x_94)) { + lean_ctor_release(x_94, 0); + x_96 = x_94; } else { - lean_dec_ref(x_98); - x_100 = lean_box(0); + lean_dec_ref(x_94); + x_96 = lean_box(0); } -if (lean_is_scalar(x_100)) { - x_101 = lean_alloc_ctor(0, 1, 0); +if (lean_is_scalar(x_96)) { + x_97 = lean_alloc_ctor(0, 1, 0); } else { - x_101 = x_100; + x_97 = x_96; } -lean_ctor_set(x_101, 0, x_99); -return x_101; +lean_ctor_set(x_97, 0, x_95); +return x_97; } else { -lean_object* x_102; size_t x_103; size_t x_104; lean_object* x_105; -x_102 = lean_ctor_get(x_98, 0); -lean_inc(x_102); -lean_dec(x_98); -x_103 = 1; -x_104 = lean_usize_add(x_2, x_103); -x_105 = lean_array_uset(x_9, x_2, x_102); -x_2 = x_104; -x_3 = x_105; +lean_object* x_98; size_t x_99; size_t x_100; lean_object* x_101; +x_98 = lean_ctor_get(x_94, 0); +lean_inc(x_98); +lean_dec(x_94); +x_99 = 1; +x_100 = lean_usize_add(x_2, x_99); +x_101 = lean_array_uset(x_9, x_2, x_98); +x_2 = x_100; +x_3 = x_101; goto _start; } } } else { -lean_object* x_107; lean_object* x_108; -lean_dec(x_74); +lean_object* x_103; lean_object* x_104; +lean_dec(x_72); lean_dec(x_7); -x_107 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__5; -x_108 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__3(x_107, x_4); -if (lean_obj_tag(x_108) == 0) +x_103 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__4; +x_104 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__3(x_103, x_4); +if (lean_obj_tag(x_104) == 0) { -lean_object* x_109; lean_object* x_110; lean_object* x_111; +lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_dec(x_9); -x_109 = lean_ctor_get(x_108, 0); -lean_inc(x_109); -if (lean_is_exclusive(x_108)) { - lean_ctor_release(x_108, 0); - x_110 = x_108; +x_105 = lean_ctor_get(x_104, 0); +lean_inc(x_105); +if (lean_is_exclusive(x_104)) { + lean_ctor_release(x_104, 0); + x_106 = x_104; } else { - lean_dec_ref(x_108); - x_110 = lean_box(0); + lean_dec_ref(x_104); + x_106 = lean_box(0); } -if (lean_is_scalar(x_110)) { - x_111 = lean_alloc_ctor(0, 1, 0); +if (lean_is_scalar(x_106)) { + x_107 = lean_alloc_ctor(0, 1, 0); } else { - x_111 = x_110; + x_107 = x_106; } -lean_ctor_set(x_111, 0, x_109); -return x_111; +lean_ctor_set(x_107, 0, x_105); +return x_107; } else { -lean_object* x_112; size_t x_113; size_t x_114; lean_object* x_115; -x_112 = lean_ctor_get(x_108, 0); -lean_inc(x_112); -lean_dec(x_108); -x_113 = 1; -x_114 = lean_usize_add(x_2, x_113); -x_115 = lean_array_uset(x_9, x_2, x_112); -x_2 = x_114; -x_3 = x_115; +lean_object* x_108; size_t x_109; size_t x_110; lean_object* x_111; +x_108 = lean_ctor_get(x_104, 0); +lean_inc(x_108); +lean_dec(x_104); +x_109 = 1; +x_110 = lean_usize_add(x_2, x_109); +x_111 = lean_array_uset(x_9, x_2, x_108); +x_2 = x_110; +x_3 = x_111; goto _start; } } @@ -5934,7 +5920,7 @@ x_402 = l_Lean_Json_pretty(x_19, x_401); x_403 = l_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____closed__1; x_404 = lean_string_append(x_403, x_402); lean_dec(x_402); -x_405 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__4; +x_405 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__3; x_406 = lean_string_append(x_404, x_405); lean_ctor_set_tag(x_15, 0); lean_ctor_set(x_15, 0, x_406); @@ -5949,7 +5935,7 @@ x_408 = l_Lean_Json_pretty(x_19, x_407); x_409 = l_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____closed__1; x_410 = lean_string_append(x_409, x_408); lean_dec(x_408); -x_411 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__4; +x_411 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__3; x_412 = lean_string_append(x_410, x_411); x_413 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_413, 0, x_412); @@ -5968,7 +5954,7 @@ x_415 = l_Lean_Json_pretty(x_6, x_414); x_416 = l_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____closed__1; x_417 = lean_string_append(x_416, x_415); lean_dec(x_415); -x_418 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__4; +x_418 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__3; x_419 = lean_string_append(x_417, x_418); lean_ctor_set_tag(x_3, 0); lean_ctor_set(x_3, 0, x_419); @@ -7047,7 +7033,7 @@ x_601 = l_Lean_Json_pretty(x_434, x_600); x_602 = l_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____closed__1; x_603 = lean_string_append(x_602, x_601); lean_dec(x_601); -x_604 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__4; +x_604 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__3; x_605 = lean_string_append(x_603, x_604); if (lean_is_scalar(x_599)) { x_606 = lean_alloc_ctor(0, 1, 0); @@ -7070,7 +7056,7 @@ x_608 = l_Lean_Json_pretty(x_421, x_607); x_609 = l_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____closed__1; x_610 = lean_string_append(x_609, x_608); lean_dec(x_608); -x_611 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__4; +x_611 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__3; x_612 = lean_string_append(x_610, x_611); x_613 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_613, 0, x_612); @@ -7088,6 +7074,15 @@ lean_dec(x_2); return x_3; } } +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___lambda__1(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { @@ -8499,114 +8494,114 @@ return x_37; } else { -lean_object* x_38; lean_object* x_39; lean_object* x_230; +lean_object* x_38; lean_object* x_39; lean_object* x_226; x_38 = lean_ctor_get(x_34, 0); lean_inc(x_38); lean_dec(x_34); -x_230 = lean_ctor_get(x_5, 3); -lean_inc(x_230); -if (lean_obj_tag(x_230) == 0) +x_226 = lean_ctor_get(x_5, 3); +lean_inc(x_226); +if (lean_obj_tag(x_226) == 0) { -lean_object* x_231; -x_231 = lean_box(0); -x_39 = x_231; -goto block_229; +lean_object* x_227; +x_227 = lean_box(0); +x_39 = x_227; +goto block_225; } else { -uint8_t x_232; -x_232 = !lean_is_exclusive(x_230); -if (x_232 == 0) +uint8_t x_228; +x_228 = !lean_is_exclusive(x_226); +if (x_228 == 0) { -lean_object* x_233; lean_object* x_234; lean_object* x_235; -x_233 = lean_ctor_get(x_230, 0); -x_234 = l_Lean_Json_getStr_x3f(x_233); -lean_dec(x_233); -x_235 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveGoal_dec____x40_Lean_Widget_InteractiveGoal___hyg_1387____spec__2(x_234, x_2); -if (lean_obj_tag(x_235) == 0) +lean_object* x_229; lean_object* x_230; lean_object* x_231; +x_229 = lean_ctor_get(x_226, 0); +x_230 = l_Lean_Json_getStr_x3f(x_229); +lean_dec(x_229); +x_231 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveGoal_dec____x40_Lean_Widget_InteractiveGoal___hyg_1387____spec__2(x_230, x_2); +if (lean_obj_tag(x_231) == 0) { -uint8_t x_236; -lean_free_object(x_230); +uint8_t x_232; +lean_free_object(x_226); lean_dec(x_38); lean_dec(x_31); lean_dec(x_19); lean_dec(x_5); lean_dec(x_2); -x_236 = !lean_is_exclusive(x_235); -if (x_236 == 0) +x_232 = !lean_is_exclusive(x_231); +if (x_232 == 0) { -return x_235; +return x_231; } else { -lean_object* x_237; lean_object* x_238; -x_237 = lean_ctor_get(x_235, 0); -lean_inc(x_237); -lean_dec(x_235); -x_238 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_238, 0, x_237); -return x_238; +lean_object* x_233; lean_object* x_234; +x_233 = lean_ctor_get(x_231, 0); +lean_inc(x_233); +lean_dec(x_231); +x_234 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_234, 0, x_233); +return x_234; } } else { -lean_object* x_239; -x_239 = lean_ctor_get(x_235, 0); -lean_inc(x_239); -lean_dec(x_235); -lean_ctor_set(x_230, 0, x_239); -x_39 = x_230; -goto block_229; +lean_object* x_235; +x_235 = lean_ctor_get(x_231, 0); +lean_inc(x_235); +lean_dec(x_231); +lean_ctor_set(x_226, 0, x_235); +x_39 = x_226; +goto block_225; } } else { -lean_object* x_240; lean_object* x_241; lean_object* x_242; -x_240 = lean_ctor_get(x_230, 0); -lean_inc(x_240); -lean_dec(x_230); -x_241 = l_Lean_Json_getStr_x3f(x_240); -lean_dec(x_240); -x_242 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveGoal_dec____x40_Lean_Widget_InteractiveGoal___hyg_1387____spec__2(x_241, x_2); -if (lean_obj_tag(x_242) == 0) +lean_object* x_236; lean_object* x_237; lean_object* x_238; +x_236 = lean_ctor_get(x_226, 0); +lean_inc(x_236); +lean_dec(x_226); +x_237 = l_Lean_Json_getStr_x3f(x_236); +lean_dec(x_236); +x_238 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveGoal_dec____x40_Lean_Widget_InteractiveGoal___hyg_1387____spec__2(x_237, x_2); +if (lean_obj_tag(x_238) == 0) { -lean_object* x_243; lean_object* x_244; lean_object* x_245; +lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_dec(x_38); lean_dec(x_31); lean_dec(x_19); lean_dec(x_5); lean_dec(x_2); -x_243 = lean_ctor_get(x_242, 0); -lean_inc(x_243); -if (lean_is_exclusive(x_242)) { - lean_ctor_release(x_242, 0); - x_244 = x_242; +x_239 = lean_ctor_get(x_238, 0); +lean_inc(x_239); +if (lean_is_exclusive(x_238)) { + lean_ctor_release(x_238, 0); + x_240 = x_238; } else { - lean_dec_ref(x_242); - x_244 = lean_box(0); + lean_dec_ref(x_238); + x_240 = lean_box(0); } -if (lean_is_scalar(x_244)) { - x_245 = lean_alloc_ctor(0, 1, 0); +if (lean_is_scalar(x_240)) { + x_241 = lean_alloc_ctor(0, 1, 0); } else { - x_245 = x_244; + x_241 = x_240; } -lean_ctor_set(x_245, 0, x_243); -return x_245; +lean_ctor_set(x_241, 0, x_239); +return x_241; } else { -lean_object* x_246; lean_object* x_247; -x_246 = lean_ctor_get(x_242, 0); -lean_inc(x_246); -lean_dec(x_242); -x_247 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_247, 0, x_246); -x_39 = x_247; -goto block_229; +lean_object* x_242; lean_object* x_243; +x_242 = lean_ctor_get(x_238, 0); +lean_inc(x_242); +lean_dec(x_238); +x_243 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_243, 0, x_242); +x_39 = x_243; +goto block_225; } } } -block_229: +block_225: { lean_object* x_40; lean_object* x_41; lean_object* x_42; x_40 = lean_ctor_get(x_5, 4); @@ -8682,99 +8677,90 @@ x_200 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypot x_201 = lean_string_dec_eq(x_199, x_200); if (x_201 == 0) { -lean_object* x_202; lean_object* x_203; lean_object* x_204; -x_202 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__2; -x_203 = lean_string_append(x_202, x_199); -lean_dec(x_199); -x_204 = l_Lean_Syntax_decodeNameLit(x_203); -if (lean_obj_tag(x_204) == 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; -x_205 = lean_unsigned_to_nat(80u); -x_206 = l_Lean_Json_pretty(x_47, x_205); -x_207 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__3; -x_208 = lean_string_append(x_207, x_206); -lean_dec(x_206); -x_209 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__4; -x_210 = lean_string_append(x_208, x_209); -lean_ctor_set_tag(x_194, 0); -lean_ctor_set(x_194, 0, x_210); +lean_object* x_202; uint8_t x_203; +x_202 = l_String_toName(x_199); +x_203 = l_Lean_Name_isAnonymous(x_202); +if (x_203 == 0) +{ +lean_dec(x_47); +lean_ctor_set(x_194, 0, x_202); x_48 = x_194; goto block_193; } else { -lean_object* x_211; -lean_dec(x_47); -x_211 = lean_ctor_get(x_204, 0); -lean_inc(x_211); -lean_dec(x_204); -lean_ctor_set(x_194, 0, x_211); +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_dec(x_202); +x_204 = lean_unsigned_to_nat(80u); +x_205 = l_Lean_Json_pretty(x_47, x_204); +x_206 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__2; +x_207 = lean_string_append(x_206, x_205); +lean_dec(x_205); +x_208 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__3; +x_209 = lean_string_append(x_207, x_208); +lean_ctor_set_tag(x_194, 0); +lean_ctor_set(x_194, 0, x_209); x_48 = x_194; goto block_193; } } else { -lean_object* x_212; +lean_object* x_210; lean_free_object(x_194); lean_dec(x_199); lean_dec(x_47); -x_212 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__5; -x_48 = x_212; +x_210 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__4; +x_48 = x_210; goto block_193; } } else { -lean_object* x_213; lean_object* x_214; uint8_t x_215; -x_213 = lean_ctor_get(x_194, 0); -lean_inc(x_213); +lean_object* x_211; lean_object* x_212; uint8_t x_213; +x_211 = lean_ctor_get(x_194, 0); +lean_inc(x_211); lean_dec(x_194); -x_214 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__1; -x_215 = lean_string_dec_eq(x_213, x_214); -if (x_215 == 0) +x_212 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__1; +x_213 = lean_string_dec_eq(x_211, x_212); +if (x_213 == 0) { -lean_object* x_216; lean_object* x_217; lean_object* x_218; -x_216 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__2; -x_217 = lean_string_append(x_216, x_213); -lean_dec(x_213); -x_218 = l_Lean_Syntax_decodeNameLit(x_217); -if (lean_obj_tag(x_218) == 0) +lean_object* x_214; uint8_t x_215; +x_214 = l_String_toName(x_211); +x_215 = l_Lean_Name_isAnonymous(x_214); +if (x_215 == 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_object* x_225; -x_219 = lean_unsigned_to_nat(80u); -x_220 = l_Lean_Json_pretty(x_47, x_219); -x_221 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__3; -x_222 = lean_string_append(x_221, x_220); -lean_dec(x_220); -x_223 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__4; -x_224 = lean_string_append(x_222, x_223); -x_225 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_225, 0, x_224); -x_48 = x_225; +lean_object* x_216; +lean_dec(x_47); +x_216 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_216, 0, x_214); +x_48 = x_216; goto block_193; } else { -lean_object* x_226; lean_object* x_227; -lean_dec(x_47); -x_226 = lean_ctor_get(x_218, 0); -lean_inc(x_226); +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_214); +x_217 = lean_unsigned_to_nat(80u); +x_218 = l_Lean_Json_pretty(x_47, x_217); +x_219 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__2; +x_220 = lean_string_append(x_219, x_218); lean_dec(x_218); -x_227 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_227, 0, x_226); -x_48 = x_227; +x_221 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__3; +x_222 = lean_string_append(x_220, x_221); +x_223 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_223, 0, x_222); +x_48 = x_223; goto block_193; } } else { -lean_object* x_228; -lean_dec(x_213); +lean_object* x_224; +lean_dec(x_211); lean_dec(x_47); -x_228 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__5; -x_48 = x_228; +x_224 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__4; +x_48 = x_224; goto block_193; } } @@ -9637,706 +9623,702 @@ return x_192; } else { -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_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_dec(x_5); lean_dec(x_2); -x_248 = lean_unsigned_to_nat(80u); -x_249 = l_Lean_Json_pretty(x_6, x_248); -x_250 = l_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____closed__1; -x_251 = lean_string_append(x_250, x_249); -lean_dec(x_249); -x_252 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__4; -x_253 = lean_string_append(x_251, x_252); +x_244 = lean_unsigned_to_nat(80u); +x_245 = l_Lean_Json_pretty(x_6, x_244); +x_246 = l_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____closed__1; +x_247 = lean_string_append(x_246, x_245); +lean_dec(x_245); +x_248 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__3; +x_249 = lean_string_append(x_247, x_248); lean_ctor_set_tag(x_3, 0); -lean_ctor_set(x_3, 0, x_253); +lean_ctor_set(x_3, 0, x_249); return x_3; } } else { -lean_object* x_254; lean_object* x_255; -x_254 = lean_ctor_get(x_3, 0); -lean_inc(x_254); +lean_object* x_250; lean_object* x_251; +x_250 = lean_ctor_get(x_3, 0); +lean_inc(x_250); lean_dec(x_3); -x_255 = lean_ctor_get(x_254, 0); -lean_inc(x_255); -if (lean_obj_tag(x_255) == 4) -{ -lean_object* x_256; lean_object* x_257; size_t x_258; size_t x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; size_t x_263; lean_object* x_264; -x_256 = lean_ctor_get(x_255, 0); -lean_inc(x_256); -lean_dec(x_255); -x_257 = lean_array_get_size(x_256); -x_258 = lean_usize_of_nat(x_257); -lean_dec(x_257); -x_259 = 0; -x_260 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1837____spec__2(x_258, x_259, x_256); -x_261 = lean_ctor_get(x_260, 0); -lean_inc(x_261); -lean_dec(x_260); -x_262 = lean_array_get_size(x_261); -x_263 = lean_usize_of_nat(x_262); -lean_dec(x_262); +x_251 = lean_ctor_get(x_250, 0); +lean_inc(x_251); +if (lean_obj_tag(x_251) == 4) +{ +lean_object* x_252; lean_object* x_253; size_t x_254; size_t x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; size_t x_259; lean_object* x_260; +x_252 = lean_ctor_get(x_251, 0); +lean_inc(x_252); +lean_dec(x_251); +x_253 = lean_array_get_size(x_252); +x_254 = lean_usize_of_nat(x_253); +lean_dec(x_253); +x_255 = 0; +x_256 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonCommand____x40_Lean_Data_Lsp_Basic___hyg_1837____spec__2(x_254, x_255, x_252); +x_257 = lean_ctor_get(x_256, 0); +lean_inc(x_257); +lean_dec(x_256); +x_258 = lean_array_get_size(x_257); +x_259 = lean_usize_of_nat(x_258); +lean_dec(x_258); lean_inc(x_2); -x_264 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveGoal_dec____x40_Lean_Widget_InteractiveGoal___hyg_1387____spec__1(x_263, x_259, x_261, x_2); -if (lean_obj_tag(x_264) == 0) +x_260 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveGoal_dec____x40_Lean_Widget_InteractiveGoal___hyg_1387____spec__1(x_259, x_255, x_257, x_2); +if (lean_obj_tag(x_260) == 0) { -lean_object* x_265; lean_object* x_266; lean_object* x_267; -lean_dec(x_254); +lean_object* x_261; lean_object* x_262; lean_object* x_263; +lean_dec(x_250); lean_dec(x_2); -x_265 = lean_ctor_get(x_264, 0); -lean_inc(x_265); -if (lean_is_exclusive(x_264)) { - lean_ctor_release(x_264, 0); - x_266 = x_264; +x_261 = lean_ctor_get(x_260, 0); +lean_inc(x_261); +if (lean_is_exclusive(x_260)) { + lean_ctor_release(x_260, 0); + x_262 = x_260; } else { - lean_dec_ref(x_264); - x_266 = lean_box(0); + lean_dec_ref(x_260); + x_262 = lean_box(0); } -if (lean_is_scalar(x_266)) { - x_267 = lean_alloc_ctor(0, 1, 0); +if (lean_is_scalar(x_262)) { + x_263 = lean_alloc_ctor(0, 1, 0); } else { - x_267 = x_266; + x_263 = x_262; } -lean_ctor_set(x_267, 0, x_265); -return x_267; +lean_ctor_set(x_263, 0, x_261); +return x_263; } else { -lean_object* x_268; lean_object* x_269; lean_object* x_270; -x_268 = lean_ctor_get(x_264, 0); -lean_inc(x_268); -lean_dec(x_264); -x_269 = lean_ctor_get(x_254, 1); -lean_inc(x_269); -x_270 = l___private_Lean_Widget_TaggedText_0__Lean_Widget_fromJsonTaggedText____x40_Lean_Widget_TaggedText___hyg_432____at_Lean_Widget_TaggedText_instRpcEncodableTaggedText___spec__5(x_269); -if (lean_obj_tag(x_270) == 0) +lean_object* x_264; lean_object* x_265; lean_object* x_266; +x_264 = lean_ctor_get(x_260, 0); +lean_inc(x_264); +lean_dec(x_260); +x_265 = lean_ctor_get(x_250, 1); +lean_inc(x_265); +x_266 = l___private_Lean_Widget_TaggedText_0__Lean_Widget_fromJsonTaggedText____x40_Lean_Widget_TaggedText___hyg_432____at_Lean_Widget_TaggedText_instRpcEncodableTaggedText___spec__5(x_265); +if (lean_obj_tag(x_266) == 0) { -lean_object* x_271; lean_object* x_272; lean_object* x_273; -lean_dec(x_268); -lean_dec(x_254); +lean_object* x_267; lean_object* x_268; lean_object* x_269; +lean_dec(x_264); +lean_dec(x_250); lean_dec(x_2); -x_271 = lean_ctor_get(x_270, 0); -lean_inc(x_271); -if (lean_is_exclusive(x_270)) { - lean_ctor_release(x_270, 0); - x_272 = x_270; +x_267 = lean_ctor_get(x_266, 0); +lean_inc(x_267); +if (lean_is_exclusive(x_266)) { + lean_ctor_release(x_266, 0); + x_268 = x_266; } else { - lean_dec_ref(x_270); - x_272 = lean_box(0); + lean_dec_ref(x_266); + x_268 = lean_box(0); } -if (lean_is_scalar(x_272)) { - x_273 = lean_alloc_ctor(0, 1, 0); +if (lean_is_scalar(x_268)) { + x_269 = lean_alloc_ctor(0, 1, 0); } else { - x_273 = x_272; + x_269 = x_268; } -lean_ctor_set(x_273, 0, x_271); -return x_273; +lean_ctor_set(x_269, 0, x_267); +return x_269; } else { -lean_object* x_274; lean_object* x_275; lean_object* x_276; -x_274 = lean_ctor_get(x_270, 0); -lean_inc(x_274); -lean_dec(x_270); -x_275 = l_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____closed__2; +lean_object* x_270; lean_object* x_271; lean_object* x_272; +x_270 = lean_ctor_get(x_266, 0); +lean_inc(x_270); +lean_dec(x_266); +x_271 = l_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____closed__2; lean_inc(x_2); -x_276 = l_Lean_Widget_TaggedText_mapM___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__5(x_275, x_274, x_2); -if (lean_obj_tag(x_276) == 0) +x_272 = l_Lean_Widget_TaggedText_mapM___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__5(x_271, x_270, x_2); +if (lean_obj_tag(x_272) == 0) { -lean_object* x_277; lean_object* x_278; lean_object* x_279; -lean_dec(x_268); -lean_dec(x_254); +lean_object* x_273; lean_object* x_274; lean_object* x_275; +lean_dec(x_264); +lean_dec(x_250); lean_dec(x_2); -x_277 = lean_ctor_get(x_276, 0); -lean_inc(x_277); -if (lean_is_exclusive(x_276)) { - lean_ctor_release(x_276, 0); - x_278 = x_276; +x_273 = lean_ctor_get(x_272, 0); +lean_inc(x_273); +if (lean_is_exclusive(x_272)) { + lean_ctor_release(x_272, 0); + x_274 = x_272; } else { - lean_dec_ref(x_276); - x_278 = lean_box(0); + lean_dec_ref(x_272); + x_274 = lean_box(0); } -if (lean_is_scalar(x_278)) { - x_279 = lean_alloc_ctor(0, 1, 0); +if (lean_is_scalar(x_274)) { + x_275 = lean_alloc_ctor(0, 1, 0); } else { - x_279 = x_278; + x_275 = x_274; } -lean_ctor_set(x_279, 0, x_277); -return x_279; +lean_ctor_set(x_275, 0, x_273); +return x_275; } else { -lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; -x_280 = lean_ctor_get(x_276, 0); -lean_inc(x_280); -lean_dec(x_276); -x_281 = lean_ctor_get(x_254, 2); -lean_inc(x_281); -x_282 = l_Lean_Widget_instTypeNameContextInfo; +lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; +x_276 = lean_ctor_get(x_272, 0); +lean_inc(x_276); +lean_dec(x_272); +x_277 = lean_ctor_get(x_250, 2); +lean_inc(x_277); +x_278 = l_Lean_Widget_instTypeNameContextInfo; lean_inc(x_2); -x_283 = l_Lean_Server_instRpcEncodableWithRpcRef_rpcDecode___rarg(x_282, x_281, x_2); -if (lean_obj_tag(x_283) == 0) +x_279 = l_Lean_Server_instRpcEncodableWithRpcRef_rpcDecode___rarg(x_278, x_277, x_2); +if (lean_obj_tag(x_279) == 0) { -lean_object* x_284; lean_object* x_285; lean_object* x_286; -lean_dec(x_280); -lean_dec(x_268); -lean_dec(x_254); +lean_object* x_280; lean_object* x_281; lean_object* x_282; +lean_dec(x_276); +lean_dec(x_264); +lean_dec(x_250); lean_dec(x_2); -x_284 = lean_ctor_get(x_283, 0); -lean_inc(x_284); -if (lean_is_exclusive(x_283)) { - lean_ctor_release(x_283, 0); - x_285 = x_283; +x_280 = lean_ctor_get(x_279, 0); +lean_inc(x_280); +if (lean_is_exclusive(x_279)) { + lean_ctor_release(x_279, 0); + x_281 = x_279; } else { - lean_dec_ref(x_283); - x_285 = lean_box(0); + lean_dec_ref(x_279); + x_281 = lean_box(0); } -if (lean_is_scalar(x_285)) { - x_286 = lean_alloc_ctor(0, 1, 0); +if (lean_is_scalar(x_281)) { + x_282 = lean_alloc_ctor(0, 1, 0); } else { - x_286 = x_285; + x_282 = x_281; } -lean_ctor_set(x_286, 0, x_284); -return x_286; +lean_ctor_set(x_282, 0, x_280); +return x_282; } else { -lean_object* x_287; lean_object* x_288; lean_object* x_375; -x_287 = lean_ctor_get(x_283, 0); -lean_inc(x_287); -lean_dec(x_283); -x_375 = lean_ctor_get(x_254, 3); -lean_inc(x_375); -if (lean_obj_tag(x_375) == 0) +lean_object* x_283; lean_object* x_284; lean_object* x_369; +x_283 = lean_ctor_get(x_279, 0); +lean_inc(x_283); +lean_dec(x_279); +x_369 = lean_ctor_get(x_250, 3); +lean_inc(x_369); +if (lean_obj_tag(x_369) == 0) { -lean_object* x_376; -x_376 = lean_box(0); -x_288 = x_376; -goto block_374; +lean_object* x_370; +x_370 = lean_box(0); +x_284 = x_370; +goto block_368; } else { -lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; -x_377 = lean_ctor_get(x_375, 0); -lean_inc(x_377); -if (lean_is_exclusive(x_375)) { - lean_ctor_release(x_375, 0); - x_378 = x_375; +lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; +x_371 = lean_ctor_get(x_369, 0); +lean_inc(x_371); +if (lean_is_exclusive(x_369)) { + lean_ctor_release(x_369, 0); + x_372 = x_369; } else { - lean_dec_ref(x_375); - x_378 = lean_box(0); + lean_dec_ref(x_369); + x_372 = lean_box(0); } -x_379 = l_Lean_Json_getStr_x3f(x_377); -lean_dec(x_377); -x_380 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveGoal_dec____x40_Lean_Widget_InteractiveGoal___hyg_1387____spec__2(x_379, x_2); -if (lean_obj_tag(x_380) == 0) +x_373 = l_Lean_Json_getStr_x3f(x_371); +lean_dec(x_371); +x_374 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveGoal_dec____x40_Lean_Widget_InteractiveGoal___hyg_1387____spec__2(x_373, x_2); +if (lean_obj_tag(x_374) == 0) { -lean_object* x_381; lean_object* x_382; lean_object* x_383; -lean_dec(x_378); -lean_dec(x_287); -lean_dec(x_280); -lean_dec(x_268); -lean_dec(x_254); +lean_object* x_375; lean_object* x_376; lean_object* x_377; +lean_dec(x_372); +lean_dec(x_283); +lean_dec(x_276); +lean_dec(x_264); +lean_dec(x_250); lean_dec(x_2); -x_381 = lean_ctor_get(x_380, 0); -lean_inc(x_381); -if (lean_is_exclusive(x_380)) { - lean_ctor_release(x_380, 0); - x_382 = x_380; +x_375 = lean_ctor_get(x_374, 0); +lean_inc(x_375); +if (lean_is_exclusive(x_374)) { + lean_ctor_release(x_374, 0); + x_376 = x_374; } else { - lean_dec_ref(x_380); - x_382 = lean_box(0); + lean_dec_ref(x_374); + x_376 = lean_box(0); } -if (lean_is_scalar(x_382)) { - x_383 = lean_alloc_ctor(0, 1, 0); +if (lean_is_scalar(x_376)) { + x_377 = lean_alloc_ctor(0, 1, 0); } else { - x_383 = x_382; + x_377 = x_376; } -lean_ctor_set(x_383, 0, x_381); -return x_383; +lean_ctor_set(x_377, 0, x_375); +return x_377; } else { -lean_object* x_384; lean_object* x_385; -x_384 = lean_ctor_get(x_380, 0); -lean_inc(x_384); -lean_dec(x_380); -if (lean_is_scalar(x_378)) { - x_385 = lean_alloc_ctor(1, 1, 0); +lean_object* x_378; lean_object* x_379; +x_378 = lean_ctor_get(x_374, 0); +lean_inc(x_378); +lean_dec(x_374); +if (lean_is_scalar(x_372)) { + x_379 = lean_alloc_ctor(1, 1, 0); } else { - x_385 = x_378; + x_379 = x_372; } -lean_ctor_set(x_385, 0, x_384); -x_288 = x_385; -goto block_374; +lean_ctor_set(x_379, 0, x_378); +x_284 = x_379; +goto block_368; } } -block_374: +block_368: { -lean_object* x_289; lean_object* x_290; lean_object* x_291; -x_289 = lean_ctor_get(x_254, 4); -lean_inc(x_289); -x_290 = l_Lean_Json_getStr_x3f(x_289); -lean_dec(x_289); -x_291 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveGoal_dec____x40_Lean_Widget_InteractiveGoal___hyg_1387____spec__2(x_290, x_2); -if (lean_obj_tag(x_291) == 0) +lean_object* x_285; lean_object* x_286; lean_object* x_287; +x_285 = lean_ctor_get(x_250, 4); +lean_inc(x_285); +x_286 = l_Lean_Json_getStr_x3f(x_285); +lean_dec(x_285); +x_287 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveGoal_dec____x40_Lean_Widget_InteractiveGoal___hyg_1387____spec__2(x_286, x_2); +if (lean_obj_tag(x_287) == 0) { -lean_object* x_292; lean_object* x_293; lean_object* x_294; -lean_dec(x_288); -lean_dec(x_287); -lean_dec(x_280); -lean_dec(x_268); -lean_dec(x_254); +lean_object* x_288; lean_object* x_289; lean_object* x_290; +lean_dec(x_284); +lean_dec(x_283); +lean_dec(x_276); +lean_dec(x_264); +lean_dec(x_250); lean_dec(x_2); -x_292 = lean_ctor_get(x_291, 0); -lean_inc(x_292); -if (lean_is_exclusive(x_291)) { - lean_ctor_release(x_291, 0); - x_293 = x_291; +x_288 = lean_ctor_get(x_287, 0); +lean_inc(x_288); +if (lean_is_exclusive(x_287)) { + lean_ctor_release(x_287, 0); + x_289 = x_287; } else { - lean_dec_ref(x_291); - x_293 = lean_box(0); + lean_dec_ref(x_287); + x_289 = lean_box(0); } -if (lean_is_scalar(x_293)) { - x_294 = lean_alloc_ctor(0, 1, 0); +if (lean_is_scalar(x_289)) { + x_290 = lean_alloc_ctor(0, 1, 0); } else { - x_294 = x_293; + x_290 = x_289; } -lean_ctor_set(x_294, 0, x_292); -return x_294; +lean_ctor_set(x_290, 0, x_288); +return x_290; } else { -lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_353; -x_295 = lean_ctor_get(x_291, 0); -lean_inc(x_295); -lean_dec(x_291); -x_296 = lean_ctor_get(x_254, 5); -lean_inc(x_296); -x_353 = l_Lean_Json_getStr_x3f(x_296); -if (lean_obj_tag(x_353) == 0) +lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_349; +x_291 = lean_ctor_get(x_287, 0); +lean_inc(x_291); +lean_dec(x_287); +x_292 = lean_ctor_get(x_250, 5); +lean_inc(x_292); +x_349 = l_Lean_Json_getStr_x3f(x_292); +if (lean_obj_tag(x_349) == 0) { -lean_object* x_354; lean_object* x_355; lean_object* x_356; -lean_dec(x_296); -x_354 = lean_ctor_get(x_353, 0); -lean_inc(x_354); -if (lean_is_exclusive(x_353)) { - lean_ctor_release(x_353, 0); - x_355 = x_353; +lean_object* x_350; lean_object* x_351; lean_object* x_352; +lean_dec(x_292); +x_350 = lean_ctor_get(x_349, 0); +lean_inc(x_350); +if (lean_is_exclusive(x_349)) { + lean_ctor_release(x_349, 0); + x_351 = x_349; } else { - lean_dec_ref(x_353); - x_355 = lean_box(0); + lean_dec_ref(x_349); + x_351 = lean_box(0); } -if (lean_is_scalar(x_355)) { - x_356 = lean_alloc_ctor(0, 1, 0); +if (lean_is_scalar(x_351)) { + x_352 = lean_alloc_ctor(0, 1, 0); } else { - x_356 = x_355; + x_352 = x_351; } -lean_ctor_set(x_356, 0, x_354); -x_297 = x_356; -goto block_352; +lean_ctor_set(x_352, 0, x_350); +x_293 = x_352; +goto block_348; } else { -lean_object* x_357; lean_object* x_358; lean_object* x_359; uint8_t x_360; -x_357 = lean_ctor_get(x_353, 0); -lean_inc(x_357); -if (lean_is_exclusive(x_353)) { - lean_ctor_release(x_353, 0); - x_358 = x_353; +lean_object* x_353; lean_object* x_354; lean_object* x_355; uint8_t x_356; +x_353 = lean_ctor_get(x_349, 0); +lean_inc(x_353); +if (lean_is_exclusive(x_349)) { + lean_ctor_release(x_349, 0); + x_354 = x_349; } else { - lean_dec_ref(x_353); - x_358 = lean_box(0); + lean_dec_ref(x_349); + x_354 = lean_box(0); } -x_359 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__1; -x_360 = lean_string_dec_eq(x_357, x_359); -if (x_360 == 0) +x_355 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__1; +x_356 = lean_string_dec_eq(x_353, x_355); +if (x_356 == 0) { -lean_object* x_361; lean_object* x_362; lean_object* x_363; -x_361 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__2; -x_362 = lean_string_append(x_361, x_357); -lean_dec(x_357); -x_363 = l_Lean_Syntax_decodeNameLit(x_362); -if (lean_obj_tag(x_363) == 0) +lean_object* x_357; uint8_t x_358; +x_357 = l_String_toName(x_353); +x_358 = l_Lean_Name_isAnonymous(x_357); +if (x_358 == 0) { -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_364 = lean_unsigned_to_nat(80u); -x_365 = l_Lean_Json_pretty(x_296, x_364); -x_366 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__3; -x_367 = lean_string_append(x_366, x_365); -lean_dec(x_365); -x_368 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__4; -x_369 = lean_string_append(x_367, x_368); -if (lean_is_scalar(x_358)) { - x_370 = lean_alloc_ctor(0, 1, 0); +lean_object* x_359; +lean_dec(x_292); +if (lean_is_scalar(x_354)) { + x_359 = lean_alloc_ctor(1, 1, 0); } else { - x_370 = x_358; - lean_ctor_set_tag(x_370, 0); + x_359 = x_354; } -lean_ctor_set(x_370, 0, x_369); -x_297 = x_370; -goto block_352; +lean_ctor_set(x_359, 0, x_357); +x_293 = x_359; +goto block_348; } else { -lean_object* x_371; lean_object* x_372; -lean_dec(x_296); -x_371 = lean_ctor_get(x_363, 0); -lean_inc(x_371); -lean_dec(x_363); -if (lean_is_scalar(x_358)) { - x_372 = lean_alloc_ctor(1, 1, 0); +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_dec(x_357); +x_360 = lean_unsigned_to_nat(80u); +x_361 = l_Lean_Json_pretty(x_292, x_360); +x_362 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__2; +x_363 = lean_string_append(x_362, x_361); +lean_dec(x_361); +x_364 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__3; +x_365 = lean_string_append(x_363, x_364); +if (lean_is_scalar(x_354)) { + x_366 = lean_alloc_ctor(0, 1, 0); } else { - x_372 = x_358; + x_366 = x_354; + lean_ctor_set_tag(x_366, 0); } -lean_ctor_set(x_372, 0, x_371); -x_297 = x_372; -goto block_352; +lean_ctor_set(x_366, 0, x_365); +x_293 = x_366; +goto block_348; } } else { -lean_object* x_373; -lean_dec(x_358); -lean_dec(x_357); -lean_dec(x_296); -x_373 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__5; -x_297 = x_373; -goto block_352; +lean_object* x_367; +lean_dec(x_354); +lean_dec(x_353); +lean_dec(x_292); +x_367 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__4; +x_293 = x_367; +goto block_348; } } -block_352: +block_348: { -lean_object* x_298; -x_298 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveGoal_dec____x40_Lean_Widget_InteractiveGoal___hyg_1387____spec__3(x_297, x_2); -if (lean_obj_tag(x_298) == 0) +lean_object* x_294; +x_294 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveGoal_dec____x40_Lean_Widget_InteractiveGoal___hyg_1387____spec__3(x_293, x_2); +if (lean_obj_tag(x_294) == 0) { -lean_object* x_299; lean_object* x_300; lean_object* x_301; -lean_dec(x_295); -lean_dec(x_288); -lean_dec(x_287); -lean_dec(x_280); -lean_dec(x_268); -lean_dec(x_254); +lean_object* x_295; lean_object* x_296; lean_object* x_297; +lean_dec(x_291); +lean_dec(x_284); +lean_dec(x_283); +lean_dec(x_276); +lean_dec(x_264); +lean_dec(x_250); lean_dec(x_2); -x_299 = lean_ctor_get(x_298, 0); -lean_inc(x_299); -if (lean_is_exclusive(x_298)) { - lean_ctor_release(x_298, 0); - x_300 = x_298; +x_295 = lean_ctor_get(x_294, 0); +lean_inc(x_295); +if (lean_is_exclusive(x_294)) { + lean_ctor_release(x_294, 0); + x_296 = x_294; } else { - lean_dec_ref(x_298); - x_300 = lean_box(0); + lean_dec_ref(x_294); + x_296 = lean_box(0); } -if (lean_is_scalar(x_300)) { - x_301 = lean_alloc_ctor(0, 1, 0); +if (lean_is_scalar(x_296)) { + x_297 = lean_alloc_ctor(0, 1, 0); } else { - x_301 = x_300; + x_297 = x_296; } -lean_ctor_set(x_301, 0, x_299); -return x_301; +lean_ctor_set(x_297, 0, x_295); +return x_297; } else { -lean_object* x_302; -x_302 = lean_ctor_get(x_254, 6); +lean_object* x_298; +x_298 = lean_ctor_get(x_250, 6); +lean_inc(x_298); +if (lean_obj_tag(x_298) == 0) +{ +lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; +x_299 = lean_ctor_get(x_294, 0); +lean_inc(x_299); +if (lean_is_exclusive(x_294)) { + lean_ctor_release(x_294, 0); + x_300 = x_294; +} else { + lean_dec_ref(x_294); + x_300 = lean_box(0); +} +x_301 = lean_box(0); +x_302 = lean_ctor_get(x_250, 7); lean_inc(x_302); +lean_dec(x_250); if (lean_obj_tag(x_302) == 0) { -lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; -x_303 = lean_ctor_get(x_298, 0); -lean_inc(x_303); -if (lean_is_exclusive(x_298)) { - lean_ctor_release(x_298, 0); - x_304 = x_298; +lean_object* x_303; lean_object* x_304; lean_object* x_305; +lean_dec(x_2); +x_303 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_303, 0, x_264); +lean_ctor_set(x_303, 1, x_276); +lean_ctor_set(x_303, 2, x_283); +x_304 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_304, 0, x_303); +lean_ctor_set(x_304, 1, x_284); +lean_ctor_set(x_304, 2, x_291); +lean_ctor_set(x_304, 3, x_299); +lean_ctor_set(x_304, 4, x_301); +lean_ctor_set(x_304, 5, x_301); +if (lean_is_scalar(x_300)) { + x_305 = lean_alloc_ctor(1, 1, 0); } else { - lean_dec_ref(x_298); - x_304 = lean_box(0); + x_305 = x_300; } -x_305 = lean_box(0); -x_306 = lean_ctor_get(x_254, 7); -lean_inc(x_306); -lean_dec(x_254); -if (lean_obj_tag(x_306) == 0) +lean_ctor_set(x_305, 0, x_304); +return x_305; +} +else { -lean_object* x_307; lean_object* x_308; lean_object* x_309; -lean_dec(x_2); -x_307 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_307, 0, x_268); -lean_ctor_set(x_307, 1, x_280); -lean_ctor_set(x_307, 2, x_287); -x_308 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_308, 0, x_307); -lean_ctor_set(x_308, 1, x_288); -lean_ctor_set(x_308, 2, x_295); -lean_ctor_set(x_308, 3, x_303); -lean_ctor_set(x_308, 4, x_305); -lean_ctor_set(x_308, 5, x_305); -if (lean_is_scalar(x_304)) { - x_309 = lean_alloc_ctor(1, 1, 0); -} else { - x_309 = x_304; -} -lean_ctor_set(x_309, 0, x_308); -return x_309; -} -else -{ -lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; -lean_dec(x_304); -x_310 = lean_ctor_get(x_306, 0); -lean_inc(x_310); -if (lean_is_exclusive(x_306)) { - lean_ctor_release(x_306, 0); - x_311 = x_306; +lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; +lean_dec(x_300); +x_306 = lean_ctor_get(x_302, 0); +lean_inc(x_306); +if (lean_is_exclusive(x_302)) { + lean_ctor_release(x_302, 0); + x_307 = x_302; } else { - lean_dec_ref(x_306); - x_311 = lean_box(0); + lean_dec_ref(x_302); + x_307 = lean_box(0); } -x_312 = l_Lean_Json_getBool_x3f(x_310); -lean_dec(x_310); -x_313 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__7(x_312, x_2); +x_308 = l_Lean_Json_getBool_x3f(x_306); +lean_dec(x_306); +x_309 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__7(x_308, x_2); lean_dec(x_2); -if (lean_obj_tag(x_313) == 0) +if (lean_obj_tag(x_309) == 0) { -lean_object* x_314; lean_object* x_315; lean_object* x_316; -lean_dec(x_311); -lean_dec(x_303); -lean_dec(x_295); -lean_dec(x_288); -lean_dec(x_287); -lean_dec(x_280); -lean_dec(x_268); -x_314 = lean_ctor_get(x_313, 0); -lean_inc(x_314); -if (lean_is_exclusive(x_313)) { - lean_ctor_release(x_313, 0); - x_315 = x_313; +lean_object* x_310; lean_object* x_311; lean_object* x_312; +lean_dec(x_307); +lean_dec(x_299); +lean_dec(x_291); +lean_dec(x_284); +lean_dec(x_283); +lean_dec(x_276); +lean_dec(x_264); +x_310 = lean_ctor_get(x_309, 0); +lean_inc(x_310); +if (lean_is_exclusive(x_309)) { + lean_ctor_release(x_309, 0); + x_311 = x_309; } else { - lean_dec_ref(x_313); - x_315 = lean_box(0); + lean_dec_ref(x_309); + x_311 = lean_box(0); } -if (lean_is_scalar(x_315)) { - x_316 = lean_alloc_ctor(0, 1, 0); +if (lean_is_scalar(x_311)) { + x_312 = lean_alloc_ctor(0, 1, 0); } else { - x_316 = x_315; + x_312 = x_311; } -lean_ctor_set(x_316, 0, x_314); -return x_316; +lean_ctor_set(x_312, 0, x_310); +return x_312; } else { -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_317 = lean_ctor_get(x_313, 0); -lean_inc(x_317); -if (lean_is_exclusive(x_313)) { - lean_ctor_release(x_313, 0); - x_318 = x_313; +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_313 = lean_ctor_get(x_309, 0); +lean_inc(x_313); +if (lean_is_exclusive(x_309)) { + lean_ctor_release(x_309, 0); + x_314 = x_309; } else { - lean_dec_ref(x_313); - x_318 = lean_box(0); + lean_dec_ref(x_309); + x_314 = lean_box(0); } -if (lean_is_scalar(x_311)) { - x_319 = lean_alloc_ctor(1, 1, 0); +if (lean_is_scalar(x_307)) { + x_315 = lean_alloc_ctor(1, 1, 0); } else { - x_319 = x_311; + x_315 = x_307; } -lean_ctor_set(x_319, 0, x_317); -x_320 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_320, 0, x_268); -lean_ctor_set(x_320, 1, x_280); -lean_ctor_set(x_320, 2, x_287); -x_321 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_321, 0, x_320); -lean_ctor_set(x_321, 1, x_288); -lean_ctor_set(x_321, 2, x_295); -lean_ctor_set(x_321, 3, x_303); -lean_ctor_set(x_321, 4, x_305); -lean_ctor_set(x_321, 5, x_319); -if (lean_is_scalar(x_318)) { - x_322 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_315, 0, x_313); +x_316 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_316, 0, x_264); +lean_ctor_set(x_316, 1, x_276); +lean_ctor_set(x_316, 2, x_283); +x_317 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_317, 0, x_316); +lean_ctor_set(x_317, 1, x_284); +lean_ctor_set(x_317, 2, x_291); +lean_ctor_set(x_317, 3, x_299); +lean_ctor_set(x_317, 4, x_301); +lean_ctor_set(x_317, 5, x_315); +if (lean_is_scalar(x_314)) { + x_318 = lean_alloc_ctor(1, 1, 0); } else { - x_322 = x_318; + x_318 = x_314; } -lean_ctor_set(x_322, 0, x_321); -return x_322; +lean_ctor_set(x_318, 0, x_317); +return x_318; } } } else { -lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; -x_323 = lean_ctor_get(x_298, 0); -lean_inc(x_323); -lean_dec(x_298); -x_324 = lean_ctor_get(x_302, 0); -lean_inc(x_324); -if (lean_is_exclusive(x_302)) { - lean_ctor_release(x_302, 0); - x_325 = x_302; +lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; +x_319 = lean_ctor_get(x_294, 0); +lean_inc(x_319); +lean_dec(x_294); +x_320 = lean_ctor_get(x_298, 0); +lean_inc(x_320); +if (lean_is_exclusive(x_298)) { + lean_ctor_release(x_298, 0); + x_321 = x_298; } else { - lean_dec_ref(x_302); - x_325 = lean_box(0); + lean_dec_ref(x_298); + x_321 = lean_box(0); } -x_326 = l_Lean_Json_getBool_x3f(x_324); -lean_dec(x_324); -x_327 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__7(x_326, x_2); -if (lean_obj_tag(x_327) == 0) +x_322 = l_Lean_Json_getBool_x3f(x_320); +lean_dec(x_320); +x_323 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__7(x_322, x_2); +if (lean_obj_tag(x_323) == 0) { -lean_object* x_328; lean_object* x_329; lean_object* x_330; -lean_dec(x_325); -lean_dec(x_323); -lean_dec(x_295); -lean_dec(x_288); -lean_dec(x_287); -lean_dec(x_280); -lean_dec(x_268); -lean_dec(x_254); +lean_object* x_324; lean_object* x_325; lean_object* x_326; +lean_dec(x_321); +lean_dec(x_319); +lean_dec(x_291); +lean_dec(x_284); +lean_dec(x_283); +lean_dec(x_276); +lean_dec(x_264); +lean_dec(x_250); lean_dec(x_2); -x_328 = lean_ctor_get(x_327, 0); -lean_inc(x_328); -if (lean_is_exclusive(x_327)) { - lean_ctor_release(x_327, 0); - x_329 = x_327; +x_324 = lean_ctor_get(x_323, 0); +lean_inc(x_324); +if (lean_is_exclusive(x_323)) { + lean_ctor_release(x_323, 0); + x_325 = x_323; } else { - lean_dec_ref(x_327); - x_329 = lean_box(0); + lean_dec_ref(x_323); + x_325 = lean_box(0); } -if (lean_is_scalar(x_329)) { - x_330 = lean_alloc_ctor(0, 1, 0); +if (lean_is_scalar(x_325)) { + x_326 = lean_alloc_ctor(0, 1, 0); } else { - x_330 = x_329; + x_326 = x_325; } -lean_ctor_set(x_330, 0, x_328); -return x_330; +lean_ctor_set(x_326, 0, x_324); +return x_326; } else { -lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; -x_331 = lean_ctor_get(x_327, 0); -lean_inc(x_331); -if (lean_is_exclusive(x_327)) { - lean_ctor_release(x_327, 0); - x_332 = x_327; +lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; +x_327 = lean_ctor_get(x_323, 0); +lean_inc(x_327); +if (lean_is_exclusive(x_323)) { + lean_ctor_release(x_323, 0); + x_328 = x_323; } else { - lean_dec_ref(x_327); - x_332 = lean_box(0); + lean_dec_ref(x_323); + x_328 = lean_box(0); } -if (lean_is_scalar(x_325)) { - x_333 = lean_alloc_ctor(1, 1, 0); +if (lean_is_scalar(x_321)) { + x_329 = lean_alloc_ctor(1, 1, 0); } else { - x_333 = x_325; + x_329 = x_321; } -lean_ctor_set(x_333, 0, x_331); -x_334 = lean_ctor_get(x_254, 7); -lean_inc(x_334); -lean_dec(x_254); -if (lean_obj_tag(x_334) == 0) +lean_ctor_set(x_329, 0, x_327); +x_330 = lean_ctor_get(x_250, 7); +lean_inc(x_330); +lean_dec(x_250); +if (lean_obj_tag(x_330) == 0) { -lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; +lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_dec(x_2); -x_335 = lean_box(0); -x_336 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_336, 0, x_268); -lean_ctor_set(x_336, 1, x_280); -lean_ctor_set(x_336, 2, x_287); -x_337 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_337, 0, x_336); -lean_ctor_set(x_337, 1, x_288); -lean_ctor_set(x_337, 2, x_295); -lean_ctor_set(x_337, 3, x_323); -lean_ctor_set(x_337, 4, x_333); -lean_ctor_set(x_337, 5, x_335); -if (lean_is_scalar(x_332)) { - x_338 = lean_alloc_ctor(1, 1, 0); -} else { - x_338 = x_332; -} -lean_ctor_set(x_338, 0, x_337); -return x_338; -} -else -{ -lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; -lean_dec(x_332); -x_339 = lean_ctor_get(x_334, 0); -lean_inc(x_339); -if (lean_is_exclusive(x_334)) { - lean_ctor_release(x_334, 0); - x_340 = x_334; +x_331 = lean_box(0); +x_332 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_332, 0, x_264); +lean_ctor_set(x_332, 1, x_276); +lean_ctor_set(x_332, 2, x_283); +x_333 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_333, 0, x_332); +lean_ctor_set(x_333, 1, x_284); +lean_ctor_set(x_333, 2, x_291); +lean_ctor_set(x_333, 3, x_319); +lean_ctor_set(x_333, 4, x_329); +lean_ctor_set(x_333, 5, x_331); +if (lean_is_scalar(x_328)) { + x_334 = lean_alloc_ctor(1, 1, 0); } else { - lean_dec_ref(x_334); - x_340 = lean_box(0); + x_334 = x_328; +} +lean_ctor_set(x_334, 0, x_333); +return x_334; +} +else +{ +lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; +lean_dec(x_328); +x_335 = lean_ctor_get(x_330, 0); +lean_inc(x_335); +if (lean_is_exclusive(x_330)) { + lean_ctor_release(x_330, 0); + x_336 = x_330; +} else { + lean_dec_ref(x_330); + x_336 = lean_box(0); } -x_341 = l_Lean_Json_getBool_x3f(x_339); -lean_dec(x_339); -x_342 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__7(x_341, x_2); +x_337 = l_Lean_Json_getBool_x3f(x_335); +lean_dec(x_335); +x_338 = l_MonadExcept_ofExcept___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__7(x_337, x_2); lean_dec(x_2); -if (lean_obj_tag(x_342) == 0) +if (lean_obj_tag(x_338) == 0) { -lean_object* x_343; lean_object* x_344; lean_object* x_345; -lean_dec(x_340); -lean_dec(x_333); -lean_dec(x_323); -lean_dec(x_295); -lean_dec(x_288); -lean_dec(x_287); -lean_dec(x_280); -lean_dec(x_268); -x_343 = lean_ctor_get(x_342, 0); -lean_inc(x_343); -if (lean_is_exclusive(x_342)) { - lean_ctor_release(x_342, 0); - x_344 = x_342; +lean_object* x_339; lean_object* x_340; lean_object* x_341; +lean_dec(x_336); +lean_dec(x_329); +lean_dec(x_319); +lean_dec(x_291); +lean_dec(x_284); +lean_dec(x_283); +lean_dec(x_276); +lean_dec(x_264); +x_339 = lean_ctor_get(x_338, 0); +lean_inc(x_339); +if (lean_is_exclusive(x_338)) { + lean_ctor_release(x_338, 0); + x_340 = x_338; } else { - lean_dec_ref(x_342); - x_344 = lean_box(0); + lean_dec_ref(x_338); + x_340 = lean_box(0); } -if (lean_is_scalar(x_344)) { - x_345 = lean_alloc_ctor(0, 1, 0); +if (lean_is_scalar(x_340)) { + x_341 = lean_alloc_ctor(0, 1, 0); } else { - x_345 = x_344; + x_341 = x_340; } -lean_ctor_set(x_345, 0, x_343); -return x_345; +lean_ctor_set(x_341, 0, x_339); +return x_341; } 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; -x_346 = lean_ctor_get(x_342, 0); -lean_inc(x_346); -if (lean_is_exclusive(x_342)) { - lean_ctor_release(x_342, 0); - x_347 = x_342; +lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; +x_342 = lean_ctor_get(x_338, 0); +lean_inc(x_342); +if (lean_is_exclusive(x_338)) { + lean_ctor_release(x_338, 0); + x_343 = x_338; } else { - lean_dec_ref(x_342); - x_347 = lean_box(0); + lean_dec_ref(x_338); + x_343 = lean_box(0); } -if (lean_is_scalar(x_340)) { - x_348 = lean_alloc_ctor(1, 1, 0); +if (lean_is_scalar(x_336)) { + x_344 = lean_alloc_ctor(1, 1, 0); } else { - x_348 = x_340; + x_344 = x_336; } -lean_ctor_set(x_348, 0, x_346); -x_349 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_349, 0, x_268); -lean_ctor_set(x_349, 1, x_280); -lean_ctor_set(x_349, 2, x_287); -x_350 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_350, 0, x_349); -lean_ctor_set(x_350, 1, x_288); -lean_ctor_set(x_350, 2, x_295); -lean_ctor_set(x_350, 3, x_323); -lean_ctor_set(x_350, 4, x_333); -lean_ctor_set(x_350, 5, x_348); -if (lean_is_scalar(x_347)) { - x_351 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_344, 0, x_342); +x_345 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_345, 0, x_264); +lean_ctor_set(x_345, 1, x_276); +lean_ctor_set(x_345, 2, x_283); +x_346 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_346, 0, x_345); +lean_ctor_set(x_346, 1, x_284); +lean_ctor_set(x_346, 2, x_291); +lean_ctor_set(x_346, 3, x_319); +lean_ctor_set(x_346, 4, x_329); +lean_ctor_set(x_346, 5, x_344); +if (lean_is_scalar(x_343)) { + x_347 = lean_alloc_ctor(1, 1, 0); } else { - x_351 = x_347; + x_347 = x_343; } -lean_ctor_set(x_351, 0, x_350); -return x_351; +lean_ctor_set(x_347, 0, x_346); +return x_347; } } } @@ -10352,19 +10334,19 @@ return x_351; } else { -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_dec(x_254); +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_dec(x_250); lean_dec(x_2); -x_386 = lean_unsigned_to_nat(80u); -x_387 = l_Lean_Json_pretty(x_255, x_386); -x_388 = l_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____closed__1; -x_389 = lean_string_append(x_388, x_387); -lean_dec(x_387); -x_390 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__4; -x_391 = lean_string_append(x_389, x_390); -x_392 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_392, 0, x_391); -return x_392; +x_380 = lean_unsigned_to_nat(80u); +x_381 = l_Lean_Json_pretty(x_251, x_380); +x_382 = l_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____closed__1; +x_383 = lean_string_append(x_382, x_381); +lean_dec(x_381); +x_384 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__3; +x_385 = lean_string_append(x_383, x_384); +x_386 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_386, 0, x_385); +return x_386; } } } @@ -11214,7 +11196,7 @@ x_61 = l_Lean_Json_pretty(x_6, x_60); x_62 = l_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____closed__1; x_63 = lean_string_append(x_62, x_61); lean_dec(x_61); -x_64 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__4; +x_64 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__3; x_65 = lean_string_append(x_63, x_64); lean_ctor_set_tag(x_3, 0); lean_ctor_set(x_3, 0, x_65); @@ -11484,7 +11466,7 @@ x_119 = l_Lean_Json_pretty(x_67, x_118); x_120 = l_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____closed__1; x_121 = lean_string_append(x_120, x_119); lean_dec(x_119); -x_122 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__4; +x_122 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__3; x_123 = lean_string_append(x_121, x_122); x_124 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_124, 0, x_123); @@ -12691,7 +12673,7 @@ x_22 = l_Lean_Json_pretty(x_5, x_21); x_23 = l_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____closed__1; x_24 = lean_string_append(x_23, x_22); lean_dec(x_22); -x_25 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__4; +x_25 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__3; x_26 = lean_string_append(x_24, x_25); lean_ctor_set_tag(x_3, 0); lean_ctor_set(x_3, 0, x_26); @@ -12772,7 +12754,7 @@ x_44 = l_Lean_Json_pretty(x_27, x_43); x_45 = l_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____closed__1; x_46 = lean_string_append(x_45, x_44); lean_dec(x_44); -x_47 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__4; +x_47 = l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__3; x_48 = lean_string_append(x_46, x_47); x_49 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_49, 0, x_48); @@ -13102,7 +13084,7 @@ lean_inc(x_23); x_24 = lean_ctor_get(x_22, 1); lean_inc(x_24); lean_dec(x_22); -x_25 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_2, x_6, x_7, x_8, x_9, x_24); +x_25 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_2, x_6, x_7, x_8, x_9, x_24); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -14071,7 +14053,7 @@ x_61 = lean_ctor_get(x_52, 3); lean_inc(x_61); lean_dec(x_52); x_62 = lean_simp_macro_scopes(x_60); -x_63 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_61, x_11, x_12, x_13, x_14, x_15); +x_63 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_61, x_11, x_12, x_13, x_14, x_15); x_64 = lean_ctor_get(x_63, 0); lean_inc(x_64); x_65 = lean_ctor_get(x_63, 1); @@ -14246,7 +14228,7 @@ goto block_130; block_119: { lean_object* x_108; -x_108 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_103, x_11, x_12, x_13, x_14, x_107); +x_108 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_103, x_11, x_12, x_13, x_14, x_107); if (x_3 == 0) { lean_object* x_109; lean_object* x_110; lean_object* x_111; @@ -14277,7 +14259,7 @@ lean_inc(x_112); x_113 = lean_ctor_get(x_108, 1); lean_inc(x_113); lean_dec(x_108); -x_114 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_104, x_11, x_12, x_13, x_14, x_113); +x_114 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_104, x_11, x_12, x_13, x_14, x_113); x_115 = lean_ctor_get(x_114, 0); lean_inc(x_115); x_116 = lean_ctor_get(x_114, 1); @@ -14506,7 +14488,7 @@ x_155 = lean_ctor_get(x_52, 3); lean_inc(x_155); lean_dec(x_52); x_156 = lean_simp_macro_scopes(x_154); -x_157 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_155, x_11, x_12, x_13, x_14, x_15); +x_157 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_155, x_11, x_12, x_13, x_14, x_15); x_158 = lean_ctor_get(x_157, 0); lean_inc(x_158); x_159 = lean_ctor_get(x_157, 1); @@ -14682,7 +14664,7 @@ goto block_224; block_213: { lean_object* x_202; -x_202 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_197, x_11, x_12, x_13, x_14, x_201); +x_202 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_197, x_11, x_12, x_13, x_14, x_201); if (x_3 == 0) { lean_object* x_203; lean_object* x_204; lean_object* x_205; @@ -14713,7 +14695,7 @@ lean_inc(x_206); x_207 = lean_ctor_get(x_202, 1); lean_inc(x_207); lean_dec(x_202); -x_208 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_198, x_11, x_12, x_13, x_14, x_207); +x_208 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_198, x_11, x_12, x_13, x_14, x_207); x_209 = lean_ctor_get(x_208, 0); lean_inc(x_209); x_210 = lean_ctor_get(x_208, 1); @@ -15376,7 +15358,7 @@ x_66 = lean_ctor_get(x_57, 3); lean_inc(x_66); lean_dec(x_57); x_67 = lean_simp_macro_scopes(x_65); -x_68 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_66, x_11, x_12, x_13, x_14, x_15); +x_68 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_66, x_11, x_12, x_13, x_14, x_15); x_69 = lean_ctor_get(x_68, 0); lean_inc(x_69); x_70 = lean_ctor_get(x_68, 1); @@ -15551,7 +15533,7 @@ goto block_135; block_124: { lean_object* x_113; -x_113 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_108, x_11, x_12, x_13, x_14, x_112); +x_113 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_108, x_11, x_12, x_13, x_14, x_112); if (x_3 == 0) { lean_object* x_114; lean_object* x_115; lean_object* x_116; @@ -15582,7 +15564,7 @@ lean_inc(x_117); x_118 = lean_ctor_get(x_113, 1); lean_inc(x_118); lean_dec(x_113); -x_119 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_109, x_11, x_12, x_13, x_14, x_118); +x_119 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_109, x_11, x_12, x_13, x_14, x_118); x_120 = lean_ctor_get(x_119, 0); lean_inc(x_120); x_121 = lean_ctor_get(x_119, 1); @@ -15811,7 +15793,7 @@ x_160 = lean_ctor_get(x_57, 3); lean_inc(x_160); lean_dec(x_57); x_161 = lean_simp_macro_scopes(x_159); -x_162 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_160, x_11, x_12, x_13, x_14, x_15); +x_162 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_160, x_11, x_12, x_13, x_14, x_15); x_163 = lean_ctor_get(x_162, 0); lean_inc(x_163); x_164 = lean_ctor_get(x_162, 1); @@ -15987,7 +15969,7 @@ goto block_229; block_218: { lean_object* x_207; -x_207 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_202, x_11, x_12, x_13, x_14, x_206); +x_207 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_202, x_11, x_12, x_13, x_14, x_206); if (x_3 == 0) { lean_object* x_208; lean_object* x_209; lean_object* x_210; @@ -16018,7 +16000,7 @@ lean_inc(x_211); x_212 = lean_ctor_get(x_207, 1); lean_inc(x_212); lean_dec(x_207); -x_213 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_203, x_11, x_12, x_13, x_14, x_212); +x_213 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_203, x_11, x_12, x_13, x_14, x_212); x_214 = lean_ctor_get(x_213, 0); lean_inc(x_214); x_215 = lean_ctor_get(x_213, 1); @@ -16816,7 +16798,7 @@ x_26 = lean_ctor_get(x_19, 4); lean_inc(x_26); lean_inc(x_25); x_27 = lean_apply_2(x_2, x_25, x_19); -x_28 = l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__3___rarg(x_25, x_26, x_27, x_3, x_4, x_5, x_6, x_10); +x_28 = l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2___rarg(x_25, x_26, x_27, x_3, x_4, x_5, x_6, x_10); return x_28; } } @@ -16979,7 +16961,7 @@ goto block_71; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; x_27 = lean_ctor_get(x_6, 2); lean_inc(x_27); -x_28 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(x_27, x_7, x_8, x_9, x_10, x_26); +x_28 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_27, x_7, x_8, x_9, x_10, x_26); x_29 = lean_ctor_get(x_28, 0); lean_inc(x_29); x_30 = lean_ctor_get(x_28, 1); @@ -17212,11 +17194,11 @@ lean_object* x_7; lean_object* x_8; uint8_t x_9; lean_object* x_10; uint8_t x_11 x_7 = lean_ctor_get(x_4, 2); lean_inc(x_7); x_8 = l_Lean_Widget_goalToInteractive___closed__1; -x_9 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_7, x_8); +x_9 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_7, x_8); x_10 = l_Lean_Widget_goalToInteractive___closed__2; -x_11 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_7, x_10); +x_11 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_7, x_10); x_12 = l_Lean_Widget_goalToInteractive___closed__3; -x_13 = l_Lean_Option_get___at_Lean_getSanitizeNames___spec__1(x_7, x_12); +x_13 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_7, x_12); lean_dec(x_7); x_14 = lean_box(x_9); x_15 = lean_box(x_11); @@ -17611,8 +17593,6 @@ l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBun lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__3); l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__4 = _init_l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__4(); lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__4); -l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__5 = _init_l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__5(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____spec__2___closed__5); l_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____closed__1 = _init_l_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____closed__1(); lean_mark_persistent(l_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____closed__1); l_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____closed__2 = _init_l_Lean_Widget_instRpcEncodableInteractiveHypothesisBundle_dec____x40_Lean_Widget_InteractiveGoal___hyg_5____closed__2(); diff --git a/stage0/stdlib/Lean/Widget/UserWidget.c b/stage0/stdlib/Lean/Widget/UserWidget.c index ead90a33e3f1..88e3dfdb4adb 100644 --- a/stage0/stdlib/Lean/Widget/UserWidget.c +++ b/stage0/stdlib/Lean/Widget/UserWidget.c @@ -58,6 +58,7 @@ static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUs static lean_object* l_Lean_Widget_getWidgets___rpc__wrapped___closed__1; static lean_object* l_Lean_Widget_instFromJsonUserWidgetDefinition___closed__1; LEAN_EXPORT lean_object* l_Lean_Widget_instToJsonGetWidgetsResponse; +static lean_object* l_Lean_Widget_widgetCmd___closed__17; LEAN_EXPORT lean_object* l_Lean_Widget_instInhabitedWidgetSource; static lean_object* l_Lean_Widget_getWidgetSource___lambda__5___closed__1; LEAN_EXPORT lean_object* l_Lean_Widget_getWidgetSource___rpc__wrapped; @@ -234,6 +235,7 @@ LEAN_EXPORT lean_object* l_Lean_Widget_saveWidgetInfo___at_Lean_Widget_elabWidge static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonUserWidgetDefinition____x40_Lean_Widget_UserWidget___hyg_223____closed__11; lean_object* l_Lean_Server_RequestM_readDoc___at_Lean_Server_RequestM_withWaitFindSnapAtPos___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Widget_elabWidgetCmd___spec__1___rarg___closed__2; +static lean_object* l_Lean_Widget_widgetCmd___closed__19; static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonGetWidgetSourceParams____x40_Lean_Widget_UserWidget___hyg_1026____closed__7; static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonUserWidget____x40_Lean_Widget_UserWidget___hyg_370____closed__2; static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonGetWidgetSourceParams____x40_Lean_Widget_UserWidget___hyg_1026____closed__11; @@ -275,6 +277,7 @@ LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_Widget_getWidgets___spec__1___ LEAN_EXPORT lean_object* l_Lean_Server_wrapRpcProcedure___at_Lean_Widget_getWidgets___rpc__wrapped___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Widget_getWidgets___rpc__wrapped___closed__2; lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_497_(lean_object*); +static lean_object* l_Lean_Widget_widgetCmd___closed__16; LEAN_EXPORT lean_object* l_Lean_Widget_instFromJsonGetWidgetSourceParams; LEAN_EXPORT uint8_t l_Lean_Widget_getWidgets___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_getWidgetSource___lambda__1(lean_object*, lean_object*, lean_object*); @@ -298,6 +301,7 @@ static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_626 static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_attributeImpl___closed__8; lean_object* l_Lean_Environment_evalConstCheck___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_InfoTree_deepestNodes___rarg(lean_object*, lean_object*); +static lean_object* l_Lean_Widget_widgetCmd___closed__18; static lean_object* l_Lean_Widget_getWidgets___rpc__wrapped___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_toJsonGetWidgetsResponse____x40_Lean_Widget_UserWidget___hyg_1946_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Widget_widgetInfosAt_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -7146,7 +7150,7 @@ static lean_object* _init_l_Lean_Widget_widgetCmd___closed__11() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("term", 4); +x_1 = lean_mk_string_from_bytes("ppSpace", 7); return x_1; } } @@ -7163,8 +7167,50 @@ return x_3; static lean_object* _init_l_Lean_Widget_widgetCmd___closed__13() { _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_Widget_widgetCmd___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_Widget_widgetCmd___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Widget_widgetCmd___closed__4; +x_2 = l_Lean_Widget_widgetCmd___closed__10; +x_3 = l_Lean_Widget_widgetCmd___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); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Widget_widgetCmd___closed__15() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("term", 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Widget_widgetCmd___closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Widget_widgetCmd___closed__15; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Widget_widgetCmd___closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Widget_widgetCmd___closed__16; x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7172,13 +7218,13 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Widget_widgetCmd___closed__14() { +static lean_object* _init_l_Lean_Widget_widgetCmd___closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Widget_widgetCmd___closed__4; -x_2 = l_Lean_Widget_widgetCmd___closed__10; -x_3 = l_Lean_Widget_widgetCmd___closed__13; +x_2 = l_Lean_Widget_widgetCmd___closed__14; +x_3 = l_Lean_Widget_widgetCmd___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); @@ -7186,13 +7232,13 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Widget_widgetCmd___closed__15() { +static lean_object* _init_l_Lean_Widget_widgetCmd___closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Widget_widgetCmd___closed__2; x_2 = lean_unsigned_to_nat(1022u); -x_3 = l_Lean_Widget_widgetCmd___closed__14; +x_3 = l_Lean_Widget_widgetCmd___closed__18; x_4 = lean_alloc_ctor(3, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -7204,7 +7250,7 @@ static lean_object* _init_l_Lean_Widget_widgetCmd() { _start: { lean_object* x_1; -x_1 = l_Lean_Widget_widgetCmd___closed__15; +x_1 = l_Lean_Widget_widgetCmd___closed__19; return x_1; } } @@ -7348,7 +7394,6 @@ x_13 = lean_unsigned_to_nat(2u); x_14 = l_Lean_Syntax_getArg(x_1, x_13); x_15 = lean_alloc_closure((void*)(l_Lean_Widget_elabWidgetCmd___lambda__1___boxed), 9, 1); lean_closure_set(x_15, 0, x_14); -lean_inc(x_2); x_16 = l_Lean_Elab_Command_runTermElabM___rarg(x_15, x_2, x_3, x_4); if (lean_obj_tag(x_16) == 0) { @@ -7817,6 +7862,14 @@ l_Lean_Widget_widgetCmd___closed__14 = _init_l_Lean_Widget_widgetCmd___closed__1 lean_mark_persistent(l_Lean_Widget_widgetCmd___closed__14); l_Lean_Widget_widgetCmd___closed__15 = _init_l_Lean_Widget_widgetCmd___closed__15(); lean_mark_persistent(l_Lean_Widget_widgetCmd___closed__15); +l_Lean_Widget_widgetCmd___closed__16 = _init_l_Lean_Widget_widgetCmd___closed__16(); +lean_mark_persistent(l_Lean_Widget_widgetCmd___closed__16); +l_Lean_Widget_widgetCmd___closed__17 = _init_l_Lean_Widget_widgetCmd___closed__17(); +lean_mark_persistent(l_Lean_Widget_widgetCmd___closed__17); +l_Lean_Widget_widgetCmd___closed__18 = _init_l_Lean_Widget_widgetCmd___closed__18(); +lean_mark_persistent(l_Lean_Widget_widgetCmd___closed__18); +l_Lean_Widget_widgetCmd___closed__19 = _init_l_Lean_Widget_widgetCmd___closed__19(); +lean_mark_persistent(l_Lean_Widget_widgetCmd___closed__19); l_Lean_Widget_widgetCmd = _init_l_Lean_Widget_widgetCmd(); lean_mark_persistent(l_Lean_Widget_widgetCmd); l___private_Lean_Widget_UserWidget_0__Lean_Widget_evalJsonUnsafe___closed__1 = _init_l___private_Lean_Widget_UserWidget_0__Lean_Widget_evalJsonUnsafe___closed__1(); diff --git a/tests/lean/1021.lean.expected.out b/tests/lean/1021.lean.expected.out index c986018f7483..59be053f7bab 100644 --- a/tests/lean/1021.lean.expected.out +++ b/tests/lean/1021.lean.expected.out @@ -1,8 +1,8 @@ -some { range := { pos := { line := 167, column := 42 }, +some { range := { pos := { line := 175, column := 42 }, charUtf16 := 42, - endPos := { line := 173, column := 31 }, + endPos := { line := 181, column := 31 }, endCharUtf16 := 31 }, - selectionRange := { pos := { line := 167, column := 46 }, + selectionRange := { pos := { line := 175, column := 46 }, charUtf16 := 46, - endPos := { line := 167, column := 58 }, + endPos := { line := 175, column := 58 }, endCharUtf16 := 58 } } diff --git a/tests/lean/StxQuot.lean.expected.out b/tests/lean/StxQuot.lean.expected.out index b1203c884529..685d3a2f8825 100644 --- a/tests/lean/StxQuot.lean.expected.out +++ b/tests/lean/StxQuot.lean.expected.out @@ -40,7 +40,7 @@ StxQuot.lean:19:15: error: expected term "(Term.match\n \"match\"\n []\n []\n [(Term.matchDiscr [] `a._@.UnhygienicMain._hyg.1)]\n \"with\"\n (Term.matchAlts\n [(Term.matchAlt \"|\" [[`a._@.UnhygienicMain._hyg.1]] \"=>\" `b._@.UnhygienicMain._hyg.1)\n (Term.matchAlt\n \"|\"\n [[(«term_+_» `a._@.UnhygienicMain._hyg.1 \"+\" (num \"1\"))]]\n \"=>\"\n («term_+_» `b._@.UnhygienicMain._hyg.1 \"+\" (num \"1\")))]))" "#[`a._@.UnhygienicMain._hyg.1, `b._@.UnhygienicMain._hyg.1]" "1" -"(Term.sufficesDecl [] `x._@.UnhygienicMain._hyg.1 (Term.fromTerm \"from\" `x._@.UnhygienicMain._hyg.1))" +"(Term.sufficesDecl\n (hygieneInfo `_@.UnhygienicMain._hyg.1)\n `x._@.UnhygienicMain._hyg.1\n (Term.fromTerm \"from\" `x._@.UnhygienicMain._hyg.1))" "#[(num \"1\"), [(num \"2\") (num \"3\")], (num \"4\")]" "#[(num \"2\")]" StxQuot.lean:95:39-95:44: error: unexpected antiquotation splice diff --git a/tests/lean/interactive/1265.lean.expected.out b/tests/lean/interactive/1265.lean.expected.out index bf1c2fa6a2cd..787c3640f3f8 100644 --- a/tests/lean/interactive/1265.lean.expected.out +++ b/tests/lean/interactive/1265.lean.expected.out @@ -1,7 +1,10 @@ {"textDocument": {"uri": "file://1265.lean"}, "position": {"line": 0, "character": 51}} {"items": - [{"label": "getId", "kind": 3, "detail": "Lean.Ident → Lean.Name"}, + [{"label": "getHygieneInfo", + "kind": 3, + "detail": "Lean.HygieneInfo → Lean.Name"}, + {"label": "getId", "kind": 3, "detail": "Lean.Ident → Lean.Name"}, {"label": "getName", "kind": 3, "detail": "Lean.NameLit → Lean.Name"}, {"label": "expandInterpolatedStr", "kind": 3, @@ -22,7 +25,10 @@ {"textDocument": {"uri": "file://1265.lean"}, "position": {"line": 2, "character": 53}} {"items": - [{"label": "getId", "kind": 3, "detail": "Lean.Ident → Lean.Name"}, + [{"label": "getHygieneInfo", + "kind": 3, + "detail": "Lean.HygieneInfo → Lean.Name"}, + {"label": "getId", "kind": 3, "detail": "Lean.Ident → Lean.Name"}, {"label": "getName", "kind": 3, "detail": "Lean.NameLit → Lean.Name"}, {"label": "expandInterpolatedStr", "kind": 3,